merged
authorhaftmann
Mon, 26 Oct 2009 09:03:57 +0100
changeset 33176 d6936fd7cda8
parent 33174 1f2051f41335 (diff)
parent 33175 2083bde13ce1 (current diff)
child 33177 edbd2c09176b
child 33178 70522979c7be
child 33184 de8cc01e8d9e
child 33189 82a40677c1f8
child 33201 e3d741e9d2fe
merged
src/HOL/IsaMakefile
src/HOL/Library/Convex_Euclidean_Space.thy
src/HOL/Library/Determinants.thy
src/HOL/Library/Euclidean_Space.thy
src/HOL/Library/Finite_Cartesian_Product.thy
src/HOL/Library/Library.thy
src/HOL/Library/Topology_Euclidean_Space.thy
--- a/Admin/isatest/isatest-makedist	Fri Oct 23 13:23:18 2009 +0200
+++ b/Admin/isatest/isatest-makedist	Mon Oct 26 09:03:57 2009 +0100
@@ -110,8 +110,8 @@
 sleep 15
 $SSH macbroy5 "$MAKEALL $HOME/settings/mac-poly"
 sleep 15
-#$SSH macbroy6 "$MAKEALL $HOME/settings/at-mac-poly-5.1-para"
-#sleep 15
+$SSH macbroy6 "sleep 10800; $MAKEALL $HOME/settings/at-mac-poly-5.1-para"
+sleep 15
 $SSH atbroy51 "$HOME/admin/isatest/isatest-annomaly"
 
 echo ------------------- spawned tests successfully --- `date` --- $HOSTNAME >> $DISTLOG 2>&1
--- a/Admin/isatest/settings/at-mac-poly-5.1-para	Fri Oct 23 13:23:18 2009 +0200
+++ b/Admin/isatest/settings/at-mac-poly-5.1-para	Mon Oct 26 09:03:57 2009 +0100
@@ -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 -t true -M 4 -q 2"
 
 HOL_USEDIR_OPTIONS="-p 2 -q 0"
--- a/Admin/isatest/settings/mac-poly-M4	Fri Oct 23 13:23:18 2009 +0200
+++ b/Admin/isatest/settings/mac-poly-M4	Mon Oct 26 09:03:57 2009 +0100
@@ -1,7 +1,7 @@
 # -*- shell-script -*- :mode=shellscript:
 
-  POLYML_HOME="/home/polyml/polyml-svn"
-  ML_SYSTEM="polyml-experimental"
+  POLYML_HOME="/home/polyml/polyml-5.2.1"
+  ML_SYSTEM="polyml-5.2.1"
   ML_PLATFORM="x86-darwin"
   ML_HOME="$POLYML_HOME/$ML_PLATFORM"
   ML_OPTIONS="--mutable 800 --immutable 2000"
@@ -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 -t true -q 2"
+ISABELLE_USEDIR_OPTIONS="-i false -d false -t true -M 4 -q 2"
 
 HOL_USEDIR_OPTIONS="-p 2 -q 0"
--- a/Admin/isatest/settings/mac-poly-M8	Fri Oct 23 13:23:18 2009 +0200
+++ b/Admin/isatest/settings/mac-poly-M8	Mon Oct 26 09:03:57 2009 +0100
@@ -1,7 +1,7 @@
 # -*- shell-script -*- :mode=shellscript:
 
-  POLYML_HOME="/home/polyml/polyml-svn"
-  ML_SYSTEM="polyml-experimental"
+  POLYML_HOME="/home/polyml/polyml-5.2.1"
+  ML_SYSTEM="polyml-5.2.1"
   ML_PLATFORM="x86-darwin"
   ML_HOME="$POLYML_HOME/$ML_PLATFORM"
   ML_OPTIONS="--mutable 800 --immutable 2000"
@@ -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 8 -t true -q 2"
+ISABELLE_USEDIR_OPTIONS="-i false -d false -t true -M 8 -q 2"
 
 HOL_USEDIR_OPTIONS="-p 2 -q 0"
--- a/Admin/isatest/settings/mac-poly64-M4	Fri Oct 23 13:23:18 2009 +0200
+++ b/Admin/isatest/settings/mac-poly64-M4	Mon Oct 26 09:03:57 2009 +0100
@@ -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 -q 2 -t true"
+ISABELLE_USEDIR_OPTIONS="-i false -d false -t true -M 4 -q 2"
 
 HOL_USEDIR_OPTIONS="-p 2 -q 2"
--- a/Admin/isatest/settings/mac-poly64-M8	Fri Oct 23 13:23:18 2009 +0200
+++ b/Admin/isatest/settings/mac-poly64-M8	Mon Oct 26 09:03:57 2009 +0100
@@ -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 8 -q 2 -t true"
+ISABELLE_USEDIR_OPTIONS="-i false -d false -t true -M 8 -q 2"
 
 HOL_USEDIR_OPTIONS="-p 2 -q 2"
--- a/NEWS	Fri Oct 23 13:23:18 2009 +0200
+++ b/NEWS	Mon Oct 26 09:03:57 2009 +0100
@@ -165,7 +165,8 @@
 
 * New implementation of quickcheck uses generic code generator;
 default generators are provided for all suitable HOL types, records
-and datatypes.
+and datatypes.  Old quickcheck can be re-activated importing
+theory Library/SML_Quickcheck.
 
 * Renamed theorems:
 Suc_eq_add_numeral_1 -> Suc_eq_plus1
--- a/doc-src/IsarImplementation/Thy/Logic.thy	Fri Oct 23 13:23:18 2009 +0200
+++ b/doc-src/IsarImplementation/Thy/Logic.thy	Mon Oct 26 09:03:57 2009 +0100
@@ -322,9 +322,9 @@
   @{index_ML fastype_of: "term -> typ"} \\
   @{index_ML lambda: "term -> term -> term"} \\
   @{index_ML betapply: "term * term -> term"} \\
-  @{index_ML Sign.declare_const: "Properties.T -> (binding * typ) * mixfix ->
+  @{index_ML Sign.declare_const: "(binding * typ) * mixfix ->
   theory -> term * theory"} \\
-  @{index_ML Sign.add_abbrev: "string -> Properties.T -> binding * term ->
+  @{index_ML Sign.add_abbrev: "string -> binding * term ->
   theory -> (term * term) * theory"} \\
   @{index_ML Sign.const_typargs: "theory -> string * typ -> typ list"} \\
   @{index_ML Sign.const_instance: "theory -> string * typ list -> typ"} \\
@@ -370,11 +370,11 @@
   "t u"}, with topmost @{text "\<beta>"}-conversion if @{text "t"} is an
   abstraction.
 
-  \item @{ML Sign.declare_const}~@{text "properties ((c, \<sigma>), mx)"}
+  \item @{ML Sign.declare_const}~@{text "((c, \<sigma>), mx)"}
   declares a new constant @{text "c :: \<sigma>"} with optional mixfix
   syntax.
 
-  \item @{ML Sign.add_abbrev}~@{text "print_mode properties (c, t)"}
+  \item @{ML Sign.add_abbrev}~@{text "print_mode (c, t)"}
   introduces a new term abbreviation @{text "c \<equiv> t"}.
 
   \item @{ML Sign.const_typargs}~@{text "thy (c, \<tau>)"} and @{ML
--- a/doc-src/IsarImplementation/Thy/ML.thy	Fri Oct 23 13:23:18 2009 +0200
+++ b/doc-src/IsarImplementation/Thy/ML.thy	Mon Oct 26 09:03:57 2009 +0100
@@ -317,7 +317,7 @@
   a theory by constant declararion and primitive definitions:
 
   \smallskip\begin{mldecls}
-  @{ML "Sign.declare_const: Properties.T -> (binding * typ) * mixfix
+  @{ML "Sign.declare_const: (binding * typ) * mixfix
   -> theory -> term * theory"} \\
   @{ML "Thm.add_def: bool -> bool -> binding * term -> theory -> thm * theory"}
   \end{mldecls}
@@ -329,7 +329,7 @@
   \smallskip\begin{mldecls}
   @{ML "(fn (t, thy) => Thm.add_def false false
   (Binding.name \"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})) thy)
-    (Sign.declare_const []
+    (Sign.declare_const
       ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn) thy)"}
   \end{mldecls}
 
@@ -344,7 +344,7 @@
 
   \smallskip\begin{mldecls}
 @{ML "thy
-|> Sign.declare_const [] ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
+|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
 |> (fn (t, thy) => thy
 |> Thm.add_def false false
      (Binding.name \"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})))"}
@@ -368,7 +368,7 @@
 
   \smallskip\begin{mldecls}
 @{ML "thy
-|> Sign.declare_const [] ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
+|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
 |-> (fn t => Thm.add_def false false
       (Binding.name \"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})))
 "}
@@ -378,7 +378,7 @@
 
   \smallskip\begin{mldecls}
 @{ML "thy
-|> Sign.declare_const [] ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
+|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
 |>> (fn t => Logic.mk_equals (t, @{term \"%x. x\"}))
 |-> (fn def => Thm.add_def false false (Binding.name \"bar_def\", def))
 "}
@@ -389,7 +389,7 @@
 
   \smallskip\begin{mldecls}
 @{ML "thy
-|> Sign.declare_const [] ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
+|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
 ||> Sign.add_path \"foobar\"
 |-> (fn t => Thm.add_def false false
       (Binding.name \"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})))
@@ -401,8 +401,8 @@
 
   \smallskip\begin{mldecls}
 @{ML "thy
-|> Sign.declare_const [] ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
-||>> Sign.declare_const [] ((Binding.name \"foobar\", @{typ \"foo => foo\"}), NoSyn)
+|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
+||>> Sign.declare_const ((Binding.name \"foobar\", @{typ \"foo => foo\"}), NoSyn)
 |-> (fn (t1, t2) => Thm.add_def false false
       (Binding.name \"bar_def\", Logic.mk_equals (t1, t2)))
 "}
@@ -447,7 +447,7 @@
   val consts = [\"foo\", \"bar\"];
 in
   thy
-  |> fold_map (fn const => Sign.declare_const []
+  |> fold_map (fn const => Sign.declare_const
        ((Binding.name const, @{typ \"foo => foo\"}), NoSyn)) consts
   |>> map (fn t => Logic.mk_equals (t, @{term \"%x. x\"}))
   |-> (fn defs => fold_map (fn def =>
@@ -486,11 +486,11 @@
   \smallskip\begin{mldecls}
 @{ML "thy
 |> tap (fn _ => writeln \"now adding constant\")
-|> Sign.declare_const [] ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
+|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
 ||>> `(fn thy => Sign.declared_const thy
          (Sign.full_name thy (Binding.name \"foobar\")))
 |-> (fn (t, b) => if b then I
-       else Sign.declare_const []
+       else Sign.declare_const
          ((Binding.name \"foobar\", @{typ \"foo => foo\"}), NoSyn) #> snd)
 "}
   \end{mldecls}
--- a/doc-src/IsarImplementation/Thy/Prelim.thy	Fri Oct 23 13:23:18 2009 +0200
+++ b/doc-src/IsarImplementation/Thy/Prelim.thy	Mon Oct 26 09:03:57 2009 +0100
@@ -689,19 +689,19 @@
   @{index_ML Long_Name.explode: "string -> string list"} \\
   \end{mldecls}
   \begin{mldecls}
-  @{index_ML_type NameSpace.naming} \\
-  @{index_ML NameSpace.default_naming: NameSpace.naming} \\
-  @{index_ML NameSpace.add_path: "string -> NameSpace.naming -> NameSpace.naming"} \\
-  @{index_ML NameSpace.full_name: "NameSpace.naming -> binding -> string"} \\
+  @{index_ML_type Name_Space.naming} \\
+  @{index_ML Name_Space.default_naming: Name_Space.naming} \\
+  @{index_ML Name_Space.add_path: "string -> Name_Space.naming -> Name_Space.naming"} \\
+  @{index_ML Name_Space.full_name: "Name_Space.naming -> binding -> string"} \\
   \end{mldecls}
   \begin{mldecls}
-  @{index_ML_type NameSpace.T} \\
-  @{index_ML NameSpace.empty: NameSpace.T} \\
-  @{index_ML NameSpace.merge: "NameSpace.T * NameSpace.T -> NameSpace.T"} \\
-  @{index_ML NameSpace.declare: "NameSpace.naming -> binding -> NameSpace.T ->
-  string * NameSpace.T"} \\
-  @{index_ML NameSpace.intern: "NameSpace.T -> string -> string"} \\
-  @{index_ML NameSpace.extern: "NameSpace.T -> string -> string"} \\
+  @{index_ML_type Name_Space.T} \\
+  @{index_ML Name_Space.empty: "string -> Name_Space.T"} \\
+  @{index_ML Name_Space.merge: "Name_Space.T * Name_Space.T -> Name_Space.T"} \\
+  @{index_ML Name_Space.declare: "bool -> Name_Space.naming -> binding -> Name_Space.T ->
+  string * Name_Space.T"} \\
+  @{index_ML Name_Space.intern: "Name_Space.T -> string -> string"} \\
+  @{index_ML Name_Space.extern: "Name_Space.T -> string -> string"} \\
   \end{mldecls}
 
   \begin{description}
@@ -719,41 +719,43 @@
   Long_Name.explode}~@{text "name"} convert between the packed string
   representation and the explicit list form of qualified names.
 
-  \item @{ML_type NameSpace.naming} represents the abstract concept of
+  \item @{ML_type Name_Space.naming} represents the abstract concept of
   a naming policy.
 
-  \item @{ML NameSpace.default_naming} is the default naming policy.
+  \item @{ML Name_Space.default_naming} is the default naming policy.
   In a theory context, this is usually augmented by a path prefix
   consisting of the theory name.
 
-  \item @{ML NameSpace.add_path}~@{text "path naming"} augments the
+  \item @{ML Name_Space.add_path}~@{text "path naming"} augments the
   naming policy by extending its path component.
 
-  \item @{ML NameSpace.full_name}~@{text "naming binding"} turns a
+  \item @{ML Name_Space.full_name}~@{text "naming binding"} turns a
   name binding (usually a basic name) into the fully qualified
   internal name, according to the given naming policy.
 
-  \item @{ML_type NameSpace.T} represents name spaces.
+  \item @{ML_type Name_Space.T} represents name spaces.
 
-  \item @{ML NameSpace.empty} and @{ML NameSpace.merge}~@{text
+  \item @{ML Name_Space.empty}~@{text "kind"} and @{ML Name_Space.merge}~@{text
   "(space\<^isub>1, space\<^isub>2)"} are the canonical operations for
   maintaining name spaces according to theory data management
-  (\secref{sec:context-data}).
+  (\secref{sec:context-data}); @{text "kind"} is a formal comment
+  to characterize the purpose of a name space.
 
-  \item @{ML NameSpace.declare}~@{text "naming bindings space"} enters a
-  name binding as fully qualified internal name into the name space,
-  with external accesses determined by the naming policy.
+  \item @{ML Name_Space.declare}~@{text "strict naming bindings
+  space"} enters a name binding as fully qualified internal name into
+  the name space, with external accesses determined by the naming
+  policy.
 
-  \item @{ML NameSpace.intern}~@{text "space name"} internalizes a
+  \item @{ML Name_Space.intern}~@{text "space name"} internalizes a
   (partially qualified) external name.
 
   This operation is mostly for parsing!  Note that fully qualified
   names stemming from declarations are produced via @{ML
-  "NameSpace.full_name"} and @{ML "NameSpace.declare"}
+  "Name_Space.full_name"} and @{ML "Name_Space.declare"}
   (or their derivatives for @{ML_type theory} and
   @{ML_type Proof.context}).
 
-  \item @{ML NameSpace.extern}~@{text "space name"} externalizes a
+  \item @{ML Name_Space.extern}~@{text "space name"} externalizes a
   (fully qualified) internal name.
 
   This operation is mostly for printing!  User code should not rely on
--- a/doc-src/IsarImplementation/Thy/document/Logic.tex	Fri Oct 23 13:23:18 2009 +0200
+++ b/doc-src/IsarImplementation/Thy/document/Logic.tex	Mon Oct 26 09:03:57 2009 +0100
@@ -325,9 +325,9 @@
   \indexdef{}{ML}{fastype\_of}\verb|fastype_of: term -> typ| \\
   \indexdef{}{ML}{lambda}\verb|lambda: term -> term -> term| \\
   \indexdef{}{ML}{betapply}\verb|betapply: term * term -> term| \\
-  \indexdef{}{ML}{Sign.declare\_const}\verb|Sign.declare_const: Properties.T -> (binding * typ) * mixfix ->|\isasep\isanewline%
+  \indexdef{}{ML}{Sign.declare\_const}\verb|Sign.declare_const: (binding * typ) * mixfix ->|\isasep\isanewline%
 \verb|  theory -> term * theory| \\
-  \indexdef{}{ML}{Sign.add\_abbrev}\verb|Sign.add_abbrev: string -> Properties.T -> binding * term ->|\isasep\isanewline%
+  \indexdef{}{ML}{Sign.add\_abbrev}\verb|Sign.add_abbrev: string -> binding * term ->|\isasep\isanewline%
 \verb|  theory -> (term * term) * theory| \\
   \indexdef{}{ML}{Sign.const\_typargs}\verb|Sign.const_typargs: theory -> string * typ -> typ list| \\
   \indexdef{}{ML}{Sign.const\_instance}\verb|Sign.const_instance: theory -> string * typ list -> typ| \\
@@ -365,11 +365,11 @@
   \item \verb|betapply|~\isa{{\isacharparenleft}t{\isacharcomma}\ u{\isacharparenright}} produces an application \isa{t\ u}, with topmost \isa{{\isasymbeta}}-conversion if \isa{t} is an
   abstraction.
 
-  \item \verb|Sign.declare_const|~\isa{properties\ {\isacharparenleft}{\isacharparenleft}c{\isacharcomma}\ {\isasymsigma}{\isacharparenright}{\isacharcomma}\ mx{\isacharparenright}}
+  \item \verb|Sign.declare_const|~\isa{{\isacharparenleft}{\isacharparenleft}c{\isacharcomma}\ {\isasymsigma}{\isacharparenright}{\isacharcomma}\ mx{\isacharparenright}}
   declares a new constant \isa{c\ {\isacharcolon}{\isacharcolon}\ {\isasymsigma}} with optional mixfix
   syntax.
 
-  \item \verb|Sign.add_abbrev|~\isa{print{\isacharunderscore}mode\ properties\ {\isacharparenleft}c{\isacharcomma}\ t{\isacharparenright}}
+  \item \verb|Sign.add_abbrev|~\isa{print{\isacharunderscore}mode\ {\isacharparenleft}c{\isacharcomma}\ t{\isacharparenright}}
   introduces a new term abbreviation \isa{c\ {\isasymequiv}\ t}.
 
   \item \verb|Sign.const_typargs|~\isa{thy\ {\isacharparenleft}c{\isacharcomma}\ {\isasymtau}{\isacharparenright}} and \verb|Sign.const_instance|~\isa{thy\ {\isacharparenleft}c{\isacharcomma}\ {\isacharbrackleft}{\isasymtau}\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ {\isasymtau}\isactrlisub n{\isacharbrackright}{\isacharparenright}}
--- a/doc-src/IsarImplementation/Thy/document/ML.tex	Fri Oct 23 13:23:18 2009 +0200
+++ b/doc-src/IsarImplementation/Thy/document/ML.tex	Mon Oct 26 09:03:57 2009 +0100
@@ -242,14 +242,14 @@
   view being presented to the user.
 
   Occasionally, such global process flags are treated like implicit
-  arguments to certain operations, by using the \verb|setmp| combinator
+  arguments to certain operations, by using the \verb|setmp_CRITICAL| combinator
   for safe temporary assignment.  Its traditional purpose was to
   ensure proper recovery of the original value when exceptions are
   raised in the body, now the functionality is extended to enter the
   \emph{critical section} (with its usual potential of degrading
   parallelism).
 
-  Note that recovery of plain value passing semantics via \verb|setmp|~\isa{ref\ value} assumes that this \isa{ref} is
+  Note that recovery of plain value passing semantics via \verb|setmp_CRITICAL|~\isa{ref\ value} assumes that this \isa{ref} is
   exclusively manipulated within the critical section.  In particular,
   any persistent global assignment of \isa{ref\ {\isacharcolon}{\isacharequal}\ value} needs to
   be marked critical as well, to prevent intruding another threads
@@ -277,7 +277,7 @@
 \begin{mldecls}
   \indexdef{}{ML}{NAMED\_CRITICAL}\verb|NAMED_CRITICAL: string -> (unit -> 'a) -> 'a| \\
   \indexdef{}{ML}{CRITICAL}\verb|CRITICAL: (unit -> 'a) -> 'a| \\
-  \indexdef{}{ML}{setmp}\verb|setmp: 'a Unsynchronized.ref -> 'a -> ('b -> 'c) -> 'b -> 'c| \\
+  \indexdef{}{ML}{setmp\_CRITICAL}\verb|setmp_CRITICAL: 'a Unsynchronized.ref -> 'a -> ('b -> 'c) -> 'b -> 'c| \\
   \end{mldecls}
 
   \begin{description}
@@ -291,7 +291,7 @@
   \item \verb|CRITICAL| is the same as \verb|NAMED_CRITICAL| with empty
   name argument.
 
-  \item \verb|setmp|~\isa{ref\ value\ f\ x} evaluates \isa{f\ x}
+  \item \verb|setmp_CRITICAL|~\isa{ref\ value\ f\ x} evaluates \isa{f\ x}
   while staying within the critical section and having \isa{ref\ {\isacharcolon}{\isacharequal}\ value} assigned temporarily.  This recovers a value-passing
   semantics involving global references, regardless of exceptions or
   concurrency.
@@ -366,7 +366,7 @@
   a theory by constant declararion and primitive definitions:
 
   \smallskip\begin{mldecls}
-  \verb|Sign.declare_const: Properties.T -> (binding * typ) * mixfix|\isasep\isanewline%
+  \verb|Sign.declare_const: (binding * typ) * mixfix|\isasep\isanewline%
 \verb|  -> theory -> term * theory| \\
   \verb|Thm.add_def: bool -> bool -> binding * term -> theory -> thm * theory|
   \end{mldecls}
@@ -378,7 +378,7 @@
   \smallskip\begin{mldecls}
   \verb|(fn (t, thy) => Thm.add_def false false|\isasep\isanewline%
 \verb|  (Binding.name "bar_def", Logic.mk_equals (t, @{term "%x. x"})) thy)|\isasep\isanewline%
-\verb|    (Sign.declare_const []|\isasep\isanewline%
+\verb|    (Sign.declare_const|\isasep\isanewline%
 \verb|      ((Binding.name "bar", @{typ "foo => foo"}), NoSyn) thy)|
   \end{mldecls}
 
@@ -394,7 +394,7 @@
 
   \smallskip\begin{mldecls}
 \verb|thy|\isasep\isanewline%
-\verb||\verb,|,\verb|> Sign.declare_const [] ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
+\verb||\verb,|,\verb|> Sign.declare_const ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
 \verb||\verb,|,\verb|> (fn (t, thy) => thy|\isasep\isanewline%
 \verb||\verb,|,\verb|> Thm.add_def false false|\isasep\isanewline%
 \verb|     (Binding.name "bar_def", Logic.mk_equals (t, @{term "%x. x"})))|
@@ -433,7 +433,7 @@
 
   \smallskip\begin{mldecls}
 \verb|thy|\isasep\isanewline%
-\verb||\verb,|,\verb|> Sign.declare_const [] ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
+\verb||\verb,|,\verb|> Sign.declare_const ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
 \verb||\verb,|,\verb|-> (fn t => Thm.add_def false false|\isasep\isanewline%
 \verb|      (Binding.name "bar_def", Logic.mk_equals (t, @{term "%x. x"})))|\isasep\isanewline%
 
@@ -443,7 +443,7 @@
 
   \smallskip\begin{mldecls}
 \verb|thy|\isasep\isanewline%
-\verb||\verb,|,\verb|> Sign.declare_const [] ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
+\verb||\verb,|,\verb|> Sign.declare_const ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
 \verb||\verb,|,\verb|>> (fn t => Logic.mk_equals (t, @{term "%x. x"}))|\isasep\isanewline%
 \verb||\verb,|,\verb|-> (fn def => Thm.add_def false false (Binding.name "bar_def", def))|\isasep\isanewline%
 
@@ -454,7 +454,7 @@
 
   \smallskip\begin{mldecls}
 \verb|thy|\isasep\isanewline%
-\verb||\verb,|,\verb|> Sign.declare_const [] ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
+\verb||\verb,|,\verb|> Sign.declare_const ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
 \verb||\verb,|,\verb||\verb,|,\verb|> Sign.add_path "foobar"|\isasep\isanewline%
 \verb||\verb,|,\verb|-> (fn t => Thm.add_def false false|\isasep\isanewline%
 \verb|      (Binding.name "bar_def", Logic.mk_equals (t, @{term "%x. x"})))|\isasep\isanewline%
@@ -466,8 +466,8 @@
 
   \smallskip\begin{mldecls}
 \verb|thy|\isasep\isanewline%
-\verb||\verb,|,\verb|> Sign.declare_const [] ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
-\verb||\verb,|,\verb||\verb,|,\verb|>> Sign.declare_const [] ((Binding.name "foobar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
+\verb||\verb,|,\verb|> Sign.declare_const ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
+\verb||\verb,|,\verb||\verb,|,\verb|>> Sign.declare_const ((Binding.name "foobar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
 \verb||\verb,|,\verb|-> (fn (t1, t2) => Thm.add_def false false|\isasep\isanewline%
 \verb|      (Binding.name "bar_def", Logic.mk_equals (t1, t2)))|\isasep\isanewline%
 
@@ -527,7 +527,7 @@
 \verb|  val consts = ["foo", "bar"];|\isasep\isanewline%
 \verb|in|\isasep\isanewline%
 \verb|  thy|\isasep\isanewline%
-\verb|  |\verb,|,\verb|> fold_map (fn const => Sign.declare_const []|\isasep\isanewline%
+\verb|  |\verb,|,\verb|> fold_map (fn const => Sign.declare_const|\isasep\isanewline%
 \verb|       ((Binding.name const, @{typ "foo => foo"}), NoSyn)) consts|\isasep\isanewline%
 \verb|  |\verb,|,\verb|>> map (fn t => Logic.mk_equals (t, @{term "%x. x"}))|\isasep\isanewline%
 \verb|  |\verb,|,\verb|-> (fn defs => fold_map (fn def =>|\isasep\isanewline%
@@ -596,11 +596,11 @@
   \smallskip\begin{mldecls}
 \verb|thy|\isasep\isanewline%
 \verb||\verb,|,\verb|> tap (fn _ => writeln "now adding constant")|\isasep\isanewline%
-\verb||\verb,|,\verb|> Sign.declare_const [] ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
+\verb||\verb,|,\verb|> Sign.declare_const ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
 \verb||\verb,|,\verb||\verb,|,\verb|>> `(fn thy => Sign.declared_const thy|\isasep\isanewline%
 \verb|         (Sign.full_name thy (Binding.name "foobar")))|\isasep\isanewline%
 \verb||\verb,|,\verb|-> (fn (t, b) => if b then I|\isasep\isanewline%
-\verb|       else Sign.declare_const []|\isasep\isanewline%
+\verb|       else Sign.declare_const|\isasep\isanewline%
 \verb|         ((Binding.name "foobar", @{typ "foo => foo"}), NoSyn) #> snd)|\isasep\isanewline%
 
   \end{mldecls}%
--- a/doc-src/IsarImplementation/Thy/document/Prelim.tex	Fri Oct 23 13:23:18 2009 +0200
+++ b/doc-src/IsarImplementation/Thy/document/Prelim.tex	Mon Oct 26 09:03:57 2009 +0100
@@ -798,19 +798,19 @@
   \indexdef{}{ML}{Long\_Name.explode}\verb|Long_Name.explode: string -> string list| \\
   \end{mldecls}
   \begin{mldecls}
-  \indexdef{}{ML type}{NameSpace.naming}\verb|type NameSpace.naming| \\
-  \indexdef{}{ML}{NameSpace.default\_naming}\verb|NameSpace.default_naming: NameSpace.naming| \\
-  \indexdef{}{ML}{NameSpace.add\_path}\verb|NameSpace.add_path: string -> NameSpace.naming -> NameSpace.naming| \\
-  \indexdef{}{ML}{NameSpace.full\_name}\verb|NameSpace.full_name: NameSpace.naming -> binding -> string| \\
+  \indexdef{}{ML type}{Name\_Space.naming}\verb|type Name_Space.naming| \\
+  \indexdef{}{ML}{Name\_Space.default\_naming}\verb|Name_Space.default_naming: Name_Space.naming| \\
+  \indexdef{}{ML}{Name\_Space.add\_path}\verb|Name_Space.add_path: string -> Name_Space.naming -> Name_Space.naming| \\
+  \indexdef{}{ML}{Name\_Space.full\_name}\verb|Name_Space.full_name: Name_Space.naming -> binding -> string| \\
   \end{mldecls}
   \begin{mldecls}
-  \indexdef{}{ML type}{NameSpace.T}\verb|type NameSpace.T| \\
-  \indexdef{}{ML}{NameSpace.empty}\verb|NameSpace.empty: NameSpace.T| \\
-  \indexdef{}{ML}{NameSpace.merge}\verb|NameSpace.merge: NameSpace.T * NameSpace.T -> NameSpace.T| \\
-  \indexdef{}{ML}{NameSpace.declare}\verb|NameSpace.declare: NameSpace.naming -> binding -> NameSpace.T ->|\isasep\isanewline%
-\verb|  string * NameSpace.T| \\
-  \indexdef{}{ML}{NameSpace.intern}\verb|NameSpace.intern: NameSpace.T -> string -> string| \\
-  \indexdef{}{ML}{NameSpace.extern}\verb|NameSpace.extern: NameSpace.T -> string -> string| \\
+  \indexdef{}{ML type}{Name\_Space.T}\verb|type Name_Space.T| \\
+  \indexdef{}{ML}{Name\_Space.empty}\verb|Name_Space.empty: string -> Name_Space.T| \\
+  \indexdef{}{ML}{Name\_Space.merge}\verb|Name_Space.merge: Name_Space.T * Name_Space.T -> Name_Space.T| \\
+  \indexdef{}{ML}{Name\_Space.declare}\verb|Name_Space.declare: bool -> Name_Space.naming -> binding -> Name_Space.T ->|\isasep\isanewline%
+\verb|  string * Name_Space.T| \\
+  \indexdef{}{ML}{Name\_Space.intern}\verb|Name_Space.intern: Name_Space.T -> string -> string| \\
+  \indexdef{}{ML}{Name\_Space.extern}\verb|Name_Space.extern: Name_Space.T -> string -> string| \\
   \end{mldecls}
 
   \begin{description}
@@ -827,39 +827,40 @@
   \item \verb|Long_Name.implode|~\isa{names} and \verb|Long_Name.explode|~\isa{name} convert between the packed string
   representation and the explicit list form of qualified names.
 
-  \item \verb|NameSpace.naming| represents the abstract concept of
+  \item \verb|Name_Space.naming| represents the abstract concept of
   a naming policy.
 
-  \item \verb|NameSpace.default_naming| is the default naming policy.
+  \item \verb|Name_Space.default_naming| is the default naming policy.
   In a theory context, this is usually augmented by a path prefix
   consisting of the theory name.
 
-  \item \verb|NameSpace.add_path|~\isa{path\ naming} augments the
+  \item \verb|Name_Space.add_path|~\isa{path\ naming} augments the
   naming policy by extending its path component.
 
-  \item \verb|NameSpace.full_name|~\isa{naming\ binding} turns a
+  \item \verb|Name_Space.full_name|~\isa{naming\ binding} turns a
   name binding (usually a basic name) into the fully qualified
   internal name, according to the given naming policy.
 
-  \item \verb|NameSpace.T| represents name spaces.
-
-  \item \verb|NameSpace.empty| and \verb|NameSpace.merge|~\isa{{\isacharparenleft}space\isactrlisub {\isadigit{1}}{\isacharcomma}\ space\isactrlisub {\isadigit{2}}{\isacharparenright}} are the canonical operations for
-  maintaining name spaces according to theory data management
-  (\secref{sec:context-data}).
+  \item \verb|Name_Space.T| represents name spaces.
 
-  \item \verb|NameSpace.declare|~\isa{naming\ bindings\ space} enters a
-  name binding as fully qualified internal name into the name space,
-  with external accesses determined by the naming policy.
+  \item \verb|Name_Space.empty|~\isa{kind} and \verb|Name_Space.merge|~\isa{{\isacharparenleft}space\isactrlisub {\isadigit{1}}{\isacharcomma}\ space\isactrlisub {\isadigit{2}}{\isacharparenright}} are the canonical operations for
+  maintaining name spaces according to theory data management
+  (\secref{sec:context-data}); \isa{kind} is a formal comment
+  to characterize the purpose of a name space.
 
-  \item \verb|NameSpace.intern|~\isa{space\ name} internalizes a
+  \item \verb|Name_Space.declare|~\isa{strict\ naming\ bindings\ space} enters a name binding as fully qualified internal name into
+  the name space, with external accesses determined by the naming
+  policy.
+
+  \item \verb|Name_Space.intern|~\isa{space\ name} internalizes a
   (partially qualified) external name.
 
   This operation is mostly for parsing!  Note that fully qualified
-  names stemming from declarations are produced via \verb|NameSpace.full_name| and \verb|NameSpace.declare|
+  names stemming from declarations are produced via \verb|Name_Space.full_name| and \verb|Name_Space.declare|
   (or their derivatives for \verb|theory| and
   \verb|Proof.context|).
 
-  \item \verb|NameSpace.extern|~\isa{space\ name} externalizes a
+  \item \verb|Name_Space.extern|~\isa{space\ name} externalizes a
   (fully qualified) internal name.
 
   This operation is mostly for printing!  User code should not rely on
--- a/src/HOL/Decision_Procs/Decision_Procs.thy	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Decision_Procs/Decision_Procs.thy	Mon Oct 26 09:03:57 2009 +0100
@@ -1,7 +1,7 @@
 header {* Various decision procedures. typically involving reflection *}
 
 theory Decision_Procs
-imports Cooper Ferrack MIR Approximation Dense_Linear_Order "ex/Approximation_Ex" "ex/Dense_Linear_Order_Ex"
+imports Cooper Ferrack MIR Approximation Dense_Linear_Order "ex/Approximation_Ex" "ex/Dense_Linear_Order_Ex" Parametric_Ferrante_Rackoff
 begin
 
 end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Mon Oct 26 09:03:57 2009 +0100
@@ -0,0 +1,3227 @@
+(*  Title:      HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy
+    Author:     Amine Chaieb
+*)
+
+header{* A formalization of Ferrante and Rackoff's procedure with polynomial parameters, see Paper in CALCULEMUS 2008 *}
+
+theory Parametric_Ferrante_Rackoff
+imports Reflected_Multivariate_Polynomial 
+  "~~/src/HOL/Decision_Procs/Dense_Linear_Order"
+begin
+
+
+subsection {* Terms *}
+
+datatype tm = CP poly | Bound nat | Add tm tm | Mul poly tm 
+  | Neg tm | Sub tm tm | CNP nat poly tm
+  (* A size for poly to make inductive proofs simpler*)
+
+consts tmsize :: "tm \<Rightarrow> nat"
+primrec 
+  "tmsize (CP c) = polysize c"
+  "tmsize (Bound n) = 1"
+  "tmsize (Neg a) = 1 + tmsize a"
+  "tmsize (Add a b) = 1 + tmsize a + tmsize b"
+  "tmsize (Sub a b) = 3 + tmsize a + tmsize b"
+  "tmsize (Mul c a) = 1 + polysize c + tmsize a"
+  "tmsize (CNP n c a) = 3 + polysize c + tmsize a "
+
+  (* Semantics of terms tm *)
+consts Itm :: "'a::{ring_char_0,division_by_zero,field} list \<Rightarrow> 'a list \<Rightarrow> tm \<Rightarrow> 'a"
+primrec
+  "Itm vs bs (CP c) = (Ipoly vs c)"
+  "Itm vs bs (Bound n) = bs!n"
+  "Itm vs bs (Neg a) = -(Itm vs bs a)"
+  "Itm vs bs (Add a b) = Itm vs bs a + Itm vs bs b"
+  "Itm vs bs (Sub a b) = Itm vs bs a - Itm vs bs b"
+  "Itm vs bs (Mul c a) = (Ipoly vs c) * Itm vs bs a"
+  "Itm vs bs (CNP n c t) = (Ipoly vs c)*(bs!n) + Itm vs bs t"	
+
+
+fun allpolys:: "(poly \<Rightarrow> bool) \<Rightarrow> tm \<Rightarrow> bool"  where
+  "allpolys P (CP c) = P c"
+| "allpolys P (CNP n c p) = (P c \<and> allpolys P p)"
+| "allpolys P (Mul c p) = (P c \<and> allpolys P p)"
+| "allpolys P (Neg p) = allpolys P p"
+| "allpolys P (Add p q) = (allpolys P p \<and> allpolys P q)"
+| "allpolys P (Sub p q) = (allpolys P p \<and> allpolys P q)"
+| "allpolys P p = True"
+
+consts 
+  tmboundslt:: "nat \<Rightarrow> tm \<Rightarrow> bool"
+  tmbound0:: "tm \<Rightarrow> bool" (* a tm is INDEPENDENT of Bound 0 *)
+  tmbound:: "nat \<Rightarrow> tm \<Rightarrow> bool" (* a tm is INDEPENDENT of Bound n *)
+  incrtm0:: "tm \<Rightarrow> tm"
+  incrtm:: "nat \<Rightarrow> tm \<Rightarrow> tm"
+  decrtm0:: "tm \<Rightarrow> tm" 
+  decrtm:: "nat \<Rightarrow> tm \<Rightarrow> tm" 
+primrec
+  "tmboundslt n (CP c) = True"
+  "tmboundslt n (Bound m) = (m < n)"
+  "tmboundslt n (CNP m c a) = (m < n \<and> tmboundslt n a)"
+  "tmboundslt n (Neg a) = tmboundslt n a"
+  "tmboundslt n (Add a b) = (tmboundslt n a \<and> tmboundslt n b)"
+  "tmboundslt n (Sub a b) = (tmboundslt n a \<and> tmboundslt n b)" 
+  "tmboundslt n (Mul i a) = tmboundslt n a"
+primrec
+  "tmbound0 (CP c) = True"
+  "tmbound0 (Bound n) = (n>0)"
+  "tmbound0 (CNP n c a) = (n\<noteq>0 \<and> tmbound0 a)"
+  "tmbound0 (Neg a) = tmbound0 a"
+  "tmbound0 (Add a b) = (tmbound0 a \<and> tmbound0 b)"
+  "tmbound0 (Sub a b) = (tmbound0 a \<and> tmbound0 b)" 
+  "tmbound0 (Mul i a) = tmbound0 a"
+lemma tmbound0_I:
+  assumes nb: "tmbound0 a"
+  shows "Itm vs (b#bs) a = Itm vs (b'#bs) a"
+using nb
+by (induct a rule: tmbound0.induct,auto simp add: nth_pos2)
+
+primrec
+  "tmbound n (CP c) = True"
+  "tmbound n (Bound m) = (n \<noteq> m)"
+  "tmbound n (CNP m c a) = (n\<noteq>m \<and> tmbound n a)"
+  "tmbound n (Neg a) = tmbound n a"
+  "tmbound n (Add a b) = (tmbound n a \<and> tmbound n b)"
+  "tmbound n (Sub a b) = (tmbound n a \<and> tmbound n b)" 
+  "tmbound n (Mul i a) = tmbound n a"
+lemma tmbound0_tmbound_iff: "tmbound 0 t = tmbound0 t" by (induct t, auto)
+
+lemma tmbound_I: 
+  assumes bnd: "tmboundslt (length bs) t" and nb: "tmbound n t" and le: "n \<le> length bs"
+  shows "Itm vs (bs[n:=x]) t = Itm vs bs t"
+  using nb le bnd
+  by (induct t rule: tmbound.induct , auto)
+
+recdef decrtm0 "measure size"
+  "decrtm0 (Bound n) = Bound (n - 1)"
+  "decrtm0 (Neg a) = Neg (decrtm0 a)"
+  "decrtm0 (Add a b) = Add (decrtm0 a) (decrtm0 b)"
+  "decrtm0 (Sub a b) = Sub (decrtm0 a) (decrtm0 b)"
+  "decrtm0 (Mul c a) = Mul c (decrtm0 a)"
+  "decrtm0 (CNP n c a) = CNP (n - 1) c (decrtm0 a)"
+  "decrtm0 a = a"
+recdef incrtm0 "measure size"
+  "incrtm0 (Bound n) = Bound (n + 1)"
+  "incrtm0 (Neg a) = Neg (incrtm0 a)"
+  "incrtm0 (Add a b) = Add (incrtm0 a) (incrtm0 b)"
+  "incrtm0 (Sub a b) = Sub (incrtm0 a) (incrtm0 b)"
+  "incrtm0 (Mul c a) = Mul c (incrtm0 a)"
+  "incrtm0 (CNP n c a) = CNP (n + 1) c (incrtm0 a)"
+  "incrtm0 a = a"
+lemma decrtm0: assumes nb: "tmbound0 t"
+  shows "Itm vs (x#bs) t = Itm vs bs (decrtm0 t)"
+  using nb by (induct t rule: decrtm0.induct, simp_all add: nth_pos2)
+lemma incrtm0: "Itm vs (x#bs) (incrtm0 t) = Itm vs bs t"
+  by (induct t rule: decrtm0.induct, simp_all add: nth_pos2)
+
+primrec
+  "decrtm m (CP c) = (CP c)"
+  "decrtm m (Bound n) = (if n < m then Bound n else Bound (n - 1))"
+  "decrtm m (Neg a) = Neg (decrtm m a)"
+  "decrtm m (Add a b) = Add (decrtm m a) (decrtm m b)"
+  "decrtm m (Sub a b) = Sub (decrtm m a) (decrtm m b)"
+  "decrtm m (Mul c a) = Mul c (decrtm m a)"
+  "decrtm m (CNP n c a) = (if n < m then CNP n c (decrtm m a) else CNP (n - 1) c (decrtm m a))"
+
+consts removen:: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
+primrec
+  "removen n [] = []"
+  "removen n (x#xs) = (if n=0 then xs else (x#(removen (n - 1) xs)))"
+
+lemma removen_same: "n \<ge> length xs \<Longrightarrow> removen n xs = xs"
+  by (induct xs arbitrary: n, auto)
+
+lemma nth_length_exceeds: "n \<ge> length xs \<Longrightarrow> xs!n = []!(n - length xs)"
+  by (induct xs arbitrary: n, auto)
+
+lemma removen_length: "length (removen n xs) = (if n \<ge> length xs then length xs else length xs - 1)"
+  by (induct xs arbitrary: n, auto)
+lemma removen_nth: "(removen n xs)!m = (if n \<ge> length xs then xs!m 
+  else if m < n then xs!m else if m \<le> length xs then xs!(Suc m) else []!(m - (length xs - 1)))"
+proof(induct xs arbitrary: n m)
+  case Nil thus ?case by simp
+next
+  case (Cons x xs n m)
+  {assume nxs: "n \<ge> length (x#xs)" hence ?case using removen_same[OF nxs] by simp}
+  moreover
+  {assume nxs: "\<not> (n \<ge> length (x#xs))" 
+    {assume mln: "m < n" hence ?case using prems by (cases m, auto)}
+    moreover
+    {assume mln: "\<not> (m < n)" 
+      
+      {assume mxs: "m \<le> length (x#xs)" hence ?case using prems by (cases m, auto)}
+      moreover
+      {assume mxs: "\<not> (m \<le> length (x#xs))" 
+	have th: "length (removen n (x#xs)) = length xs" 
+	  using removen_length[where n="n" and xs="x#xs"] nxs by simp
+	with mxs have mxs':"m \<ge> length (removen n (x#xs))" by auto
+	hence "(removen n (x#xs))!m = [] ! (m - length xs)" 
+	  using th nth_length_exceeds[OF mxs'] by auto
+	hence th: "(removen n (x#xs))!m = [] ! (m - (length (x#xs) - 1))" 
+	  by auto
+	hence ?case using nxs mln mxs by auto }
+      ultimately have ?case by blast
+    }
+    ultimately have ?case by blast
+    
+  }      ultimately show ?case by blast
+qed
+
+lemma decrtm: assumes bnd: "tmboundslt (length bs) t" and nb: "tmbound m t" 
+  and nle: "m \<le> length bs" 
+  shows "Itm vs (removen m bs) (decrtm m t) = Itm vs bs t"
+  using bnd nb nle
+  by (induct t rule: decrtm.induct, auto simp add: removen_nth)
+
+consts tmsubst0:: "tm \<Rightarrow> tm \<Rightarrow> tm"
+primrec
+  "tmsubst0 t (CP c) = CP c"
+  "tmsubst0 t (Bound n) = (if n=0 then t else Bound n)"
+  "tmsubst0 t (CNP n c a) = (if n=0 then Add (Mul c t) (tmsubst0 t a) else CNP n c (tmsubst0 t a))"
+  "tmsubst0 t (Neg a) = Neg (tmsubst0 t a)"
+  "tmsubst0 t (Add a b) = Add (tmsubst0 t a) (tmsubst0 t b)"
+  "tmsubst0 t (Sub a b) = Sub (tmsubst0 t a) (tmsubst0 t b)" 
+  "tmsubst0 t (Mul i a) = Mul i (tmsubst0 t a)"
+lemma tmsubst0:
+  shows "Itm vs (x#bs) (tmsubst0 t a) = Itm vs ((Itm vs (x#bs) t)#bs) a"
+by (induct a rule: tmsubst0.induct,auto simp add: nth_pos2)
+
+lemma tmsubst0_nb: "tmbound0 t \<Longrightarrow> tmbound0 (tmsubst0 t a)"
+by (induct a rule: tmsubst0.induct,auto simp add: nth_pos2)
+
+consts tmsubst:: "nat \<Rightarrow> tm \<Rightarrow> tm \<Rightarrow> tm" 
+
+primrec
+  "tmsubst n t (CP c) = CP c"
+  "tmsubst n t (Bound m) = (if n=m then t else Bound m)"
+  "tmsubst n t (CNP m c a) = (if n=m then Add (Mul c t) (tmsubst n t a) 
+             else CNP m c (tmsubst n t a))"
+  "tmsubst n t (Neg a) = Neg (tmsubst n t a)"
+  "tmsubst n t (Add a b) = Add (tmsubst n t a) (tmsubst n t b)"
+  "tmsubst n t (Sub a b) = Sub (tmsubst n t a) (tmsubst n t b)" 
+  "tmsubst n t (Mul i a) = Mul i (tmsubst n t a)"
+
+lemma tmsubst: assumes nb: "tmboundslt (length bs) a" and nlt: "n \<le> length bs"
+  shows "Itm vs bs (tmsubst n t a) = Itm vs (bs[n:= Itm vs bs t]) a"
+using nb nlt
+by (induct a rule: tmsubst0.induct,auto simp add: nth_pos2)
+
+lemma tmsubst_nb0: assumes tnb: "tmbound0 t"
+shows "tmbound0 (tmsubst 0 t a)"
+using tnb
+by (induct a rule: tmsubst.induct, auto)
+
+lemma tmsubst_nb: assumes tnb: "tmbound m t"
+shows "tmbound m (tmsubst m t a)"
+using tnb
+by (induct a rule: tmsubst.induct, auto)
+lemma incrtm0_tmbound: "tmbound n t \<Longrightarrow> tmbound (Suc n) (incrtm0 t)"
+  by (induct t, auto)
+  (* Simplification *)
+
+consts
+  simptm:: "tm \<Rightarrow> tm"
+  tmadd:: "tm \<times> tm \<Rightarrow> tm"
+  tmmul:: "tm \<Rightarrow> poly \<Rightarrow> tm"
+recdef tmadd "measure (\<lambda> (t,s). size t + size s)"
+  "tmadd (CNP n1 c1 r1,CNP n2 c2 r2) =
+  (if n1=n2 then 
+  (let c = c1 +\<^sub>p c2
+  in if c = 0\<^sub>p then tmadd(r1,r2) else CNP n1 c (tmadd (r1,r2)))
+  else if n1 \<le> n2 then (CNP n1 c1 (tmadd (r1,CNP n2 c2 r2))) 
+  else (CNP n2 c2 (tmadd (CNP n1 c1 r1,r2))))"
+  "tmadd (CNP n1 c1 r1,t) = CNP n1 c1 (tmadd (r1, t))"  
+  "tmadd (t,CNP n2 c2 r2) = CNP n2 c2 (tmadd (t,r2))" 
+  "tmadd (CP b1, CP b2) = CP (b1 +\<^sub>p b2)"
+  "tmadd (a,b) = Add a b"
+
+lemma tmadd[simp]: "Itm vs bs (tmadd (t,s)) = Itm vs bs (Add t s)"
+apply (induct t s rule: tmadd.induct, simp_all add: Let_def)
+apply (case_tac "c1 +\<^sub>p c2 = 0\<^sub>p",case_tac "n1 \<le> n2", simp_all)
+apply (case_tac "n1 = n2", simp_all add: ring_simps)
+apply (simp only: right_distrib[symmetric]) 
+by (auto simp del: polyadd simp add: polyadd[symmetric])
+
+lemma tmadd_nb0[simp]: "\<lbrakk> tmbound0 t ; tmbound0 s\<rbrakk> \<Longrightarrow> tmbound0 (tmadd (t,s))"
+by (induct t s rule: tmadd.induct, auto simp add: Let_def)
+
+lemma tmadd_nb[simp]: "\<lbrakk> tmbound n t ; tmbound n s\<rbrakk> \<Longrightarrow> tmbound n (tmadd (t,s))"
+by (induct t s rule: tmadd.induct, auto simp add: Let_def)
+lemma tmadd_blt[simp]: "\<lbrakk>tmboundslt n t ; tmboundslt n s\<rbrakk> \<Longrightarrow> tmboundslt n (tmadd (t,s))"
+by (induct t s rule: tmadd.induct, auto simp add: Let_def)
+
+lemma tmadd_allpolys_npoly[simp]: "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly s \<Longrightarrow> allpolys isnpoly (tmadd(t,s))" by (induct t s rule: tmadd.induct, simp_all add: Let_def polyadd_norm)
+
+recdef tmmul "measure size"
+  "tmmul (CP j) = (\<lambda> i. CP (i *\<^sub>p j))"
+  "tmmul (CNP n c a) = (\<lambda> i. CNP n (i *\<^sub>p c) (tmmul a i))"
+  "tmmul t = (\<lambda> i. Mul i t)"
+
+lemma tmmul[simp]: "Itm vs bs (tmmul t i) = Itm vs bs (Mul i t)"
+by (induct t arbitrary: i rule: tmmul.induct, simp_all add: ring_simps)
+
+lemma tmmul_nb0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 (tmmul t i)"
+by (induct t arbitrary: i rule: tmmul.induct, auto )
+
+lemma tmmul_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n (tmmul t i)"
+by (induct t arbitrary: n rule: tmmul.induct, auto )
+lemma tmmul_blt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n (tmmul t i)"
+by (induct t arbitrary: i rule: tmmul.induct, auto simp add: Let_def)
+
+lemma tmmul_allpolys_npoly[simp]: 
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+  shows "allpolys isnpoly t \<Longrightarrow> isnpoly c \<Longrightarrow> allpolys isnpoly (tmmul t c)" by (induct t rule: tmmul.induct, simp_all add: Let_def polymul_norm)
+
+constdefs tmneg :: "tm \<Rightarrow> tm"
+  "tmneg t \<equiv> tmmul t (C (- 1,1))"
+
+constdefs tmsub :: "tm \<Rightarrow> tm \<Rightarrow> tm"
+  "tmsub s t \<equiv> (if s = t then CP 0\<^sub>p else tmadd (s,tmneg t))"
+
+lemma tmneg[simp]: "Itm vs bs (tmneg t) = Itm vs bs (Neg t)"
+using tmneg_def[of t] 
+apply simp
+apply (subst number_of_Min)
+apply (simp only: of_int_minus)
+apply simp
+done
+
+lemma tmneg_nb0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 (tmneg t)"
+using tmneg_def by simp
+
+lemma tmneg_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n (tmneg t)"
+using tmneg_def by simp
+lemma tmneg_blt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n (tmneg t)"
+using tmneg_def by simp
+lemma [simp]: "isnpoly (C (-1,1))" unfolding isnpoly_def by simp
+lemma tmneg_allpolys_npoly[simp]: 
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+  shows "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly (tmneg t)" 
+  unfolding tmneg_def by auto
+
+lemma tmsub[simp]: "Itm vs bs (tmsub a b) = Itm vs bs (Sub a b)"
+using tmsub_def by simp
+
+lemma tmsub_nb0[simp]: "\<lbrakk> tmbound0 t ; tmbound0 s\<rbrakk> \<Longrightarrow> tmbound0 (tmsub t s)"
+using tmsub_def by simp
+lemma tmsub_nb[simp]: "\<lbrakk> tmbound n t ; tmbound n s\<rbrakk> \<Longrightarrow> tmbound n (tmsub t s)"
+using tmsub_def by simp
+lemma tmsub_blt[simp]: "\<lbrakk>tmboundslt n t ; tmboundslt n s\<rbrakk> \<Longrightarrow> tmboundslt n (tmsub t s )"
+using tmsub_def by simp
+lemma tmsub_allpolys_npoly[simp]: 
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+  shows "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly s \<Longrightarrow> allpolys isnpoly (tmsub t s)" 
+  unfolding tmsub_def by (simp add: isnpoly_def)
+
+recdef simptm "measure size"
+  "simptm (CP j) = CP (polynate j)"
+  "simptm (Bound n) = CNP n 1\<^sub>p (CP 0\<^sub>p)"
+  "simptm (Neg t) = tmneg (simptm t)"
+  "simptm (Add t s) = tmadd (simptm t,simptm s)"
+  "simptm (Sub t s) = tmsub (simptm t) (simptm s)"
+  "simptm (Mul i t) = (let i' = polynate i in if i' = 0\<^sub>p then CP 0\<^sub>p else tmmul (simptm t) i')"
+  "simptm (CNP n c t) = (let c' = polynate c in if c' = 0\<^sub>p then simptm t else tmadd (CNP n c' (CP 0\<^sub>p ), simptm t))"
+
+lemma polynate_stupid: 
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+  shows "polynate t = 0\<^sub>p \<Longrightarrow> Ipoly bs t = (0::'a::{ring_char_0,division_by_zero, field})" 
+apply (subst polynate[symmetric])
+apply simp
+done
+
+lemma simptm_ci[simp]: "Itm vs bs (simptm t) = Itm vs bs t"
+by (induct t rule: simptm.induct, auto simp add: tmneg tmadd tmsub tmmul Let_def polynate_stupid) 
+
+lemma simptm_tmbound0[simp]: 
+  "tmbound0 t \<Longrightarrow> tmbound0 (simptm t)"
+by (induct t rule: simptm.induct, auto simp add: Let_def)
+
+lemma simptm_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n (simptm t)"
+by (induct t rule: simptm.induct, auto simp add: Let_def)
+lemma simptm_nlt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n (simptm t)"
+by (induct t rule: simptm.induct, auto simp add: Let_def)
+
+lemma [simp]: "isnpoly 0\<^sub>p" and [simp]: "isnpoly (C(1,1))" 
+  by (simp_all add: isnpoly_def)
+lemma simptm_allpolys_npoly[simp]: 
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+  shows "allpolys isnpoly (simptm p)"
+  by (induct p rule: simptm.induct, auto simp add: Let_def)
+
+consts split0 :: "tm \<Rightarrow> (poly \<times> tm)"
+recdef split0 "measure tmsize"
+  "split0 (Bound 0) = (1\<^sub>p, CP 0\<^sub>p)"
+  "split0 (CNP 0 c t) = (let (c',t') = split0 t in (c +\<^sub>p c',t'))"
+  "split0 (Neg t) = (let (c,t') = split0 t in (~\<^sub>p c,Neg t'))"
+  "split0 (CNP n c t) = (let (c',t') = split0 t in (c',CNP n c t'))"
+  "split0 (Add s t) = (let (c1,s') = split0 s ; (c2,t') = split0 t in (c1 +\<^sub>p c2, Add s' t'))"
+  "split0 (Sub s t) = (let (c1,s') = split0 s ; (c2,t') = split0 t in (c1 -\<^sub>p c2, Sub s' t'))"
+  "split0 (Mul c t) = (let (c',t') = split0 t in (c *\<^sub>p c', Mul c t'))"
+  "split0 t = (0\<^sub>p, t)"
+
+lemma split0_stupid[simp]: "\<exists>x y. (x,y) = split0 p"
+  apply (rule exI[where x="fst (split0 p)"])
+  apply (rule exI[where x="snd (split0 p)"])
+  by simp
+
+lemma split0:
+  "tmbound 0 (snd (split0 t)) \<and> (Itm vs bs (CNP 0 (fst (split0 t)) (snd (split0 t))) = Itm vs bs t)"
+  apply (induct t rule: split0.induct)
+  apply simp
+  apply (simp add: Let_def split_def ring_simps)
+  apply (simp add: Let_def split_def ring_simps)
+  apply (simp add: Let_def split_def ring_simps)
+  apply (simp add: Let_def split_def ring_simps)
+  apply (simp add: Let_def split_def ring_simps)
+  apply (simp add: Let_def split_def mult_assoc right_distrib[symmetric])
+  apply (simp add: Let_def split_def ring_simps)
+  apply (simp add: Let_def split_def ring_simps)
+  done
+
+lemma split0_ci: "split0 t = (c',t') \<Longrightarrow> Itm vs bs t = Itm vs bs (CNP 0 c' t')"
+proof-
+  fix c' t'
+  assume "split0 t = (c', t')" hence "c' = fst (split0 t)" and "t' = snd (split0 t)" by auto
+  with split0[where t="t" and bs="bs"] show "Itm vs bs t = Itm vs bs (CNP 0 c' t')" by simp
+qed
+
+lemma split0_nb0: 
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+  shows "split0 t = (c',t') \<Longrightarrow>  tmbound 0 t'"
+proof-
+  fix c' t'
+  assume "split0 t = (c', t')" hence "c' = fst (split0 t)" and "t' = snd (split0 t)" by auto
+  with conjunct1[OF split0[where t="t"]] show "tmbound 0 t'" by simp
+qed
+
+lemma split0_nb0'[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+  shows "tmbound0 (snd (split0 t))"
+  using split0_nb0[of t "fst (split0 t)" "snd (split0 t)"] by (simp add: tmbound0_tmbound_iff)
+
+
+lemma split0_nb: assumes nb:"tmbound n t" shows "tmbound n (snd (split0 t))"
+  using nb by (induct t rule: split0.induct, auto simp add: Let_def split_def split0_stupid)
+
+lemma split0_blt: assumes nb:"tmboundslt n t" shows "tmboundslt n (snd (split0 t))"
+  using nb by (induct t rule: split0.induct, auto simp add: Let_def split_def split0_stupid)
+
+lemma tmbound_split0: "tmbound 0 t \<Longrightarrow> Ipoly vs (fst(split0 t)) = 0"
+ by (induct t rule: split0.induct, auto simp add: Let_def split_def split0_stupid)
+
+lemma tmboundslt_split0: "tmboundslt n t \<Longrightarrow> Ipoly vs (fst(split0 t)) = 0 \<or> n > 0"
+by (induct t rule: split0.induct, auto simp add: Let_def split_def split0_stupid)
+
+lemma tmboundslt0_split0: "tmboundslt 0 t \<Longrightarrow> Ipoly vs (fst(split0 t)) = 0"
+ by (induct t rule: split0.induct, auto simp add: Let_def split_def split0_stupid)
+
+lemma allpolys_split0: "allpolys isnpoly p \<Longrightarrow> allpolys isnpoly (snd (split0 p))"
+by (induct p rule: split0.induct, auto simp  add: isnpoly_def Let_def split_def split0_stupid)
+
+lemma isnpoly_fst_split0:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+  shows 
+  "allpolys isnpoly p \<Longrightarrow> isnpoly (fst (split0 p))"
+  by (induct p rule: split0.induct, 
+    auto simp  add: polyadd_norm polysub_norm polyneg_norm polymul_norm 
+    Let_def split_def split0_stupid)
+
+subsection{* Formulae *}
+
+datatype fm  =  T| F| Le tm | Lt tm | Eq tm | NEq tm|
+  NOT fm| And fm fm|  Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm
+
+
+  (* A size for fm *)
+consts fmsize :: "fm \<Rightarrow> nat"
+recdef fmsize "measure size"
+  "fmsize (NOT p) = 1 + fmsize p"
+  "fmsize (And p q) = 1 + fmsize p + fmsize q"
+  "fmsize (Or p q) = 1 + fmsize p + fmsize q"
+  "fmsize (Imp p q) = 3 + fmsize p + fmsize q"
+  "fmsize (Iff p q) = 3 + 2*(fmsize p + fmsize q)"
+  "fmsize (E p) = 1 + fmsize p"
+  "fmsize (A p) = 4+ fmsize p"
+  "fmsize p = 1"
+  (* several lemmas about fmsize *)
+lemma fmsize_pos: "fmsize p > 0"	
+by (induct p rule: fmsize.induct) simp_all
+
+  (* Semantics of formulae (fm) *)
+consts Ifm ::"'a::{division_by_zero,ordered_field} list \<Rightarrow> 'a list \<Rightarrow> fm \<Rightarrow> bool"
+primrec
+  "Ifm vs bs T = True"
+  "Ifm vs bs F = False"
+  "Ifm vs bs (Lt a) = (Itm vs bs a < 0)"
+  "Ifm vs bs (Le a) = (Itm vs bs a \<le> 0)"
+  "Ifm vs bs (Eq a) = (Itm vs bs a = 0)"
+  "Ifm vs bs (NEq a) = (Itm vs bs a \<noteq> 0)"
+  "Ifm vs bs (NOT p) = (\<not> (Ifm vs bs p))"
+  "Ifm vs bs (And p q) = (Ifm vs bs p \<and> Ifm vs bs q)"
+  "Ifm vs bs (Or p q) = (Ifm vs bs p \<or> Ifm vs bs q)"
+  "Ifm vs bs (Imp p q) = ((Ifm vs bs p) \<longrightarrow> (Ifm vs bs q))"
+  "Ifm vs bs (Iff p q) = (Ifm vs bs p = Ifm vs bs q)"
+  "Ifm vs bs (E p) = (\<exists> x. Ifm vs (x#bs) p)"
+  "Ifm vs bs (A p) = (\<forall> x. Ifm vs (x#bs) p)"
+
+consts not:: "fm \<Rightarrow> fm"
+recdef not "measure size"
+  "not (NOT (NOT p)) = not p"
+  "not (NOT p) = p"
+  "not T = F"
+  "not F = T"
+  "not (Lt t) = Le (tmneg t)"
+  "not (Le t) = Lt (tmneg t)"
+  "not (Eq t) = NEq t"
+  "not (NEq t) = Eq t"
+  "not p = NOT p"
+lemma not[simp]: "Ifm vs bs (not p) = Ifm vs bs (NOT p)"
+by (induct p rule: not.induct) auto
+
+constdefs conj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+  "conj p q \<equiv> (if (p = F \<or> q=F) then F else if p=T then q else if q=T then p else 
+   if p = q then p else And p q)"
+lemma conj[simp]: "Ifm vs bs (conj p q) = Ifm vs bs (And p q)"
+by (cases "p=F \<or> q=F",simp_all add: conj_def) (cases p,simp_all)
+
+constdefs disj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+  "disj p q \<equiv> (if (p = T \<or> q=T) then T else if p=F then q else if q=F then p 
+       else if p=q then p else Or p q)"
+
+lemma disj[simp]: "Ifm vs bs (disj p q) = Ifm vs bs (Or p q)"
+by (cases "p=T \<or> q=T",simp_all add: disj_def) (cases p,simp_all)
+
+constdefs  imp :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+  "imp p q \<equiv> (if (p = F \<or> q=T \<or> p=q) then T else if p=T then q else if q=F then not p 
+    else Imp p q)"
+lemma imp[simp]: "Ifm vs bs (imp p q) = Ifm vs bs (Imp p q)"
+by (cases "p=F \<or> q=T",simp_all add: imp_def) 
+
+constdefs   iff :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+  "iff p q \<equiv> (if (p = q) then T else if (p = NOT q \<or> NOT p = q) then F else 
+       if p=F then not q else if q=F then not p else if p=T then q else if q=T then p else 
+  Iff p q)"
+lemma iff[simp]: "Ifm vs bs (iff p q) = Ifm vs bs (Iff p q)"
+  by (unfold iff_def,cases "p=q", simp,cases "p=NOT q", simp) (cases "NOT p= q", auto)
+  (* Quantifier freeness *)
+consts qfree:: "fm \<Rightarrow> bool"
+recdef qfree "measure size"
+  "qfree (E p) = False"
+  "qfree (A p) = False"
+  "qfree (NOT p) = qfree p" 
+  "qfree (And p q) = (qfree p \<and> qfree q)" 
+  "qfree (Or  p q) = (qfree p \<and> qfree q)" 
+  "qfree (Imp p q) = (qfree p \<and> qfree q)" 
+  "qfree (Iff p q) = (qfree p \<and> qfree q)"
+  "qfree p = True"
+
+  (* Boundedness and substitution *)
+
+consts boundslt :: "nat \<Rightarrow> fm \<Rightarrow> bool"
+primrec
+  "boundslt n T = True"
+  "boundslt n F = True"
+  "boundslt n (Lt t) = (tmboundslt n t)"
+  "boundslt n (Le t) = (tmboundslt n t)"
+  "boundslt n (Eq t) = (tmboundslt n t)"
+  "boundslt n (NEq t) = (tmboundslt n t)"
+  "boundslt n (NOT p) = boundslt n p"
+  "boundslt n (And p q) = (boundslt n p \<and> boundslt n q)"
+  "boundslt n (Or p q) = (boundslt n p \<and> boundslt n q)"
+  "boundslt n (Imp p q) = ((boundslt n p) \<and> (boundslt n q))"
+  "boundslt n (Iff p q) = (boundslt n p \<and> boundslt n q)"
+  "boundslt n (E p) = boundslt (Suc n) p"
+  "boundslt n (A p) = boundslt (Suc n) p"
+
+consts 
+  bound0:: "fm \<Rightarrow> bool" (* A Formula is independent of Bound 0 *)
+  bound:: "nat \<Rightarrow> fm \<Rightarrow> bool" (* A Formula is independent of Bound n *)
+  decr0 :: "fm \<Rightarrow> fm"
+  decr :: "nat \<Rightarrow> fm \<Rightarrow> fm"
+recdef bound0 "measure size"
+  "bound0 T = True"
+  "bound0 F = True"
+  "bound0 (Lt a) = tmbound0 a"
+  "bound0 (Le a) = tmbound0 a"
+  "bound0 (Eq a) = tmbound0 a"
+  "bound0 (NEq a) = tmbound0 a"
+  "bound0 (NOT p) = bound0 p"
+  "bound0 (And p q) = (bound0 p \<and> bound0 q)"
+  "bound0 (Or p q) = (bound0 p \<and> bound0 q)"
+  "bound0 (Imp p q) = ((bound0 p) \<and> (bound0 q))"
+  "bound0 (Iff p q) = (bound0 p \<and> bound0 q)"
+  "bound0 p = False"
+lemma bound0_I:
+  assumes bp: "bound0 p"
+  shows "Ifm vs (b#bs) p = Ifm vs (b'#bs) p"
+using bp tmbound0_I[where b="b" and bs="bs" and b'="b'"]
+by (induct p rule: bound0.induct,auto simp add: nth_pos2)
+
+primrec
+  "bound m T = True"
+  "bound m F = True"
+  "bound m (Lt t) = tmbound m t"
+  "bound m (Le t) = tmbound m t"
+  "bound m (Eq t) = tmbound m t"
+  "bound m (NEq t) = tmbound m t"
+  "bound m (NOT p) = bound m p"
+  "bound m (And p q) = (bound m p \<and> bound m q)"
+  "bound m (Or p q) = (bound m p \<and> bound m q)"
+  "bound m (Imp p q) = ((bound m p) \<and> (bound m q))"
+  "bound m (Iff p q) = (bound m p \<and> bound m q)"
+  "bound m (E p) = bound (Suc m) p"
+  "bound m (A p) = bound (Suc m) p"
+
+lemma bound_I:
+  assumes bnd: "boundslt (length bs) p" and nb: "bound n p" and le: "n \<le> length bs"
+  shows "Ifm vs (bs[n:=x]) p = Ifm vs bs p"
+  using bnd nb le tmbound_I[where bs=bs and vs = vs]
+proof(induct p arbitrary: bs n rule: bound.induct)
+  case (E p bs n) 
+  {fix y
+    from prems have bnd: "boundslt (length (y#bs)) p" 
+      and nb: "bound (Suc n) p" and le: "Suc n \<le> length (y#bs)" by simp+
+    from E.hyps[OF bnd nb le tmbound_I] have "Ifm vs ((y#bs)[Suc n:=x]) p = Ifm vs (y#bs) p" .   }
+  thus ?case by simp 
+next
+  case (A p bs n) {fix y
+    from prems have bnd: "boundslt (length (y#bs)) p" 
+      and nb: "bound (Suc n) p" and le: "Suc n \<le> length (y#bs)" by simp+
+    from A.hyps[OF bnd nb le tmbound_I] have "Ifm vs ((y#bs)[Suc n:=x]) p = Ifm vs (y#bs) p" .   }
+  thus ?case by simp 
+qed auto
+
+recdef decr0 "measure size"
+  "decr0 (Lt a) = Lt (decrtm0 a)"
+  "decr0 (Le a) = Le (decrtm0 a)"
+  "decr0 (Eq a) = Eq (decrtm0 a)"
+  "decr0 (NEq a) = NEq (decrtm0 a)"
+  "decr0 (NOT p) = NOT (decr0 p)" 
+  "decr0 (And p q) = conj (decr0 p) (decr0 q)"
+  "decr0 (Or p q) = disj (decr0 p) (decr0 q)"
+  "decr0 (Imp p q) = imp (decr0 p) (decr0 q)"
+  "decr0 (Iff p q) = iff (decr0 p) (decr0 q)"
+  "decr0 p = p"
+
+lemma decr0: assumes nb: "bound0 p"
+  shows "Ifm vs (x#bs) p = Ifm vs bs (decr0 p)"
+  using nb 
+  by (induct p rule: decr0.induct, simp_all add: decrtm0)
+
+primrec
+  "decr m T = T"
+  "decr m F = F"
+  "decr m (Lt t) = (Lt (decrtm m t))"
+  "decr m (Le t) = (Le (decrtm m t))"
+  "decr m (Eq t) = (Eq (decrtm m t))"
+  "decr m (NEq t) = (NEq (decrtm m t))"
+  "decr m (NOT p) = NOT (decr m p)" 
+  "decr m (And p q) = conj (decr m p) (decr m q)"
+  "decr m (Or p q) = disj (decr m p) (decr m q)"
+  "decr m (Imp p q) = imp (decr m p) (decr m q)"
+  "decr m (Iff p q) = iff (decr m p) (decr m q)"
+  "decr m (E p) = E (decr (Suc m) p)"
+  "decr m (A p) = A (decr (Suc m) p)"
+
+lemma decr: assumes  bnd: "boundslt (length bs) p" and nb: "bound m p" 
+  and nle: "m < length bs" 
+  shows "Ifm vs (removen m bs) (decr m p) = Ifm vs bs p"
+  using bnd nb nle
+proof(induct p arbitrary: bs m rule: decr.induct)
+  case (E p bs m) 
+  {fix x
+    from prems have bnd: "boundslt (length (x#bs)) p" and nb: "bound (Suc m) p" 
+  and nle: "Suc m < length (x#bs)" by auto
+    from prems(4)[OF bnd nb nle] have "Ifm vs (removen (Suc m) (x#bs)) (decr (Suc m) p) = Ifm vs (x#bs) p".
+  } thus ?case by auto 
+next
+  case (A p bs m)  
+  {fix x
+    from prems have bnd: "boundslt (length (x#bs)) p" and nb: "bound (Suc m) p" 
+  and nle: "Suc m < length (x#bs)" by auto
+    from prems(4)[OF bnd nb nle] have "Ifm vs (removen (Suc m) (x#bs)) (decr (Suc m) p) = Ifm vs (x#bs) p".
+  } thus ?case by auto
+qed (auto simp add: decrtm removen_nth)
+
+consts
+  subst0:: "tm \<Rightarrow> fm \<Rightarrow> fm"
+
+primrec
+  "subst0 t T = T"
+  "subst0 t F = F"
+  "subst0 t (Lt a) = Lt (tmsubst0 t a)"
+  "subst0 t (Le a) = Le (tmsubst0 t a)"
+  "subst0 t (Eq a) = Eq (tmsubst0 t a)"
+  "subst0 t (NEq a) = NEq (tmsubst0 t a)"
+  "subst0 t (NOT p) = NOT (subst0 t p)"
+  "subst0 t (And p q) = And (subst0 t p) (subst0 t q)"
+  "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)"
+  "subst0 t (Imp p q) = Imp (subst0 t p)  (subst0 t q)"
+  "subst0 t (Iff p q) = Iff (subst0 t p) (subst0 t q)"
+  "subst0 t (E p) = E p"
+  "subst0 t (A p) = A p"
+
+lemma subst0: assumes qf: "qfree p"
+  shows "Ifm vs (x#bs) (subst0 t p) = Ifm vs ((Itm vs (x#bs) t)#bs) p"
+using qf tmsubst0[where x="x" and bs="bs" and t="t"]
+by (induct p rule: subst0.induct, auto)
+
+lemma subst0_nb:
+  assumes bp: "tmbound0 t" and qf: "qfree p"
+  shows "bound0 (subst0 t p)"
+using qf tmsubst0_nb[OF bp] bp
+by (induct p rule: subst0.induct, auto)
+
+consts   subst:: "nat \<Rightarrow> tm \<Rightarrow> fm \<Rightarrow> fm" 
+primrec
+  "subst n t T = T"
+  "subst n t F = F"
+  "subst n t (Lt a) = Lt (tmsubst n t a)"
+  "subst n t (Le a) = Le (tmsubst n t a)"
+  "subst n t (Eq a) = Eq (tmsubst n t a)"
+  "subst n t (NEq a) = NEq (tmsubst n t a)"
+  "subst n t (NOT p) = NOT (subst n t p)"
+  "subst n t (And p q) = And (subst n t p) (subst n t q)"
+  "subst n t (Or p q) = Or (subst n t p) (subst n t q)"
+  "subst n t (Imp p q) = Imp (subst n t p)  (subst n t q)"
+  "subst n t (Iff p q) = Iff (subst n t p) (subst n t q)"
+  "subst n t (E p) = E (subst (Suc n) (incrtm0 t) p)"
+  "subst n t (A p) = A (subst (Suc n) (incrtm0 t) p)"
+
+lemma subst: assumes nb: "boundslt (length bs) p" and nlm: "n \<le> length bs"
+  shows "Ifm vs bs (subst n t p) = Ifm vs (bs[n:= Itm vs bs t]) p"
+  using nb nlm
+proof (induct p arbitrary: bs n t rule: subst0.induct)
+  case (E p bs n) 
+  {fix x 
+    from prems have bn: "boundslt (length (x#bs)) p" by simp 
+      from prems have nlm: "Suc n \<le> length (x#bs)" by simp
+    from prems(3)[OF bn nlm] have "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) = Ifm vs ((x#bs)[Suc n:= Itm vs (x#bs) (incrtm0 t)]) p" by simp 
+    hence "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) = Ifm vs (x#bs[n:= Itm vs bs t]) p"
+    by (simp add: incrtm0[where x="x" and bs="bs" and t="t"]) }  
+thus ?case by simp 
+next
+  case (A p bs n)   
+  {fix x 
+    from prems have bn: "boundslt (length (x#bs)) p" by simp 
+      from prems have nlm: "Suc n \<le> length (x#bs)" by simp
+    from prems(3)[OF bn nlm] have "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) = Ifm vs ((x#bs)[Suc n:= Itm vs (x#bs) (incrtm0 t)]) p" by simp 
+    hence "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) = Ifm vs (x#bs[n:= Itm vs bs t]) p"
+    by (simp add: incrtm0[where x="x" and bs="bs" and t="t"]) }  
+thus ?case by simp 
+qed(auto simp add: tmsubst)
+
+lemma subst_nb: assumes tnb: "tmbound m t"
+shows "bound m (subst m t p)"
+using tnb tmsubst_nb incrtm0_tmbound
+by (induct p arbitrary: m t rule: subst.induct, auto)
+
+lemma not_qf[simp]: "qfree p \<Longrightarrow> qfree (not p)"
+by (induct p rule: not.induct, auto)
+lemma not_bn0[simp]: "bound0 p \<Longrightarrow> bound0 (not p)"
+by (induct p rule: not.induct, auto)
+lemma not_nb[simp]: "bound n p \<Longrightarrow> bound n (not p)"
+by (induct p rule: not.induct, auto)
+lemma not_blt[simp]: "boundslt n p \<Longrightarrow> boundslt n (not p)"
+ by (induct p rule: not.induct, auto)
+
+lemma conj_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (conj p q)"
+using conj_def by auto 
+lemma conj_nb0[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (conj p q)"
+using conj_def by auto 
+lemma conj_nb[simp]: "\<lbrakk>bound n p ; bound n q\<rbrakk> \<Longrightarrow> bound n (conj p q)"
+using conj_def by auto 
+lemma conj_blt[simp]: "boundslt n p \<Longrightarrow> boundslt n q \<Longrightarrow> boundslt n (conj p q)"
+using conj_def by auto 
+
+lemma disj_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (disj p q)"
+using disj_def by auto 
+lemma disj_nb0[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (disj p q)"
+using disj_def by auto 
+lemma disj_nb[simp]: "\<lbrakk>bound n p ; bound n q\<rbrakk> \<Longrightarrow> bound n (disj p q)"
+using disj_def by auto 
+lemma disj_blt[simp]: "boundslt n p \<Longrightarrow> boundslt n q \<Longrightarrow> boundslt n (disj p q)"
+using disj_def by auto 
+
+lemma imp_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (imp p q)"
+using imp_def by (cases "p=F \<or> q=T",simp_all add: imp_def)
+lemma imp_nb0[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (imp p q)"
+using imp_def by (cases "p=F \<or> q=T \<or> p=q",simp_all add: imp_def)
+lemma imp_nb[simp]: "\<lbrakk>bound n p ; bound n q\<rbrakk> \<Longrightarrow> bound n (imp p q)"
+using imp_def by (cases "p=F \<or> q=T \<or> p=q",simp_all add: imp_def)
+lemma imp_blt[simp]: "boundslt n p \<Longrightarrow> boundslt n q \<Longrightarrow> boundslt n (imp p q)"
+using imp_def by auto 
+
+lemma iff_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (iff p q)"
+  by (unfold iff_def,cases "p=q", auto)
+lemma iff_nb0[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (iff p q)"
+using iff_def by (unfold iff_def,cases "p=q", auto)
+lemma iff_nb[simp]: "\<lbrakk>bound n p ; bound n q\<rbrakk> \<Longrightarrow> bound n (iff p q)"
+using iff_def by (unfold iff_def,cases "p=q", auto)
+lemma iff_blt[simp]: "boundslt n p \<Longrightarrow> boundslt n q \<Longrightarrow> boundslt n (iff p q)"
+using iff_def by auto 
+lemma decr0_qf: "bound0 p \<Longrightarrow> qfree (decr0 p)"
+by (induct p, simp_all)
+
+consts 
+  isatom :: "fm \<Rightarrow> bool" (* test for atomicity *)
+recdef isatom "measure size"
+  "isatom T = True"
+  "isatom F = True"
+  "isatom (Lt a) = True"
+  "isatom (Le a) = True"
+  "isatom (Eq a) = True"
+  "isatom (NEq a) = True"
+  "isatom p = False"
+
+lemma bound0_qf: "bound0 p \<Longrightarrow> qfree p"
+by (induct p, simp_all)
+
+constdefs djf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm"
+  "djf f p q \<equiv> (if q=T then T else if q=F then f p else 
+  (let fp = f p in case fp of T \<Rightarrow> T | F \<Rightarrow> q | _ \<Rightarrow> Or (f p) q))"
+constdefs evaldjf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm"
+  "evaldjf f ps \<equiv> foldr (djf f) ps F"
+
+lemma djf_Or: "Ifm vs bs (djf f p q) = Ifm vs bs (Or (f p) q)"
+by (cases "q=T", simp add: djf_def,cases "q=F",simp add: djf_def) 
+(cases "f p", simp_all add: Let_def djf_def) 
+
+lemma evaldjf_ex: "Ifm vs bs (evaldjf f ps) = (\<exists> p \<in> set ps. Ifm vs bs (f p))"
+  by(induct ps, simp_all add: evaldjf_def djf_Or)
+
+lemma evaldjf_bound0: 
+  assumes nb: "\<forall> x\<in> set xs. bound0 (f x)"
+  shows "bound0 (evaldjf f xs)"
+  using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto) 
+
+lemma evaldjf_qf: 
+  assumes nb: "\<forall> x\<in> set xs. qfree (f x)"
+  shows "qfree (evaldjf f xs)"
+  using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto) 
+
+consts disjuncts :: "fm \<Rightarrow> fm list"
+recdef disjuncts "measure size"
+  "disjuncts (Or p q) = (disjuncts p) @ (disjuncts q)"
+  "disjuncts F = []"
+  "disjuncts p = [p]"
+
+lemma disjuncts: "(\<exists> q\<in> set (disjuncts p). Ifm vs bs q) = Ifm vs bs p"
+by(induct p rule: disjuncts.induct, auto)
+
+lemma disjuncts_nb: "bound0 p \<Longrightarrow> \<forall> q\<in> set (disjuncts p). bound0 q"
+proof-
+  assume nb: "bound0 p"
+  hence "list_all bound0 (disjuncts p)" by (induct p rule:disjuncts.induct,auto)
+  thus ?thesis by (simp only: list_all_iff)
+qed
+
+lemma disjuncts_qf: "qfree p \<Longrightarrow> \<forall> q\<in> set (disjuncts p). qfree q"
+proof-
+  assume qf: "qfree p"
+  hence "list_all qfree (disjuncts p)"
+    by (induct p rule: disjuncts.induct, auto)
+  thus ?thesis by (simp only: list_all_iff)
+qed
+
+constdefs DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
+  "DJ f p \<equiv> evaldjf f (disjuncts p)"
+
+lemma DJ: assumes fdj: "\<forall> p q. Ifm vs bs (f (Or p q)) = Ifm vs bs (Or (f p) (f q))"
+  and fF: "f F = F"
+  shows "Ifm vs bs (DJ f p) = Ifm vs bs (f p)"
+proof-
+  have "Ifm vs bs (DJ f p) = (\<exists> q \<in> set (disjuncts p). Ifm vs bs (f q))"
+    by (simp add: DJ_def evaldjf_ex) 
+  also have "\<dots> = Ifm vs bs (f p)" using fdj fF by (induct p rule: disjuncts.induct, auto)
+  finally show ?thesis .
+qed
+
+lemma DJ_qf: assumes 
+  fqf: "\<forall> p. qfree p \<longrightarrow> qfree (f p)"
+  shows "\<forall>p. qfree p \<longrightarrow> qfree (DJ f p) "
+proof(clarify)
+  fix  p assume qf: "qfree p"
+  have th: "DJ f p = evaldjf f (disjuncts p)" by (simp add: DJ_def)
+  from disjuncts_qf[OF qf] have "\<forall> q\<in> set (disjuncts p). qfree q" .
+  with fqf have th':"\<forall> q\<in> set (disjuncts p). qfree (f q)" by blast
+  
+  from evaldjf_qf[OF th'] th show "qfree (DJ f p)" by simp
+qed
+
+lemma DJ_qe: assumes qe: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm vs bs (qe p) = Ifm vs bs (E p))"
+  shows "\<forall> bs p. qfree p \<longrightarrow> qfree (DJ qe p) \<and> (Ifm vs bs ((DJ qe p)) = Ifm vs bs (E p))"
+proof(clarify)
+  fix p::fm and bs
+  assume qf: "qfree p"
+  from qe have qth: "\<forall> p. qfree p \<longrightarrow> qfree (qe p)" by blast
+  from DJ_qf[OF qth] qf have qfth:"qfree (DJ qe p)" by auto
+  have "Ifm vs bs (DJ qe p) = (\<exists> q\<in> set (disjuncts p). Ifm vs bs (qe q))"
+    by (simp add: DJ_def evaldjf_ex)
+  also have "\<dots> = (\<exists> q \<in> set(disjuncts p). Ifm vs bs (E q))" using qe disjuncts_qf[OF qf] by auto
+  also have "\<dots> = Ifm vs bs (E p)" by (induct p rule: disjuncts.induct, auto)
+  finally show "qfree (DJ qe p) \<and> Ifm vs bs (DJ qe p) = Ifm vs bs (E p)" using qfth by blast
+qed
+
+consts conjuncts :: "fm \<Rightarrow> fm list"
+
+recdef conjuncts "measure size"
+  "conjuncts (And p q) = (conjuncts p) @ (conjuncts q)"
+  "conjuncts T = []"
+  "conjuncts p = [p]"
+
+constdefs list_conj :: "fm list \<Rightarrow> fm"
+  "list_conj ps \<equiv> foldr conj ps T"
+
+constdefs CJNB:: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
+  "CJNB f p \<equiv> (let cjs = conjuncts p ; (yes,no) = partition bound0 cjs
+                   in conj (decr0 (list_conj yes)) (f (list_conj no)))"
+
+lemma conjuncts_qf: "qfree p \<Longrightarrow> \<forall> q\<in> set (conjuncts p). qfree q"
+proof-
+  assume qf: "qfree p"
+  hence "list_all qfree (conjuncts p)"
+    by (induct p rule: conjuncts.induct, auto)
+  thus ?thesis by (simp only: list_all_iff)
+qed
+
+lemma conjuncts: "(\<forall> q\<in> set (conjuncts p). Ifm vs bs q) = Ifm vs bs p"
+by(induct p rule: conjuncts.induct, auto)
+
+lemma conjuncts_nb: "bound0 p \<Longrightarrow> \<forall> q\<in> set (conjuncts p). bound0 q"
+proof-
+  assume nb: "bound0 p"
+  hence "list_all bound0 (conjuncts p)" by (induct p rule:conjuncts.induct,auto)
+  thus ?thesis by (simp only: list_all_iff)
+qed
+
+fun islin :: "fm \<Rightarrow> bool" where
+  "islin (And p q) = (islin p \<and> islin q \<and> p \<noteq> T \<and> p \<noteq> F \<and> q \<noteq> T \<and> q \<noteq> F)"
+| "islin (Or p q) = (islin p \<and> islin q \<and> p \<noteq> T \<and> p \<noteq> F \<and> q \<noteq> T \<and> q \<noteq> F)"
+| "islin (Eq (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
+| "islin (NEq (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
+| "islin (Lt (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
+| "islin (Le (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
+| "islin (NOT p) = False"
+| "islin (Imp p q) = False"
+| "islin (Iff p q) = False"
+| "islin p = bound0 p"
+
+lemma islin_stupid: assumes nb: "tmbound0 p"
+  shows "islin (Lt p)" and "islin (Le p)" and "islin (Eq p)" and "islin (NEq p)"
+  using nb by (cases p, auto, case_tac nat, auto)+
+
+definition "lt p = (case p of CP (C c) \<Rightarrow> if 0>\<^sub>N c then T else F| _ \<Rightarrow> Lt p)"
+definition "le p = (case p of CP (C c) \<Rightarrow> if 0\<ge>\<^sub>N c then T else F | _ \<Rightarrow> Le p)"
+definition "eq p = (case p of CP (C c) \<Rightarrow> if c = 0\<^sub>N then T else F | _ \<Rightarrow> Eq p)"
+definition "neq p = not (eq p)"
+
+lemma lt: "allpolys isnpoly p \<Longrightarrow> Ifm vs bs (lt p) = Ifm vs bs (Lt p)"
+  apply(simp add: lt_def)
+  apply(cases p, simp_all)
+  apply (case_tac poly, simp_all add: isnpoly_def)
+  done
+
+lemma le: "allpolys isnpoly p \<Longrightarrow> Ifm vs bs (le p) = Ifm vs bs (Le p)"
+  apply(simp add: le_def)
+  apply(cases p, simp_all)
+  apply (case_tac poly, simp_all add: isnpoly_def)
+  done
+
+lemma eq: "allpolys isnpoly p \<Longrightarrow> Ifm vs bs (eq p) = Ifm vs bs (Eq p)"
+  apply(simp add: eq_def)
+  apply(cases p, simp_all)
+  apply (case_tac poly, simp_all add: isnpoly_def)
+  done
+
+lemma neq: "allpolys isnpoly p \<Longrightarrow> Ifm vs bs (neq p) = Ifm vs bs (NEq p)"
+  by(simp add: neq_def eq)
+
+lemma lt_lin: "tmbound0 p \<Longrightarrow> islin (lt p)"
+  apply (simp add: lt_def)
+  apply (cases p, simp_all)
+  apply (case_tac poly, simp_all)
+  apply (case_tac nat, simp_all)
+  done
+
+lemma le_lin: "tmbound0 p \<Longrightarrow> islin (le p)"
+  apply (simp add: le_def)
+  apply (cases p, simp_all)
+  apply (case_tac poly, simp_all)
+  apply (case_tac nat, simp_all)
+  done
+
+lemma eq_lin: "tmbound0 p \<Longrightarrow> islin (eq p)"
+  apply (simp add: eq_def)
+  apply (cases p, simp_all)
+  apply (case_tac poly, simp_all)
+  apply (case_tac nat, simp_all)
+  done
+
+lemma neq_lin: "tmbound0 p \<Longrightarrow> islin (neq p)"
+  apply (simp add: neq_def eq_def)
+  apply (cases p, simp_all)
+  apply (case_tac poly, simp_all)
+  apply (case_tac nat, simp_all)
+  done
+
+definition "simplt t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then lt s else Lt (CNP 0 c s))"
+definition "simple t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then le s else Le (CNP 0 c s))"
+definition "simpeq t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then eq s else Eq (CNP 0 c s))"
+definition "simpneq t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then neq s else NEq (CNP 0 c s))"
+
+lemma simplt_islin[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "islin (simplt t)"
+  unfolding simplt_def 
+  using split0_nb0'
+by (auto simp add: lt_lin Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly] islin_stupid allpolys_split0[OF simptm_allpolys_npoly])
+  
+lemma simple_islin[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "islin (simple t)"
+  unfolding simple_def 
+  using split0_nb0'
+by (auto simp add: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly] islin_stupid allpolys_split0[OF simptm_allpolys_npoly] le_lin)
+lemma simpeq_islin[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "islin (simpeq t)"
+  unfolding simpeq_def 
+  using split0_nb0'
+by (auto simp add: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly] islin_stupid allpolys_split0[OF simptm_allpolys_npoly] eq_lin)
+
+lemma simpneq_islin[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "islin (simpneq t)"
+  unfolding simpneq_def 
+  using split0_nb0'
+by (auto simp add: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly] islin_stupid allpolys_split0[OF simptm_allpolys_npoly] neq_lin)
+
+lemma really_stupid: "\<not> (\<forall>c1 s'. (c1, s') \<noteq> split0 s)"
+  by (cases "split0 s", auto)
+lemma split0_npoly:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  and n: "allpolys isnpoly t"
+  shows "isnpoly (fst (split0 t))" and "allpolys isnpoly (snd (split0 t))"
+  using n
+  by (induct t rule: split0.induct, auto simp add: Let_def split_def polyadd_norm polymul_norm polyneg_norm polysub_norm really_stupid)
+lemma simplt[simp]:
+  shows "Ifm vs bs (simplt t) = Ifm vs bs (Lt t)"
+proof-
+  have n: "allpolys isnpoly (simptm t)" by simp
+  let ?t = "simptm t"
+  {assume "fst (split0 ?t) = 0\<^sub>p" hence ?thesis
+      using split0[of "simptm t" vs bs] lt[OF split0_npoly(2)[OF n], of vs bs]
+      by (simp add: simplt_def Let_def split_def lt)}
+  moreover
+  {assume "fst (split0 ?t) \<noteq> 0\<^sub>p"
+    hence ?thesis using  split0[of "simptm t" vs bs] by (simp add: simplt_def Let_def split_def)
+  }
+  ultimately show ?thesis by blast
+qed
+
+lemma simple[simp]:
+  shows "Ifm vs bs (simple t) = Ifm vs bs (Le t)"
+proof-
+  have n: "allpolys isnpoly (simptm t)" by simp
+  let ?t = "simptm t"
+  {assume "fst (split0 ?t) = 0\<^sub>p" hence ?thesis
+      using split0[of "simptm t" vs bs] le[OF split0_npoly(2)[OF n], of vs bs]
+      by (simp add: simple_def Let_def split_def le)}
+  moreover
+  {assume "fst (split0 ?t) \<noteq> 0\<^sub>p"
+    hence ?thesis using  split0[of "simptm t" vs bs] by (simp add: simple_def Let_def split_def)
+  }
+  ultimately show ?thesis by blast
+qed
+
+lemma simpeq[simp]:
+  shows "Ifm vs bs (simpeq t) = Ifm vs bs (Eq t)"
+proof-
+  have n: "allpolys isnpoly (simptm t)" by simp
+  let ?t = "simptm t"
+  {assume "fst (split0 ?t) = 0\<^sub>p" hence ?thesis
+      using split0[of "simptm t" vs bs] eq[OF split0_npoly(2)[OF n], of vs bs]
+      by (simp add: simpeq_def Let_def split_def)}
+  moreover
+  {assume "fst (split0 ?t) \<noteq> 0\<^sub>p"
+    hence ?thesis using  split0[of "simptm t" vs bs] by (simp add: simpeq_def Let_def split_def)
+  }
+  ultimately show ?thesis by blast
+qed
+
+lemma simpneq[simp]:
+  shows "Ifm vs bs (simpneq t) = Ifm vs bs (NEq t)"
+proof-
+  have n: "allpolys isnpoly (simptm t)" by simp
+  let ?t = "simptm t"
+  {assume "fst (split0 ?t) = 0\<^sub>p" hence ?thesis
+      using split0[of "simptm t" vs bs] neq[OF split0_npoly(2)[OF n], of vs bs]
+      by (simp add: simpneq_def Let_def split_def )}
+  moreover
+  {assume "fst (split0 ?t) \<noteq> 0\<^sub>p"
+    hence ?thesis using  split0[of "simptm t" vs bs] by (simp add: simpneq_def Let_def split_def)
+  }
+  ultimately show ?thesis by blast
+qed
+
+lemma lt_nb: "tmbound0 t \<Longrightarrow> bound0 (lt t)"
+  apply (simp add: lt_def)
+  apply (cases t, auto)
+  apply (case_tac poly, auto)
+  done
+
+lemma le_nb: "tmbound0 t \<Longrightarrow> bound0 (le t)"
+  apply (simp add: le_def)
+  apply (cases t, auto)
+  apply (case_tac poly, auto)
+  done
+
+lemma eq_nb: "tmbound0 t \<Longrightarrow> bound0 (eq t)"
+  apply (simp add: eq_def)
+  apply (cases t, auto)
+  apply (case_tac poly, auto)
+  done
+
+lemma neq_nb: "tmbound0 t \<Longrightarrow> bound0 (neq t)"
+  apply (simp add: neq_def eq_def)
+  apply (cases t, auto)
+  apply (case_tac poly, auto)
+  done
+
+lemma simplt_nb[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "tmbound0 t \<Longrightarrow> bound0 (simplt t)"
+  using split0 [of "simptm t" vs bs]
+proof(simp add: simplt_def Let_def split_def)
+  assume nb: "tmbound0 t"
+  hence nb': "tmbound0 (simptm t)" by simp
+  let ?c = "fst (split0 (simptm t))"
+  from tmbound_split0[OF nb'[unfolded tmbound0_tmbound_iff[symmetric]]]
+  have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p" by auto
+  from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
+  have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0" by (simp_all add: isnpoly_def)
+  from iffD1[OF isnpolyh_unique[OF ths] th]
+  have "fst (split0 (simptm t)) = 0\<^sub>p" . 
+  thus "(fst (split0 (simptm t)) = 0\<^sub>p \<longrightarrow> bound0 (lt (snd (split0 (simptm t))))) \<and>
+       fst (split0 (simptm t)) = 0\<^sub>p" by (simp add: simplt_def Let_def split_def lt_nb)
+qed
+
+lemma simple_nb[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "tmbound0 t \<Longrightarrow> bound0 (simple t)"
+  using split0 [of "simptm t" vs bs]
+proof(simp add: simple_def Let_def split_def)
+  assume nb: "tmbound0 t"
+  hence nb': "tmbound0 (simptm t)" by simp
+  let ?c = "fst (split0 (simptm t))"
+  from tmbound_split0[OF nb'[unfolded tmbound0_tmbound_iff[symmetric]]]
+  have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p" by auto
+  from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
+  have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0" by (simp_all add: isnpoly_def)
+  from iffD1[OF isnpolyh_unique[OF ths] th]
+  have "fst (split0 (simptm t)) = 0\<^sub>p" . 
+  thus "(fst (split0 (simptm t)) = 0\<^sub>p \<longrightarrow> bound0 (le (snd (split0 (simptm t))))) \<and>
+       fst (split0 (simptm t)) = 0\<^sub>p" by (simp add: simplt_def Let_def split_def le_nb)
+qed
+
+lemma simpeq_nb[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "tmbound0 t \<Longrightarrow> bound0 (simpeq t)"
+  using split0 [of "simptm t" vs bs]
+proof(simp add: simpeq_def Let_def split_def)
+  assume nb: "tmbound0 t"
+  hence nb': "tmbound0 (simptm t)" by simp
+  let ?c = "fst (split0 (simptm t))"
+  from tmbound_split0[OF nb'[unfolded tmbound0_tmbound_iff[symmetric]]]
+  have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p" by auto
+  from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
+  have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0" by (simp_all add: isnpoly_def)
+  from iffD1[OF isnpolyh_unique[OF ths] th]
+  have "fst (split0 (simptm t)) = 0\<^sub>p" . 
+  thus "(fst (split0 (simptm t)) = 0\<^sub>p \<longrightarrow> bound0 (eq (snd (split0 (simptm t))))) \<and>
+       fst (split0 (simptm t)) = 0\<^sub>p" by (simp add: simpeq_def Let_def split_def eq_nb)
+qed
+
+lemma simpneq_nb[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "tmbound0 t \<Longrightarrow> bound0 (simpneq t)"
+  using split0 [of "simptm t" vs bs]
+proof(simp add: simpneq_def Let_def split_def)
+  assume nb: "tmbound0 t"
+  hence nb': "tmbound0 (simptm t)" by simp
+  let ?c = "fst (split0 (simptm t))"
+  from tmbound_split0[OF nb'[unfolded tmbound0_tmbound_iff[symmetric]]]
+  have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p" by auto
+  from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
+  have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0" by (simp_all add: isnpoly_def)
+  from iffD1[OF isnpolyh_unique[OF ths] th]
+  have "fst (split0 (simptm t)) = 0\<^sub>p" . 
+  thus "(fst (split0 (simptm t)) = 0\<^sub>p \<longrightarrow> bound0 (neq (snd (split0 (simptm t))))) \<and>
+       fst (split0 (simptm t)) = 0\<^sub>p" by (simp add: simpneq_def Let_def split_def neq_nb)
+qed
+
+consts conjs   :: "fm \<Rightarrow> fm list"
+recdef conjs "measure size"
+  "conjs (And p q) = (conjs p)@(conjs q)"
+  "conjs T = []"
+  "conjs p = [p]"
+lemma conjs_ci: "(\<forall> q \<in> set (conjs p). Ifm vs bs q) = Ifm vs bs p"
+by (induct p rule: conjs.induct, auto)
+constdefs list_disj :: "fm list \<Rightarrow> fm"
+  "list_disj ps \<equiv> foldr disj ps F"
+
+lemma list_conj: "Ifm vs bs (list_conj ps) = (\<forall>p\<in> set ps. Ifm vs bs p)"
+  by (induct ps, auto simp add: list_conj_def)
+lemma list_conj_qf: " \<forall>p\<in> set ps. qfree p \<Longrightarrow> qfree (list_conj ps)"
+  by (induct ps, auto simp add: list_conj_def conj_qf)
+lemma list_disj: "Ifm vs bs (list_disj ps) = (\<exists>p\<in> set ps. Ifm vs bs p)"
+  by (induct ps, auto simp add: list_disj_def)
+
+lemma conj_boundslt: "boundslt n p \<Longrightarrow> boundslt n q \<Longrightarrow> boundslt n (conj p q)"
+  unfolding conj_def by auto
+
+lemma conjs_nb: "bound n p \<Longrightarrow> \<forall>q\<in> set (conjs p). bound n q"
+  apply (induct p rule: conjs.induct) 
+  apply (unfold conjs.simps)
+  apply (unfold set_append)
+  apply (unfold ball_Un)
+  apply (unfold bound.simps)
+  apply auto
+  done
+
+lemma conjs_boundslt: "boundslt n p \<Longrightarrow> \<forall>q\<in> set (conjs p). boundslt n q"
+  apply (induct p rule: conjs.induct) 
+  apply (unfold conjs.simps)
+  apply (unfold set_append)
+  apply (unfold ball_Un)
+  apply (unfold boundslt.simps)
+  apply blast
+by simp_all
+
+lemma list_conj_boundslt: " \<forall>p\<in> set ps. boundslt n p \<Longrightarrow> boundslt n (list_conj ps)"
+  unfolding list_conj_def
+  by (induct ps, auto simp add: conj_boundslt)
+
+lemma list_conj_nb: assumes bnd: "\<forall>p\<in> set ps. bound n p"
+  shows "bound n (list_conj ps)"
+  using bnd
+  unfolding list_conj_def
+  by (induct ps, auto simp add: conj_nb)
+
+lemma list_conj_nb': "\<forall>p\<in>set ps. bound0 p \<Longrightarrow> bound0 (list_conj ps)"
+unfolding list_conj_def by (induct ps , auto)
+
+lemma CJNB_qe: 
+  assumes qe: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm vs bs (qe p) = Ifm vs bs (E p))"
+  shows "\<forall> bs p. qfree p \<longrightarrow> qfree (CJNB qe p) \<and> (Ifm vs bs ((CJNB qe p)) = Ifm vs bs (E p))"
+proof(clarify)
+  fix bs p
+  assume qfp: "qfree p"
+  let ?cjs = "conjuncts p"
+  let ?yes = "fst (partition bound0 ?cjs)"
+  let ?no = "snd (partition bound0 ?cjs)"
+  let ?cno = "list_conj ?no"
+  let ?cyes = "list_conj ?yes"
+  have part: "partition bound0 ?cjs = (?yes,?no)" by simp
+  from partition_P[OF part] have "\<forall> q\<in> set ?yes. bound0 q" by blast 
+  hence yes_nb: "bound0 ?cyes" by (simp add: list_conj_nb') 
+  hence yes_qf: "qfree (decr0 ?cyes )" by (simp add: decr0_qf)
+  from conjuncts_qf[OF qfp] partition_set[OF part] 
+  have " \<forall>q\<in> set ?no. qfree q" by auto
+  hence no_qf: "qfree ?cno"by (simp add: list_conj_qf)
+  with qe have cno_qf:"qfree (qe ?cno )" 
+    and noE: "Ifm vs bs (qe ?cno) = Ifm vs bs (E ?cno)" by blast+
+  from cno_qf yes_qf have qf: "qfree (CJNB qe p)" 
+    by (simp add: CJNB_def Let_def conj_qf split_def)
+  {fix bs
+    from conjuncts have "Ifm vs bs p = (\<forall>q\<in> set ?cjs. Ifm vs bs q)" by blast
+    also have "\<dots> = ((\<forall>q\<in> set ?yes. Ifm vs bs q) \<and> (\<forall>q\<in> set ?no. Ifm vs bs q))"
+      using partition_set[OF part] by auto
+    finally have "Ifm vs bs p = ((Ifm vs bs ?cyes) \<and> (Ifm vs bs ?cno))" using list_conj[of vs bs] by simp}
+  hence "Ifm vs bs (E p) = (\<exists>x. (Ifm vs (x#bs) ?cyes) \<and> (Ifm vs (x#bs) ?cno))" by simp
+  also have "\<dots> = (\<exists>x. (Ifm vs (y#bs) ?cyes) \<and> (Ifm vs (x#bs) ?cno))"
+    using bound0_I[OF yes_nb, where bs="bs" and b'="y"] by blast
+  also have "\<dots> = (Ifm vs bs (decr0 ?cyes) \<and> Ifm vs bs (E ?cno))"
+    by (auto simp add: decr0[OF yes_nb])
+  also have "\<dots> = (Ifm vs bs (conj (decr0 ?cyes) (qe ?cno)))"
+    using qe[rule_format, OF no_qf] by auto
+  finally have "Ifm vs bs (E p) = Ifm vs bs (CJNB qe p)" 
+    by (simp add: Let_def CJNB_def split_def)
+  with qf show "qfree (CJNB qe p) \<and> Ifm vs bs (CJNB qe p) = Ifm vs bs (E p)" by blast
+qed
+
+consts simpfm :: "fm \<Rightarrow> fm"
+recdef simpfm "measure fmsize"
+  "simpfm (Lt t) = simplt (simptm t)"
+  "simpfm (Le t) = simple (simptm t)"
+  "simpfm (Eq t) = simpeq(simptm t)"
+  "simpfm (NEq t) = simpneq(simptm t)"
+  "simpfm (And p q) = conj (simpfm p) (simpfm q)"
+  "simpfm (Or p q) = disj (simpfm p) (simpfm q)"
+  "simpfm (Imp p q) = disj (simpfm (NOT p)) (simpfm q)"
+  "simpfm (Iff p q) = disj (conj (simpfm p) (simpfm q)) (conj (simpfm (NOT p)) (simpfm (NOT q)))"
+  "simpfm (NOT (And p q)) = disj (simpfm (NOT p)) (simpfm (NOT q))"
+  "simpfm (NOT (Or p q)) = conj (simpfm (NOT p)) (simpfm (NOT q))"
+  "simpfm (NOT (Imp p q)) = conj (simpfm p) (simpfm (NOT q))"
+  "simpfm (NOT (Iff p q)) = disj (conj (simpfm p) (simpfm (NOT q))) (conj (simpfm (NOT p)) (simpfm q))"
+  "simpfm (NOT (Eq t)) = simpneq t"
+  "simpfm (NOT (NEq t)) = simpeq t"
+  "simpfm (NOT (Le t)) = simplt (Neg t)"
+  "simpfm (NOT (Lt t)) = simple (Neg t)"
+  "simpfm (NOT (NOT p)) = simpfm p"
+  "simpfm (NOT T) = F"
+  "simpfm (NOT F) = T"
+  "simpfm p = p"
+
+lemma simpfm[simp]: "Ifm vs bs (simpfm p) = Ifm vs bs p"
+by(induct p arbitrary: bs rule: simpfm.induct, auto)
+
+lemma simpfm_bound0:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "bound0 p \<Longrightarrow> bound0 (simpfm p)"
+by (induct p rule: simpfm.induct, auto)
+
+lemma lt_qf[simp]: "qfree (lt t)"
+  apply (cases t, auto simp add: lt_def)
+  by (case_tac poly, auto)
+
+lemma le_qf[simp]: "qfree (le t)"
+  apply (cases t, auto simp add: le_def)
+  by (case_tac poly, auto)
+
+lemma eq_qf[simp]: "qfree (eq t)"
+  apply (cases t, auto simp add: eq_def)
+  by (case_tac poly, auto)
+
+lemma neq_qf[simp]: "qfree (neq t)" by (simp add: neq_def)
+
+lemma simplt_qf[simp]: "qfree (simplt t)" by (simp add: simplt_def Let_def split_def)
+lemma simple_qf[simp]: "qfree (simple t)" by (simp add: simple_def Let_def split_def)
+lemma simpeq_qf[simp]: "qfree (simpeq t)" by (simp add: simpeq_def Let_def split_def)
+lemma simpneq_qf[simp]: "qfree (simpneq t)" by (simp add: simpneq_def Let_def split_def)
+
+lemma simpfm_qf[simp]: "qfree p \<Longrightarrow> qfree (simpfm p)"
+by (induct p rule: simpfm.induct, auto simp add: disj_qf imp_qf iff_qf conj_qf not_qf Let_def)
+
+lemma disj_lin: "islin p \<Longrightarrow> islin q \<Longrightarrow> islin (disj p q)" by (simp add: disj_def)
+lemma conj_lin: "islin p \<Longrightarrow> islin q \<Longrightarrow> islin (conj p q)" by (simp add: conj_def)
+
+lemma   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "qfree p \<Longrightarrow> islin (simpfm p)" 
+  apply (induct p rule: simpfm.induct)
+  apply (simp_all add: conj_lin disj_lin)
+  done
+
+consts prep :: "fm \<Rightarrow> fm"
+recdef prep "measure fmsize"
+  "prep (E T) = T"
+  "prep (E F) = F"
+  "prep (E (Or p q)) = disj (prep (E p)) (prep (E q))"
+  "prep (E (Imp p q)) = disj (prep (E (NOT p))) (prep (E q))"
+  "prep (E (Iff p q)) = disj (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))" 
+  "prep (E (NOT (And p q))) = disj (prep (E (NOT p))) (prep (E(NOT q)))"
+  "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))"
+  "prep (E (NOT (Iff p q))) = disj (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))"
+  "prep (E p) = E (prep p)"
+  "prep (A (And p q)) = conj (prep (A p)) (prep (A q))"
+  "prep (A p) = prep (NOT (E (NOT p)))"
+  "prep (NOT (NOT p)) = prep p"
+  "prep (NOT (And p q)) = disj (prep (NOT p)) (prep (NOT q))"
+  "prep (NOT (A p)) = prep (E (NOT p))"
+  "prep (NOT (Or p q)) = conj (prep (NOT p)) (prep (NOT q))"
+  "prep (NOT (Imp p q)) = conj (prep p) (prep (NOT q))"
+  "prep (NOT (Iff p q)) = disj (prep (And p (NOT q))) (prep (And (NOT p) q))"
+  "prep (NOT p) = not (prep p)"
+  "prep (Or p q) = disj (prep p) (prep q)"
+  "prep (And p q) = conj (prep p) (prep q)"
+  "prep (Imp p q) = prep (Or (NOT p) q)"
+  "prep (Iff p q) = disj (prep (And p q)) (prep (And (NOT p) (NOT q)))"
+  "prep p = p"
+(hints simp add: fmsize_pos)
+lemma prep: "Ifm vs bs (prep p) = Ifm vs bs p"
+by (induct p arbitrary: bs rule: prep.induct, auto)
+
+
+
+  (* Generic quantifier elimination *)
+consts qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"
+recdef qelim "measure fmsize"
+  "qelim (E p) = (\<lambda> qe. DJ (CJNB qe) (qelim p qe))"
+  "qelim (A p) = (\<lambda> qe. not (qe ((qelim (NOT p) qe))))"
+  "qelim (NOT p) = (\<lambda> qe. not (qelim p qe))"
+  "qelim (And p q) = (\<lambda> qe. conj (qelim p qe) (qelim q qe))" 
+  "qelim (Or  p q) = (\<lambda> qe. disj (qelim p qe) (qelim q qe))" 
+  "qelim (Imp p q) = (\<lambda> qe. imp (qelim p qe) (qelim q qe))"
+  "qelim (Iff p q) = (\<lambda> qe. iff (qelim p qe) (qelim q qe))"
+  "qelim p = (\<lambda> y. simpfm p)"
+
+
+lemma qelim:
+  assumes qe_inv: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm vs bs (qe p) = Ifm vs bs (E p))"
+  shows "\<And> bs. qfree (qelim p qe) \<and> (Ifm vs bs (qelim p qe) = Ifm vs bs p)"
+using qe_inv DJ_qe[OF CJNB_qe[OF qe_inv]]
+by (induct p rule: qelim.induct) auto
+
+subsection{* Core Procedure *}
+
+consts 
+  plusinf:: "fm \<Rightarrow> fm" (* Virtual substitution of +\<infinity>*)
+  minusinf:: "fm \<Rightarrow> fm" (* Virtual substitution of -\<infinity>*)
+recdef minusinf "measure size"
+  "minusinf (And p q) = conj (minusinf p) (minusinf q)" 
+  "minusinf (Or p q) = disj (minusinf p) (minusinf q)" 
+  "minusinf (Eq  (CNP 0 c e)) = conj (eq (CP c)) (eq e)"
+  "minusinf (NEq (CNP 0 c e)) = disj (not (eq e)) (not (eq (CP c)))"
+  "minusinf (Lt  (CNP 0 c e)) = disj (conj (eq (CP c)) (lt e)) (lt (CP (~\<^sub>p c)))"
+  "minusinf (Le  (CNP 0 c e)) = disj (conj (eq (CP c)) (le e)) (lt (CP (~\<^sub>p c)))"
+  "minusinf p = p"
+
+recdef plusinf "measure size"
+  "plusinf (And p q) = conj (plusinf p) (plusinf q)" 
+  "plusinf (Or p q) = disj (plusinf p) (plusinf q)" 
+  "plusinf (Eq  (CNP 0 c e)) = conj (eq (CP c)) (eq e)"
+  "plusinf (NEq (CNP 0 c e)) = disj (not (eq e)) (not (eq (CP c)))"
+  "plusinf (Lt  (CNP 0 c e)) = disj (conj (eq (CP c)) (lt e)) (lt (CP c))"
+  "plusinf (Le  (CNP 0 c e)) = disj (conj (eq (CP c)) (le e)) (lt (CP c))"
+  "plusinf p = p"
+
+lemma minusinf_inf: assumes lp:"islin p"
+  shows "\<exists>z. \<forall>x < z. Ifm vs (x#bs) (minusinf p) \<longleftrightarrow> Ifm vs (x#bs) p"
+  using lp
+proof (induct p rule: minusinf.induct)
+  case 1 thus ?case by (auto,rule_tac x="min z za" in exI, auto)
+next
+  case 2 thus ?case by (auto,rule_tac x="min z za" in exI, auto)
+next
+  case (3 c e) hence nbe: "tmbound0 e" by simp
+  from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all
+  note eqs = eq[OF nc(1), where ?'a = 'a] eq[OF nc(2), where ?'a = 'a]
+  let ?c = "Ipoly vs c"
+  let ?e = "Itm vs (y#bs) e"
+  have "?c=0 \<or> ?c > 0 \<or> ?c < 0" by arith
+  moreover {assume "?c = 0" hence ?case 
+      using eq[OF nc(2), of vs] eq[OF nc(1), of vs] by auto}
+  moreover {assume cp: "?c > 0"
+    {fix x assume xz: "x < -?e / ?c" hence "?c * x < - ?e"
+	using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+      hence "?c * x + ?e < 0" by simp
+      hence "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Eq (CNP 0 c e)))"
+	using eqs tmbound0_I[OF nbe, where b="y" and b'="x" and vs=vs and bs=bs] by auto} hence ?case by auto}
+  moreover {assume cp: "?c < 0"
+    {fix x assume xz: "x < -?e / ?c" hence "?c * x > - ?e"
+	using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+      hence "?c * x + ?e > 0" by simp
+      hence "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Eq (CNP 0 c e)))"
+	using tmbound0_I[OF nbe, where b="y" and b'="x"] eqs by auto} hence ?case by auto}
+  ultimately show ?case by blast
+next
+  case (4 c e)  hence nbe: "tmbound0 e" by simp
+  from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all
+  note eqs = eq[OF nc(1), where ?'a = 'a] eq[OF nc(2), where ?'a = 'a]
+  let ?c = "Ipoly vs c"
+  let ?e = "Itm vs (y#bs) e"
+  have "?c=0 \<or> ?c > 0 \<or> ?c < 0" by arith
+  moreover {assume "?c = 0" hence ?case using eqs by auto}
+  moreover {assume cp: "?c > 0"
+    {fix x assume xz: "x < -?e / ?c" hence "?c * x < - ?e"
+	using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+      hence "?c * x + ?e < 0" by simp
+      hence "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (NEq (CNP 0 c e)))"
+	using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto} hence ?case by auto}
+  moreover {assume cp: "?c < 0"
+    {fix x assume xz: "x < -?e / ?c" hence "?c * x > - ?e"
+	using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+      hence "?c * x + ?e > 0" by simp
+      hence "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (NEq (CNP 0 c e)))"
+	using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto} hence ?case by auto}
+  ultimately show ?case by blast
+next
+  case (5 c e)  hence nbe: "tmbound0 e" by simp
+  from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all
+  hence nc': "allpolys isnpoly (CP (~\<^sub>p c))" by (simp add: polyneg_norm)
+  note eqs = lt[OF nc', where ?'a = 'a] eq [OF nc(1), where ?'a = 'a] lt[OF nc(2), where ?'a = 'a]
+  let ?c = "Ipoly vs c"
+  let ?e = "Itm vs (y#bs) e"
+  have "?c=0 \<or> ?c > 0 \<or> ?c < 0" by arith
+  moreover {assume "?c = 0" hence ?case using eqs by auto}
+  moreover {assume cp: "?c > 0"
+    {fix x assume xz: "x < -?e / ?c" hence "?c * x < - ?e"
+	using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+      hence "?c * x + ?e < 0" by simp
+      hence "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Lt (CNP 0 c e)))"
+	using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
+  moreover {assume cp: "?c < 0"
+    {fix x assume xz: "x < -?e / ?c" hence "?c * x > - ?e"
+	using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+      hence "?c * x + ?e > 0" by simp
+      hence "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Lt (CNP 0 c e)))"
+	using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] cp by auto} hence ?case by auto}
+  ultimately show ?case by blast
+next
+  case (6 c e)  hence nbe: "tmbound0 e" by simp
+  from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all
+  hence nc': "allpolys isnpoly (CP (~\<^sub>p c))" by (simp add: polyneg_norm)
+  note eqs = lt[OF nc', where ?'a = 'a] eq [OF nc(1), where ?'a = 'a] le[OF nc(2), where ?'a = 'a]
+  let ?c = "Ipoly vs c"
+  let ?e = "Itm vs (y#bs) e"
+  have "?c=0 \<or> ?c > 0 \<or> ?c < 0" by arith
+  moreover {assume "?c = 0" hence ?case using eqs by auto}
+  moreover {assume cp: "?c > 0"
+    {fix x assume xz: "x < -?e / ?c" hence "?c * x < - ?e"
+	using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+      hence "?c * x + ?e < 0" by simp
+      hence "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Le (CNP 0 c e)))"
+	using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
+  moreover {assume cp: "?c < 0"
+    {fix x assume xz: "x < -?e / ?c" hence "?c * x > - ?e"
+	using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+      hence "?c * x + ?e > 0" by simp
+      hence "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Le (CNP 0 c e)))"
+	using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
+  ultimately show ?case by blast
+qed (auto)
+
+lemma plusinf_inf: assumes lp:"islin p"
+  shows "\<exists>z. \<forall>x > z. Ifm vs (x#bs) (plusinf p) \<longleftrightarrow> Ifm vs (x#bs) p"
+  using lp
+proof (induct p rule: plusinf.induct)
+  case 1 thus ?case by (auto,rule_tac x="max z za" in exI, auto)
+next
+  case 2 thus ?case by (auto,rule_tac x="max z za" in exI, auto)
+next
+  case (3 c e) hence nbe: "tmbound0 e" by simp
+  from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all
+  note eqs = eq[OF nc(1), where ?'a = 'a] eq[OF nc(2), where ?'a = 'a]
+  let ?c = "Ipoly vs c"
+  let ?e = "Itm vs (y#bs) e"
+  have "?c=0 \<or> ?c > 0 \<or> ?c < 0" by arith
+  moreover {assume "?c = 0" hence ?case 
+      using eq[OF nc(2), of vs] eq[OF nc(1), of vs] by auto}
+  moreover {assume cp: "?c > 0"
+    {fix x assume xz: "x > -?e / ?c" hence "?c * x > - ?e" 
+	using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+      hence "?c * x + ?e > 0" by simp
+      hence "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Eq (CNP 0 c e)))"
+	using eqs tmbound0_I[OF nbe, where b="y" and b'="x" and vs=vs and bs=bs] by auto} hence ?case by auto}
+  moreover {assume cp: "?c < 0"
+    {fix x assume xz: "x > -?e / ?c" hence "?c * x < - ?e"
+	using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+      hence "?c * x + ?e < 0" by simp
+      hence "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Eq (CNP 0 c e)))"
+	using tmbound0_I[OF nbe, where b="y" and b'="x"] eqs by auto} hence ?case by auto}
+  ultimately show ?case by blast
+next
+  case (4 c e)  hence nbe: "tmbound0 e" by simp
+  from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all
+  note eqs = eq[OF nc(1), where ?'a = 'a] eq[OF nc(2), where ?'a = 'a]
+  let ?c = "Ipoly vs c"
+  let ?e = "Itm vs (y#bs) e"
+  have "?c=0 \<or> ?c > 0 \<or> ?c < 0" by arith
+  moreover {assume "?c = 0" hence ?case using eqs by auto}
+  moreover {assume cp: "?c > 0"
+    {fix x assume xz: "x > -?e / ?c" hence "?c * x > - ?e"
+	using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+      hence "?c * x + ?e > 0" by simp
+      hence "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (NEq (CNP 0 c e)))"
+	using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto} hence ?case by auto}
+  moreover {assume cp: "?c < 0"
+    {fix x assume xz: "x > -?e / ?c" hence "?c * x < - ?e"
+	using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+      hence "?c * x + ?e < 0" by simp
+      hence "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (NEq (CNP 0 c e)))"
+	using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto} hence ?case by auto}
+  ultimately show ?case by blast
+next
+  case (5 c e)  hence nbe: "tmbound0 e" by simp
+  from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all
+  hence nc': "allpolys isnpoly (CP (~\<^sub>p c))" by (simp add: polyneg_norm)
+  note eqs = lt[OF nc(1), where ?'a = 'a] lt[OF nc', where ?'a = 'a] eq [OF nc(1), where ?'a = 'a] lt[OF nc(2), where ?'a = 'a]
+  let ?c = "Ipoly vs c"
+  let ?e = "Itm vs (y#bs) e"
+  have "?c=0 \<or> ?c > 0 \<or> ?c < 0" by arith
+  moreover {assume "?c = 0" hence ?case using eqs by auto}
+  moreover {assume cp: "?c > 0"
+    {fix x assume xz: "x > -?e / ?c" hence "?c * x > - ?e"
+	using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+      hence "?c * x + ?e > 0" by simp
+      hence "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Lt (CNP 0 c e)))"
+	using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
+  moreover {assume cp: "?c < 0"
+    {fix x assume xz: "x > -?e / ?c" hence "?c * x < - ?e"
+	using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+      hence "?c * x + ?e < 0" by simp
+      hence "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Lt (CNP 0 c e)))"
+	using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] cp by auto} hence ?case by auto}
+  ultimately show ?case by blast
+next
+  case (6 c e)  hence nbe: "tmbound0 e" by simp
+  from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all
+  hence nc': "allpolys isnpoly (CP (~\<^sub>p c))" by (simp add: polyneg_norm)
+  note eqs = lt[OF nc(1), where ?'a = 'a] eq [OF nc(1), where ?'a = 'a] le[OF nc(2), where ?'a = 'a]
+  let ?c = "Ipoly vs c"
+  let ?e = "Itm vs (y#bs) e"
+  have "?c=0 \<or> ?c > 0 \<or> ?c < 0" by arith
+  moreover {assume "?c = 0" hence ?case using eqs by auto}
+  moreover {assume cp: "?c > 0"
+    {fix x assume xz: "x > -?e / ?c" hence "?c * x > - ?e"
+	using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+      hence "?c * x + ?e > 0" by simp
+      hence "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Le (CNP 0 c e)))"
+	using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
+  moreover {assume cp: "?c < 0"
+    {fix x assume xz: "x > -?e / ?c" hence "?c * x < - ?e"
+	using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+      hence "?c * x + ?e < 0" by simp
+      hence "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Le (CNP 0 c e)))"
+	using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
+  ultimately show ?case by blast
+qed (auto)
+
+lemma minusinf_nb: "islin p \<Longrightarrow> bound0 (minusinf p)" 
+  by (induct p rule: minusinf.induct, auto simp add: eq_nb lt_nb le_nb)
+lemma plusinf_nb: "islin p \<Longrightarrow> bound0 (plusinf p)" 
+  by (induct p rule: minusinf.induct, auto simp add: eq_nb lt_nb le_nb)
+
+lemma minusinf_ex: assumes lp: "islin p" and ex: "Ifm vs (x#bs) (minusinf p)"
+  shows "\<exists>x. Ifm vs (x#bs) p"
+proof-
+  from bound0_I [OF minusinf_nb[OF lp], where b="a" and bs ="bs"] ex
+  have th: "\<forall> x. Ifm vs (x#bs) (minusinf p)" by auto
+  from minusinf_inf[OF lp, where bs="bs"] 
+  obtain z where z_def: "\<forall>x<z. Ifm vs (x # bs) (minusinf p) = Ifm vs (x # bs) p" by blast
+  from th have "Ifm vs ((z - 1)#bs) (minusinf p)" by simp
+  moreover have "z - 1 < z" by simp
+  ultimately show ?thesis using z_def by auto
+qed
+
+lemma plusinf_ex: assumes lp: "islin p" and ex: "Ifm vs (x#bs) (plusinf p)"
+  shows "\<exists>x. Ifm vs (x#bs) p"
+proof-
+  from bound0_I [OF plusinf_nb[OF lp], where b="a" and bs ="bs"] ex
+  have th: "\<forall> x. Ifm vs (x#bs) (plusinf p)" by auto
+  from plusinf_inf[OF lp, where bs="bs"] 
+  obtain z where z_def: "\<forall>x>z. Ifm vs (x # bs) (plusinf p) = Ifm vs (x # bs) p" by blast
+  from th have "Ifm vs ((z + 1)#bs) (plusinf p)" by simp
+  moreover have "z + 1 > z" by simp
+  ultimately show ?thesis using z_def by auto
+qed
+
+fun uset :: "fm \<Rightarrow> (poly \<times> tm) list" where
+  "uset (And p q) = uset p @ uset q"
+| "uset (Or p q) = uset p @ uset q"
+| "uset (Eq (CNP 0 a e))  = [(a,e)]"
+| "uset (Le (CNP 0 a e))  = [(a,e)]"
+| "uset (Lt (CNP 0 a e))  = [(a,e)]"
+| "uset (NEq (CNP 0 a e)) = [(a,e)]"
+| "uset p = []"
+
+lemma uset_l:
+  assumes lp: "islin p"
+  shows "\<forall> (c,s) \<in> set (uset p). isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s"
+using lp by(induct p rule: uset.induct,auto)
+
+lemma minusinf_uset0:
+  assumes lp: "islin p"
+  and nmi: "\<not> (Ifm vs (x#bs) (minusinf p))"
+  and ex: "Ifm vs (x#bs) p" (is "?I x p")
+  shows "\<exists> (c,s) \<in> set (uset p). x \<ge> - Itm vs (x#bs) s / Ipoly vs c" 
+proof-
+  have "\<exists> (c,s) \<in> set (uset p). (Ipoly vs c < 0 \<and> Ipoly vs c * x \<le> - Itm vs (x#bs) s) \<or>  (Ipoly vs c > 0 \<and> Ipoly vs c * x \<ge> - Itm vs (x#bs) s)" 
+    using lp nmi ex
+    apply (induct p rule: minusinf.induct, auto simp add: eq le lt nth_pos2 polyneg_norm)
+    apply (auto simp add: linorder_not_less order_le_less)
+    done 
+  then obtain c s where csU: "(c,s) \<in> set (uset p)" and x: "(Ipoly vs c < 0 \<and> Ipoly vs c * x \<le> - Itm vs (x#bs) s) \<or>  (Ipoly vs c > 0 \<and> Ipoly vs c * x \<ge> - Itm vs (x#bs) s)" by blast
+  hence "x \<ge> (- Itm vs (x#bs) s) / Ipoly vs c"
+    using divide_le_eq[of "- Itm vs (x#bs) s" "Ipoly vs c" x]
+    by (auto simp add: mult_commute del: divide_minus_left)
+  thus ?thesis using csU by auto
+qed
+
+lemma minusinf_uset:
+  assumes lp: "islin p"
+  and nmi: "\<not> (Ifm vs (a#bs) (minusinf p))"
+  and ex: "Ifm vs (x#bs) p" (is "?I x p")
+  shows "\<exists> (c,s) \<in> set (uset p). x \<ge> - Itm vs (a#bs) s / Ipoly vs c" 
+proof-
+  from nmi have nmi': "\<not> (Ifm vs (x#bs) (minusinf p))" 
+    by (simp add: bound0_I[OF minusinf_nb[OF lp], where b=x and b'=a])
+  from minusinf_uset0[OF lp nmi' ex] 
+  obtain c s where csU: "(c,s) \<in> set (uset p)" and th: "x \<ge> - Itm vs (x#bs) s / Ipoly vs c" by blast
+  from uset_l[OF lp, rule_format, OF csU] have nb: "tmbound0 s" by simp
+  from th tmbound0_I[OF nb, of vs x bs a] csU show ?thesis by auto
+qed
+
+
+lemma plusinf_uset0:
+  assumes lp: "islin p"
+  and nmi: "\<not> (Ifm vs (x#bs) (plusinf p))"
+  and ex: "Ifm vs (x#bs) p" (is "?I x p")
+  shows "\<exists> (c,s) \<in> set (uset p). x \<le> - Itm vs (x#bs) s / Ipoly vs c" 
+proof-
+  have "\<exists> (c,s) \<in> set (uset p). (Ipoly vs c < 0 \<and> Ipoly vs c * x \<ge> - Itm vs (x#bs) s) \<or>  (Ipoly vs c > 0 \<and> Ipoly vs c * x \<le> - Itm vs (x#bs) s)" 
+    using lp nmi ex
+    apply (induct p rule: minusinf.induct, auto simp add: eq le lt nth_pos2 polyneg_norm)
+    apply (auto simp add: linorder_not_less order_le_less)
+    done 
+  then obtain c s where csU: "(c,s) \<in> set (uset p)" and x: "(Ipoly vs c < 0 \<and> Ipoly vs c * x \<ge> - Itm vs (x#bs) s) \<or>  (Ipoly vs c > 0 \<and> Ipoly vs c * x \<le> - Itm vs (x#bs) s)" by blast
+  hence "x \<le> (- Itm vs (x#bs) s) / Ipoly vs c"
+    using le_divide_eq[of x "- Itm vs (x#bs) s" "Ipoly vs c"]
+    by (auto simp add: mult_commute del: divide_minus_left)
+  thus ?thesis using csU by auto
+qed
+
+lemma plusinf_uset:
+  assumes lp: "islin p"
+  and nmi: "\<not> (Ifm vs (a#bs) (plusinf p))"
+  and ex: "Ifm vs (x#bs) p" (is "?I x p")
+  shows "\<exists> (c,s) \<in> set (uset p). x \<le> - Itm vs (a#bs) s / Ipoly vs c" 
+proof-
+  from nmi have nmi': "\<not> (Ifm vs (x#bs) (plusinf p))" 
+    by (simp add: bound0_I[OF plusinf_nb[OF lp], where b=x and b'=a])
+  from plusinf_uset0[OF lp nmi' ex] 
+  obtain c s where csU: "(c,s) \<in> set (uset p)" and th: "x \<le> - Itm vs (x#bs) s / Ipoly vs c" by blast
+  from uset_l[OF lp, rule_format, OF csU] have nb: "tmbound0 s" by simp
+  from th tmbound0_I[OF nb, of vs x bs a] csU show ?thesis by auto
+qed
+
+lemma lin_dense: 
+  assumes lp: "islin p"
+  and noS: "\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> (\<lambda> (c,t). - Itm vs (x#bs) t / Ipoly vs c) ` set (uset p)" 
+  (is "\<forall> t. _ \<and> _ \<longrightarrow> t \<notin> (\<lambda> (c,t). - ?Nt x t / ?N c) ` ?U p")
+  and lx: "l < x" and xu:"x < u" and px:" Ifm vs (x#bs) p"
+  and ly: "l < y" and yu: "y < u"
+  shows "Ifm vs (y#bs) p"
+using lp px noS
+proof (induct p rule: islin.induct) 
+  case (5 c s)
+  from "5.prems" 
+  have lin: "isnpoly c" "c \<noteq> 0\<^sub>p" "tmbound0 s" "allpolys isnpoly s"
+    and px: "Ifm vs (x # bs) (Lt (CNP 0 c s))"
+    and noS: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<noteq> - Itm vs (x # bs) s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" by simp_all
+  from ly yu noS have yne: "y \<noteq> - ?Nt x s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" by simp
+  hence ycs: "y < - ?Nt x s / ?N c \<or> y > -?Nt x s / ?N c" by auto
+  have ccs: "?N c = 0 \<or> ?N c < 0 \<or> ?N c > 0" by dlo
+  moreover
+  {assume "?N c = 0" hence ?case using px by (simp add: tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"])}
+  moreover
+  {assume c: "?N c > 0"
+      from px pos_less_divide_eq[OF c, where a="x" and b="-?Nt x s"]  
+      have px': "x < - ?Nt x s / ?N c" 
+	by (auto simp add: not_less ring_simps) 
+    {assume y: "y < - ?Nt x s / ?N c" 
+      hence "y * ?N c < - ?Nt x s"
+	by (simp add: pos_less_divide_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
+      hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps)
+      hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
+    moreover
+    {assume y: "y > -?Nt x s / ?N c" 
+      with yu have eu: "u > - ?Nt x s / ?N c" by auto
+      with noS ly yu have th: "- ?Nt x s / ?N c \<le> l" by (cases "- ?Nt x s / ?N c > l", auto)
+      with lx px' have "False" by simp  hence ?case by simp }
+    ultimately have ?case using ycs by blast
+  }
+  moreover
+  {assume c: "?N c < 0"
+      from px neg_divide_less_eq[OF c, where a="x" and b="-?Nt x s"]  
+      have px': "x > - ?Nt x s / ?N c" 
+	by (auto simp add: not_less ring_simps) 
+    {assume y: "y > - ?Nt x s / ?N c" 
+      hence "y * ?N c < - ?Nt x s"
+	by (simp add: neg_divide_less_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
+      hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps)
+      hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
+    moreover
+    {assume y: "y < -?Nt x s / ?N c" 
+      with ly have eu: "l < - ?Nt x s / ?N c" by auto
+      with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u" by (cases "- ?Nt x s / ?N c < u", auto)
+      with xu px' have "False" by simp  hence ?case by simp }
+    ultimately have ?case using ycs by blast
+  }
+  ultimately show ?case by blast
+next
+  case (6 c s)
+  from "6.prems" 
+  have lin: "isnpoly c" "c \<noteq> 0\<^sub>p" "tmbound0 s" "allpolys isnpoly s"
+    and px: "Ifm vs (x # bs) (Le (CNP 0 c s))"
+    and noS: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<noteq> - Itm vs (x # bs) s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" by simp_all
+  from ly yu noS have yne: "y \<noteq> - ?Nt x s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" by simp
+  hence ycs: "y < - ?Nt x s / ?N c \<or> y > -?Nt x s / ?N c" by auto
+  have ccs: "?N c = 0 \<or> ?N c < 0 \<or> ?N c > 0" by dlo
+  moreover
+  {assume "?N c = 0" hence ?case using px by (simp add: tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"])}
+  moreover
+  {assume c: "?N c > 0"
+      from px pos_le_divide_eq[OF c, where a="x" and b="-?Nt x s"]  
+      have px': "x <= - ?Nt x s / ?N c" by (simp add: not_less ring_simps) 
+    {assume y: "y < - ?Nt x s / ?N c" 
+      hence "y * ?N c < - ?Nt x s"
+	by (simp add: pos_less_divide_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
+      hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps)
+      hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
+    moreover
+    {assume y: "y > -?Nt x s / ?N c" 
+      with yu have eu: "u > - ?Nt x s / ?N c" by auto
+      with noS ly yu have th: "- ?Nt x s / ?N c \<le> l" by (cases "- ?Nt x s / ?N c > l", auto)
+      with lx px' have "False" by simp  hence ?case by simp }
+    ultimately have ?case using ycs by blast
+  }
+  moreover
+  {assume c: "?N c < 0"
+      from px neg_divide_le_eq[OF c, where a="x" and b="-?Nt x s"]  
+      have px': "x >= - ?Nt x s / ?N c" by (simp add: ring_simps) 
+    {assume y: "y > - ?Nt x s / ?N c" 
+      hence "y * ?N c < - ?Nt x s"
+	by (simp add: neg_divide_less_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
+      hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps)
+      hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
+    moreover
+    {assume y: "y < -?Nt x s / ?N c" 
+      with ly have eu: "l < - ?Nt x s / ?N c" by auto
+      with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u" by (cases "- ?Nt x s / ?N c < u", auto)
+      with xu px' have "False" by simp  hence ?case by simp }
+    ultimately have ?case using ycs by blast
+  }
+  ultimately show ?case by blast
+next
+    case (3 c s)
+  from "3.prems" 
+  have lin: "isnpoly c" "c \<noteq> 0\<^sub>p" "tmbound0 s" "allpolys isnpoly s"
+    and px: "Ifm vs (x # bs) (Eq (CNP 0 c s))"
+    and noS: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<noteq> - Itm vs (x # bs) s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" by simp_all
+  from ly yu noS have yne: "y \<noteq> - ?Nt x s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" by simp
+  hence ycs: "y < - ?Nt x s / ?N c \<or> y > -?Nt x s / ?N c" by auto
+  have ccs: "?N c = 0 \<or> ?N c < 0 \<or> ?N c > 0" by dlo
+  moreover
+  {assume "?N c = 0" hence ?case using px by (simp add: tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"])}
+  moreover
+  {assume c: "?N c > 0" hence cnz: "?N c \<noteq> 0" by simp
+    from px eq_divide_eq[of "x" "-?Nt x s" "?N c"]  cnz
+    have px': "x = - ?Nt x s / ?N c" by (simp add: ring_simps)
+    {assume y: "y < -?Nt x s / ?N c" 
+      with ly have eu: "l < - ?Nt x s / ?N c" by auto
+      with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u" by (cases "- ?Nt x s / ?N c < u", auto)
+      with xu px' have "False" by simp  hence ?case by simp }
+    moreover
+    {assume y: "y > -?Nt x s / ?N c" 
+      with yu have eu: "u > - ?Nt x s / ?N c" by auto
+      with noS ly yu have th: "- ?Nt x s / ?N c \<le> l" by (cases "- ?Nt x s / ?N c > l", auto)
+      with lx px' have "False" by simp  hence ?case by simp }
+    ultimately have ?case using ycs by blast
+  }
+  moreover
+  {assume c: "?N c < 0" hence cnz: "?N c \<noteq> 0" by simp
+    from px eq_divide_eq[of "x" "-?Nt x s" "?N c"]  cnz
+    have px': "x = - ?Nt x s / ?N c" by (simp add: ring_simps)
+    {assume y: "y < -?Nt x s / ?N c" 
+      with ly have eu: "l < - ?Nt x s / ?N c" by auto
+      with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u" by (cases "- ?Nt x s / ?N c < u", auto)
+      with xu px' have "False" by simp  hence ?case by simp }
+    moreover
+    {assume y: "y > -?Nt x s / ?N c" 
+      with yu have eu: "u > - ?Nt x s / ?N c" by auto
+      with noS ly yu have th: "- ?Nt x s / ?N c \<le> l" by (cases "- ?Nt x s / ?N c > l", auto)
+      with lx px' have "False" by simp  hence ?case by simp }
+    ultimately have ?case using ycs by blast
+  }
+  ultimately show ?case by blast
+next
+    case (4 c s)
+  from "4.prems" 
+  have lin: "isnpoly c" "c \<noteq> 0\<^sub>p" "tmbound0 s" "allpolys isnpoly s"
+    and px: "Ifm vs (x # bs) (NEq (CNP 0 c s))"
+    and noS: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<noteq> - Itm vs (x # bs) s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" by simp_all
+  from ly yu noS have yne: "y \<noteq> - ?Nt x s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" by simp
+  hence ycs: "y < - ?Nt x s / ?N c \<or> y > -?Nt x s / ?N c" by auto
+  have ccs: "?N c = 0 \<or> ?N c \<noteq> 0" by dlo
+  moreover
+  {assume "?N c = 0" hence ?case using px by (simp add: tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"])}
+  moreover
+  {assume c: "?N c \<noteq> 0"
+    from yne c eq_divide_eq[of "y" "- ?Nt x s" "?N c"] have ?case
+      by (simp add: ring_simps tmbound0_I[OF lin(3), of vs x bs y] sum_eq[symmetric]) }
+  ultimately show ?case by blast
+qed (auto simp add: nth_pos2 tmbound0_I[where vs=vs and bs="bs" and b="y" and b'="x"] bound0_I[where vs=vs and bs="bs" and b="y" and b'="x"])
+
+lemma one_plus_one_pos[simp]: "(1::'a::{ordered_field}) + 1 > 0"
+proof-
+  have op: "(1::'a) > 0" by simp
+  from add_pos_pos[OF op op] show ?thesis . 
+qed
+
+lemma one_plus_one_nonzero[simp]: "(1::'a::{ordered_field}) + 1 \<noteq> 0" 
+  using one_plus_one_pos[where ?'a = 'a] by (simp add: less_le) 
+
+lemma half_sum_eq: "(u + u) / (1+1) = (u::'a::{ordered_field})" 
+proof-
+  have "(u + u) = (1 + 1) * u" by (simp add: ring_simps)
+  hence "(u + u) / (1+1) = (1 + 1)*u / (1 + 1)" by simp
+  with nonzero_mult_divide_cancel_left[OF one_plus_one_nonzero, of u] show ?thesis by simp
+qed
+
+lemma inf_uset:
+  assumes lp: "islin p"
+  and nmi: "\<not> (Ifm vs (x#bs) (minusinf p))" (is "\<not> (Ifm vs (x#bs) (?M p))")
+  and npi: "\<not> (Ifm vs (x#bs) (plusinf p))" (is "\<not> (Ifm vs (x#bs) (?P p))")
+  and ex: "\<exists> x.  Ifm vs (x#bs) p" (is "\<exists> x. ?I x p")
+  shows "\<exists> (c,t) \<in> set (uset p). \<exists> (d,s) \<in> set (uset p). ?I ((- Itm vs (x#bs) t / Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) / (1 + 1)) p" 
+proof-
+  let ?Nt = "\<lambda> x t. Itm vs (x#bs) t"
+  let ?N = "Ipoly vs"
+  let ?U = "set (uset p)"
+  from ex obtain a where pa: "?I a p" by blast
+  from bound0_I[OF minusinf_nb[OF lp], where bs="bs" and b="x" and b'="a"] nmi
+  have nmi': "\<not> (?I a (?M p))" by simp
+  from bound0_I[OF plusinf_nb[OF lp], where bs="bs" and b="x" and b'="a"] npi
+  have npi': "\<not> (?I a (?P p))" by simp
+  have "\<exists> (c,t) \<in> set (uset p). \<exists> (d,s) \<in> set (uset p). ?I ((- ?Nt a t/?N c + - ?Nt a s /?N d) / (1 + 1)) p"
+  proof-
+    let ?M = "(\<lambda> (c,t). - ?Nt a t / ?N c) ` ?U"
+    have fM: "finite ?M" by auto
+    from minusinf_uset[OF lp nmi pa] plusinf_uset[OF lp npi pa] 
+    have "\<exists> (c,t) \<in> set (uset p). \<exists> (d,s) \<in> set (uset p). a \<le> - ?Nt x t / ?N c \<and> a \<ge> - ?Nt x s / ?N d" by blast
+    then obtain "c" "t" "d" "s" where 
+      ctU: "(c,t) \<in> ?U" and dsU: "(d,s) \<in> ?U" 
+      and xs1: "a \<le> - ?Nt x s / ?N d" and tx1: "a \<ge> - ?Nt x t / ?N c" by blast
+    from uset_l[OF lp] ctU dsU tmbound0_I[where bs="bs" and b="x" and b'="a"] xs1 tx1 
+    have xs: "a \<le> - ?Nt a s / ?N d" and tx: "a \<ge> - ?Nt a t / ?N c" by auto
+    from ctU have Mne: "?M \<noteq> {}" by auto
+    hence Une: "?U \<noteq> {}" by simp
+    let ?l = "Min ?M"
+    let ?u = "Max ?M"
+    have linM: "?l \<in> ?M" using fM Mne by simp
+    have uinM: "?u \<in> ?M" using fM Mne by simp
+    have ctM: "- ?Nt a t / ?N c \<in> ?M" using ctU by auto
+    have dsM: "- ?Nt a s / ?N d \<in> ?M" using dsU by auto 
+    have lM: "\<forall> t\<in> ?M. ?l \<le> t" using Mne fM by auto
+    have Mu: "\<forall> t\<in> ?M. t \<le> ?u" using Mne fM by auto
+    have "?l \<le> - ?Nt a t / ?N c" using ctM Mne by simp hence lx: "?l \<le> a" using tx by simp
+    have "- ?Nt a s / ?N d \<le> ?u" using dsM Mne by simp hence xu: "a \<le> ?u" using xs by simp
+    from finite_set_intervals2[where P="\<lambda> x. ?I x p",OF pa lx xu linM uinM fM lM Mu]
+    have "(\<exists> s\<in> ?M. ?I s p) \<or> 
+      (\<exists> t1\<in> ?M. \<exists> t2 \<in> ?M. (\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M) \<and> t1 < a \<and> a < t2 \<and> ?I a p)" .
+    moreover {fix u assume um: "u\<in> ?M" and pu: "?I u p"
+      hence "\<exists> (nu,tu) \<in> ?U. u = - ?Nt a tu / ?N nu" by auto
+      then obtain "tu" "nu" where tuU: "(nu,tu) \<in> ?U" and tuu:"u= - ?Nt a tu / ?N nu" by blast
+      from half_sum_eq[of u] pu tuu 
+      have "?I (((- ?Nt a tu / ?N nu) + (- ?Nt a tu / ?N nu)) / (1 + 1)) p" by simp
+      with tuU have ?thesis by blast}
+    moreover{
+      assume "\<exists> t1\<in> ?M. \<exists> t2 \<in> ?M. (\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M) \<and> t1 < a \<and> a < t2 \<and> ?I a p"
+      then obtain t1 and t2 where t1M: "t1 \<in> ?M" and t2M: "t2\<in> ?M" 
+	and noM: "\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M" and t1x: "t1 < a" and xt2: "a < t2" and px: "?I a p"
+	by blast
+      from t1M have "\<exists> (t1n,t1u) \<in> ?U. t1 = - ?Nt a t1u / ?N t1n" by auto
+      then obtain "t1u" "t1n" where t1uU: "(t1n,t1u) \<in> ?U" and t1u: "t1 = - ?Nt a t1u / ?N t1n" by blast
+      from t2M have "\<exists> (t2n,t2u) \<in> ?U. t2 = - ?Nt a t2u / ?N t2n" by auto
+      then obtain "t2u" "t2n" where t2uU: "(t2n,t2u) \<in> ?U" and t2u: "t2 = - ?Nt a t2u / ?N t2n" by blast
+      from t1x xt2 have t1t2: "t1 < t2" by simp
+      let ?u = "(t1 + t2) / (1 + 1)"
+      from less_half_sum[OF t1t2] gt_half_sum[OF t1t2] have t1lu: "t1 < ?u" and ut2: "?u < t2" by auto
+      from lin_dense[OF lp noM t1x xt2 px t1lu ut2] have "?I ?u p" .
+      with t1uU t2uU t1u t2u have ?thesis by blast}
+    ultimately show ?thesis by blast
+  qed
+  then obtain "l" "n" "s"  "m" where lnU: "(n,l) \<in> ?U" and smU:"(m,s) \<in> ?U" 
+    and pu: "?I ((- ?Nt a l / ?N n + - ?Nt a s / ?N m) / (1 + 1)) p" by blast
+  from lnU smU uset_l[OF lp] have nbl: "tmbound0 l" and nbs: "tmbound0 s" by auto
+  from tmbound0_I[OF nbl, where bs="bs" and b="a" and b'="x"] 
+    tmbound0_I[OF nbs, where bs="bs" and b="a" and b'="x"] pu
+  have "?I ((- ?Nt x l / ?N n + - ?Nt x s / ?N m) / (1 + 1)) p" by simp
+  with lnU smU
+  show ?thesis by auto
+qed
+
+    (* The Ferrante - Rackoff Theorem *)
+
+theorem fr_eq: 
+  assumes lp: "islin p"
+  shows "(\<exists> x. Ifm vs (x#bs) p) = ((Ifm vs (x#bs) (minusinf p)) \<or> (Ifm vs (x#bs) (plusinf p)) \<or> (\<exists> (n,t) \<in> set (uset p). \<exists> (m,s) \<in> set (uset p). Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs n + - Itm vs (x#bs) s / Ipoly vs m) /(1 + 1))#bs) p))"
+  (is "(\<exists> x. ?I x p) = (?M \<or> ?P \<or> ?F)" is "?E = ?D")
+proof
+  assume px: "\<exists> x. ?I x p"
+  have "?M \<or> ?P \<or> (\<not> ?M \<and> \<not> ?P)" by blast
+  moreover {assume "?M \<or> ?P" hence "?D" by blast}
+  moreover {assume nmi: "\<not> ?M" and npi: "\<not> ?P"
+    from inf_uset[OF lp nmi npi] have "?F" using px by blast hence "?D" by blast}
+  ultimately show "?D" by blast
+next
+  assume "?D" 
+  moreover {assume m:"?M" from minusinf_ex[OF lp m] have "?E" .}
+  moreover {assume p: "?P" from plusinf_ex[OF lp p] have "?E" . }
+  moreover {assume f:"?F" hence "?E" by blast}
+  ultimately show "?E" by blast
+qed
+
+section{* First implementation : Naive by encoding all case splits locally *}
+definition "msubsteq c t d s a r = 
+  evaldjf (split conj) 
+  [(let cd = c *\<^sub>p d in (NEq (CP cd), Eq (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+   (conj (Eq (CP c)) (NEq (CP d)) , Eq (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+   (conj (NEq (CP c)) (Eq (CP d)) , Eq (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+   (conj (Eq (CP c)) (Eq (CP d)) , Eq r)]"
+
+lemma msubsteq_nb: assumes lp: "islin (Eq (CNP 0 a r))" and t: "tmbound0 t" and s: "tmbound0 s"
+  shows "bound0 (msubsteq c t d s a r)"
+proof-
+  have th: "\<forall>x\<in> set [(let cd = c *\<^sub>p d in (NEq (CP cd), Eq (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+   (conj (Eq (CP c)) (NEq (CP d)) , Eq (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+   (conj (NEq (CP c)) (Eq (CP d)) , Eq (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+   (conj (Eq (CP c)) (Eq (CP d)) , Eq r)]. bound0 (split conj x)"
+    using lp by (simp add: Let_def t s )
+  from evaldjf_bound0[OF th] show ?thesis by (simp add: msubsteq_def)
+qed
+
+lemma msubsteq: assumes lp: "islin (Eq (CNP 0 a r))"
+  shows "Ifm vs (x#bs) (msubsteq c t d s a r) = Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /(1 + 1))#bs) (Eq (CNP 0 a r))" (is "?lhs = ?rhs")
+proof-
+  let ?Nt = "\<lambda>(x::'a) t. Itm vs (x#bs) t"
+  let ?N = "\<lambda>p. Ipoly vs p"
+  let ?c = "?N c"
+  let ?d = "?N d"
+  let ?t = "?Nt x t"
+  let ?s = "?Nt x s"
+  let ?a = "?N a"
+  let ?r = "?Nt x r"
+  from lp have lin:"isnpoly a" "a \<noteq> 0\<^sub>p" "tmbound0 r" "allpolys isnpoly r" by simp_all
+  note r= tmbound0_I[OF lin(3), of vs _ bs x]
+  have cd_cs: "?c * ?d \<noteq> 0 \<or> (?c = 0 \<and> ?d = 0) \<or> (?c = 0 \<and> ?d \<noteq> 0) \<or> (?c \<noteq> 0 \<and> ?d = 0)" by auto
+  moreover
+  {assume c: "?c = 0" and d: "?d=0"
+    hence ?thesis  by (simp add: r[of 0] msubsteq_def Let_def evaldjf_ex)}
+  moreover 
+  {assume c: "?c = 0" and d: "?d\<noteq>0"
+    from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = -?s / ((1 + 1)*?d)" by simp
+    have "?rhs = Ifm vs (-?s / ((1 + 1)*?d) # bs) (Eq (CNP 0 a r))" by (simp only: th)
+    also have "\<dots> \<longleftrightarrow> ?a * (-?s / ((1 + 1)*?d)) + ?r = 0" by (simp add: r[of "- (Itm vs (x # bs) s / ((1 + 1) * \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>))"])
+    also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a * (-?s / ((1 + 1)*?d)) + ?r) = 0" 
+      using d mult_cancel_left[of "(1 + 1)*?d" "(?a * (-?s / ((1 + 1)*?d)) + ?r)" 0] by simp
+    also have "\<dots> \<longleftrightarrow> (- ?a * ?s) * ((1 + 1)*?d / ((1 + 1)*?d)) + (1 + 1)*?d*?r= 0"
+      by (simp add: ring_simps right_distrib[of "(1 + 1)*?d"] del: right_distrib)
+    
+    also have "\<dots> \<longleftrightarrow> - (?a * ?s) + (1 + 1)*?d*?r = 0" using d by simp 
+    finally have ?thesis using c d 
+      apply (simp add: r[of "- (Itm vs (x # bs) s / ((1 + 1) * \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>))"] msubsteq_def Let_def evaldjf_ex del: one_add_one_is_two)
+      apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+      apply simp
+      done}
+  moreover
+  {assume c: "?c \<noteq> 0" and d: "?d=0"
+    from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = -?t / ((1 + 1)*?c)" by simp
+    have "?rhs = Ifm vs (-?t / ((1 + 1)*?c) # bs) (Eq (CNP 0 a r))" by (simp only: th)
+    also have "\<dots> \<longleftrightarrow> ?a * (-?t / ((1 + 1)*?c)) + ?r = 0" by (simp add: r[of "- (?t/ ((1 + 1)* ?c))"])
+    also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a * (-?t / ((1 + 1)*?c)) + ?r) = 0" 
+      using c mult_cancel_left[of "(1 + 1)*?c" "(?a * (-?t / ((1 + 1)*?c)) + ?r)" 0] by simp
+    also have "\<dots> \<longleftrightarrow> (?a * -?t)* ((1 + 1)*?c) / ((1 + 1)*?c) + (1 + 1)*?c*?r= 0"
+      by (simp add: ring_simps right_distrib[of "(1 + 1)*?c"] del: right_distrib)
+    also have "\<dots> \<longleftrightarrow> - (?a * ?t) + (1 + 1)*?c*?r = 0" using c by simp 
+    finally have ?thesis using c d 
+      apply (simp add: r[of "- (?t/ ((1 + 1)*?c))"] msubsteq_def Let_def evaldjf_ex del: one_add_one_is_two)
+      apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+      apply simp
+      done }
+  moreover
+  {assume c: "?c \<noteq> 0" and d: "?d\<noteq>0" hence dc: "?c * ?d *(1 + 1) \<noteq> 0" by simp
+    from add_frac_eq[OF c d, of "- ?t" "- ?s"]
+    have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)" 
+      by (simp add: ring_simps)
+    have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Eq (CNP 0 a r))" by (simp only: th)
+    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r = 0" 
+      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
+    also have "\<dots> \<longleftrightarrow> ((1 + 1) * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r) =0 "
+      using c d mult_cancel_left[of "(1 + 1) * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r" 0] by simp
+    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + (1 + 1)*?c*?d*?r =0" 
+      using nonzero_mult_divide_cancel_left[OF dc] c d
+      by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+    finally  have ?thesis using c d 
+      apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubsteq_def Let_def evaldjf_ex ring_simps)
+      apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+      apply (simp add: ring_simps)
+      done }
+  ultimately show ?thesis by blast
+qed
+
+
+definition "msubstneq c t d s a r = 
+  evaldjf (split conj) 
+  [(let cd = c *\<^sub>p d in (NEq (CP cd), NEq (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+   (conj (Eq (CP c)) (NEq (CP d)) , NEq (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+   (conj (NEq (CP c)) (Eq (CP d)) , NEq (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+   (conj (Eq (CP c)) (Eq (CP d)) , NEq r)]"
+
+lemma msubstneq_nb: assumes lp: "islin (NEq (CNP 0 a r))" and t: "tmbound0 t" and s: "tmbound0 s"
+  shows "bound0 (msubstneq c t d s a r)"
+proof-
+  have th: "\<forall>x\<in> set [(let cd = c *\<^sub>p d in (NEq (CP cd), NEq (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))), 
+    (conj (Eq (CP c)) (NEq (CP d)) , NEq (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+    (conj (NEq (CP c)) (Eq (CP d)) , NEq (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+    (conj (Eq (CP c)) (Eq (CP d)) , NEq r)]. bound0 (split conj x)"
+    using lp by (simp add: Let_def t s )
+  from evaldjf_bound0[OF th] show ?thesis by (simp add: msubstneq_def)
+qed
+
+lemma msubstneq: assumes lp: "islin (Eq (CNP 0 a r))"
+  shows "Ifm vs (x#bs) (msubstneq c t d s a r) = Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /(1 + 1))#bs) (NEq (CNP 0 a r))" (is "?lhs = ?rhs")
+proof-
+  let ?Nt = "\<lambda>(x::'a) t. Itm vs (x#bs) t"
+  let ?N = "\<lambda>p. Ipoly vs p"
+  let ?c = "?N c"
+  let ?d = "?N d"
+  let ?t = "?Nt x t"
+  let ?s = "?Nt x s"
+  let ?a = "?N a"
+  let ?r = "?Nt x r"
+  from lp have lin:"isnpoly a" "a \<noteq> 0\<^sub>p" "tmbound0 r" "allpolys isnpoly r" by simp_all
+  note r= tmbound0_I[OF lin(3), of vs _ bs x]
+  have cd_cs: "?c * ?d \<noteq> 0 \<or> (?c = 0 \<and> ?d = 0) \<or> (?c = 0 \<and> ?d \<noteq> 0) \<or> (?c \<noteq> 0 \<and> ?d = 0)" by auto
+  moreover
+  {assume c: "?c = 0" and d: "?d=0"
+    hence ?thesis  by (simp add: r[of 0] msubstneq_def Let_def evaldjf_ex)}
+  moreover 
+  {assume c: "?c = 0" and d: "?d\<noteq>0"
+    from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = -?s / ((1 + 1)*?d)" by simp
+    have "?rhs = Ifm vs (-?s / ((1 + 1)*?d) # bs) (NEq (CNP 0 a r))" by (simp only: th)
+    also have "\<dots> \<longleftrightarrow> ?a * (-?s / ((1 + 1)*?d)) + ?r \<noteq> 0" by (simp add: r[of "- (Itm vs (x # bs) s / ((1 + 1) * \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>))"])
+    also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a * (-?s / ((1 + 1)*?d)) + ?r) \<noteq> 0" 
+      using d mult_cancel_left[of "(1 + 1)*?d" "(?a * (-?s / ((1 + 1)*?d)) + ?r)" 0] by simp
+    also have "\<dots> \<longleftrightarrow> (- ?a * ?s) * ((1 + 1)*?d / ((1 + 1)*?d)) + (1 + 1)*?d*?r\<noteq> 0"
+      by (simp add: ring_simps right_distrib[of "(1 + 1)*?d"] del: right_distrib)
+    
+    also have "\<dots> \<longleftrightarrow> - (?a * ?s) + (1 + 1)*?d*?r \<noteq> 0" using d by simp 
+    finally have ?thesis using c d 
+      apply (simp add: r[of "- (Itm vs (x # bs) s / ((1 + 1) * \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>))"] msubstneq_def Let_def evaldjf_ex del: one_add_one_is_two)
+      apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+      apply simp
+      done}
+  moreover
+  {assume c: "?c \<noteq> 0" and d: "?d=0"
+    from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = -?t / ((1 + 1)*?c)" by simp
+    have "?rhs = Ifm vs (-?t / ((1 + 1)*?c) # bs) (NEq (CNP 0 a r))" by (simp only: th)
+    also have "\<dots> \<longleftrightarrow> ?a * (-?t / ((1 + 1)*?c)) + ?r \<noteq> 0" by (simp add: r[of "- (?t/ ((1 + 1)* ?c))"])
+    also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a * (-?t / ((1 + 1)*?c)) + ?r) \<noteq> 0" 
+      using c mult_cancel_left[of "(1 + 1)*?c" "(?a * (-?t / ((1 + 1)*?c)) + ?r)" 0] by simp
+    also have "\<dots> \<longleftrightarrow> (?a * -?t)* ((1 + 1)*?c) / ((1 + 1)*?c) + (1 + 1)*?c*?r \<noteq> 0"
+      by (simp add: ring_simps right_distrib[of "(1 + 1)*?c"] del: right_distrib)
+    also have "\<dots> \<longleftrightarrow> - (?a * ?t) + (1 + 1)*?c*?r \<noteq> 0" using c by simp 
+    finally have ?thesis using c d 
+      apply (simp add: r[of "- (?t/ ((1 + 1)*?c))"] msubstneq_def Let_def evaldjf_ex del: one_add_one_is_two)
+      apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+      apply simp
+      done }
+  moreover
+  {assume c: "?c \<noteq> 0" and d: "?d\<noteq>0" hence dc: "?c * ?d *(1 + 1) \<noteq> 0" by simp
+    from add_frac_eq[OF c d, of "- ?t" "- ?s"]
+    have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)" 
+      by (simp add: ring_simps)
+    have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (NEq (CNP 0 a r))" by (simp only: th)
+    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r \<noteq> 0" 
+      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
+    also have "\<dots> \<longleftrightarrow> ((1 + 1) * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r) \<noteq> 0 "
+      using c d mult_cancel_left[of "(1 + 1) * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r" 0] by simp
+    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + (1 + 1)*?c*?d*?r \<noteq> 0" 
+      using nonzero_mult_divide_cancel_left[OF dc] c d
+      by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+    finally  have ?thesis using c d 
+      apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstneq_def Let_def evaldjf_ex ring_simps)
+      apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+      apply (simp add: ring_simps)
+      done }
+  ultimately show ?thesis by blast
+qed
+
+definition "msubstlt c t d s a r = 
+  evaldjf (split conj) 
+  [(let cd = c *\<^sub>p d in (lt (CP (~\<^sub>p cd)), Lt (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+  (let cd = c *\<^sub>p d in (lt (CP cd), Lt (Sub (Mul a (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+   (conj (lt (CP (~\<^sub>p c))) (Eq (CP d)) , Lt (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+   (conj (lt (CP c)) (Eq (CP d)) , Lt (Sub (Mul a t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+   (conj (lt (CP (~\<^sub>p d))) (Eq (CP c)) , Lt (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+   (conj (lt (CP d)) (Eq (CP c)) , Lt (Sub (Mul a s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+   (conj (Eq (CP c)) (Eq (CP d)) , Lt r)]"
+
+lemma msubstlt_nb: assumes lp: "islin (Lt (CNP 0 a r))" and t: "tmbound0 t" and s: "tmbound0 s"
+  shows "bound0 (msubstlt c t d s a r)"
+proof-
+  have th: "\<forall>x\<in> set [(let cd = c *\<^sub>p d in (lt (CP (~\<^sub>p cd)), Lt (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+  (let cd = c *\<^sub>p d in (lt (CP cd), Lt (Sub (Mul a (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+   (conj (lt (CP (~\<^sub>p c))) (Eq (CP d)) , Lt (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+   (conj (lt (CP c)) (Eq (CP d)) , Lt (Sub (Mul a t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+   (conj (lt (CP (~\<^sub>p d))) (Eq (CP c)) , Lt (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+   (conj (lt (CP d)) (Eq (CP c)) , Lt (Sub (Mul a s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+   (conj (Eq (CP c)) (Eq (CP d)) , Lt r)]. bound0 (split conj x)"
+    using lp by (simp add: Let_def t s lt_nb )
+  from evaldjf_bound0[OF th] show ?thesis by (simp add: msubstlt_def)
+qed
+
+
+lemma msubstlt: assumes nc: "isnpoly c" and nd: "isnpoly d" and lp: "islin (Lt (CNP 0 a r))" 
+  shows "Ifm vs (x#bs) (msubstlt c t d s a r) \<longleftrightarrow> 
+  Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /(1 + 1))#bs) (Lt (CNP 0 a r))" (is "?lhs = ?rhs")
+proof-
+  let ?Nt = "\<lambda>x t. Itm vs (x#bs) t"
+  let ?N = "\<lambda>p. Ipoly vs p"
+  let ?c = "?N c"
+  let ?d = "?N d"
+  let ?t = "?Nt x t"
+  let ?s = "?Nt x s"
+  let ?a = "?N a"
+  let ?r = "?Nt x r"
+  from lp have lin:"isnpoly a" "a \<noteq> 0\<^sub>p" "tmbound0 r" "allpolys isnpoly r" by simp_all
+  note r= tmbound0_I[OF lin(3), of vs _ bs x]
+  have cd_cs: "?c * ?d < 0 \<or> ?c * ?d > 0 \<or> (?c = 0 \<and> ?d = 0) \<or> (?c = 0 \<and> ?d < 0) \<or> (?c = 0 \<and> ?d > 0) \<or> (?c < 0 \<and> ?d = 0) \<or> (?c > 0 \<and> ?d = 0)" by auto
+  moreover
+  {assume c: "?c=0" and d: "?d=0"
+    hence ?thesis  using nc nd by (simp add: polyneg_norm lt r[of 0] msubstlt_def Let_def evaldjf_ex)}
+  moreover
+  {assume dc: "?c*?d > 0" 
+    from mult_pos_pos[OF one_plus_one_pos dc] have dc': "(1 + 1)*?c *?d > 0" by simp
+    hence c:"?c \<noteq> 0" and d: "?d\<noteq> 0" by auto
+    from dc' have dc'': "\<not> (1 + 1)*?c *?d < 0" by simp
+    from add_frac_eq[OF c d, of "- ?t" "- ?s"]
+    have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)" 
+      by (simp add: ring_simps)
+    have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
+    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r < 0" 
+      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
+    also have "\<dots> \<longleftrightarrow> ((1 + 1) * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r) < 0"
+      
+      using dc' dc'' mult_less_cancel_left_disj[of "(1 + 1) * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r" 0] by simp
+    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + (1 + 1)*?c*?d*?r < 0" 
+      using nonzero_mult_divide_cancel_left[of "(1 + 1)*?c*?d"] c d
+      by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+    finally  have ?thesis using dc c d  nc nd dc'
+      apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm) 
+    apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+    by (simp add: ring_simps order_less_not_sym[OF dc])}
+  moreover
+  {assume dc: "?c*?d < 0" 
+
+    from dc one_plus_one_pos[where ?'a='a] have dc': "(1 + 1)*?c *?d < 0"
+      apply (simp add: mult_less_0_iff field_simps) 
+      apply (rule add_neg_neg)
+      apply (simp_all add: mult_less_0_iff)
+      done
+    hence c:"?c \<noteq> 0" and d: "?d\<noteq> 0" by auto
+    from add_frac_eq[OF c d, of "- ?t" "- ?s"]
+    have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)" 
+      by (simp add: ring_simps)
+    have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
+    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r < 0" 
+      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
+
+    also have "\<dots> \<longleftrightarrow> ((1 + 1) * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r) > 0"
+      
+      using dc' order_less_not_sym[OF dc'] mult_less_cancel_left_disj[of "(1 + 1) * ?c * ?d" 0 "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r"] by simp
+    also have "\<dots> \<longleftrightarrow> ?a * ((?d * ?t + ?c* ?s )) - (1 + 1)*?c*?d*?r < 0" 
+      using nonzero_mult_divide_cancel_left[of "(1 + 1)*?c*?d"] c d
+      by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+    finally  have ?thesis using dc c d  nc nd
+      apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm) 
+      apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+      by (simp add: ring_simps order_less_not_sym[OF dc]) }
+  moreover
+  {assume c: "?c > 0" and d: "?d=0"  
+    from c have c'': "(1 + 1)*?c > 0" by (simp add: zero_less_mult_iff)
+    from c have c': "(1 + 1)*?c \<noteq> 0" by simp
+    from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)"  by (simp add: ring_simps)
+    have "?rhs \<longleftrightarrow> Ifm vs (- ?t / ((1 + 1)*?c) # bs) (Lt (CNP 0 a r))" by (simp only: th)
+    also have "\<dots> \<longleftrightarrow> ?a* (- ?t / ((1 + 1)*?c))+ ?r < 0" by (simp add: r[of "- (?t / ((1 + 1)*?c))"])
+    also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a* (- ?t / ((1 + 1)*?c))+ ?r) < 0"
+      using c mult_less_cancel_left_disj[of "(1 + 1) * ?c" "?a* (- ?t / ((1 + 1)*?c))+ ?r" 0] c' c'' order_less_not_sym[OF c''] by simp
+    also have "\<dots> \<longleftrightarrow> - ?a*?t+  (1 + 1)*?c *?r < 0" 
+      using nonzero_mult_divide_cancel_left[OF c'] c
+      by (simp add: ring_simps diff_divide_distrib less_le del: left_distrib)
+    finally have ?thesis using c d nc nd 
+      apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+      apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+      using c order_less_not_sym[OF c] less_imp_neq[OF c]
+      by (simp add: ring_simps )  }
+  moreover
+  {assume c: "?c < 0" and d: "?d=0"  hence c': "(1 + 1)*?c \<noteq> 0" by simp
+    from c have c'': "(1 + 1)*?c < 0" by (simp add: mult_less_0_iff)
+    from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)"  by (simp add: ring_simps)
+    have "?rhs \<longleftrightarrow> Ifm vs (- ?t / ((1 + 1)*?c) # bs) (Lt (CNP 0 a r))" by (simp only: th)
+    also have "\<dots> \<longleftrightarrow> ?a* (- ?t / ((1 + 1)*?c))+ ?r < 0" by (simp add: r[of "- (?t / ((1 + 1)*?c))"])
+    also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a* (- ?t / ((1 + 1)*?c))+ ?r) > 0"
+      using c order_less_not_sym[OF c''] less_imp_neq[OF c''] c'' mult_less_cancel_left_disj[of "(1 + 1) * ?c" 0 "?a* (- ?t / ((1 + 1)*?c))+ ?r"] by simp
+    also have "\<dots> \<longleftrightarrow> ?a*?t -  (1 + 1)*?c *?r < 0" 
+      using nonzero_mult_divide_cancel_left[OF c'] c order_less_not_sym[OF c''] less_imp_neq[OF c''] c''
+	by (simp add: ring_simps diff_divide_distrib del:  left_distrib)
+    finally have ?thesis using c d nc nd 
+      apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+      apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+      using c order_less_not_sym[OF c] less_imp_neq[OF c]
+      by (simp add: ring_simps )    }
+  moreover
+  moreover
+  {assume c: "?c = 0" and d: "?d>0"  
+    from d have d'': "(1 + 1)*?d > 0" by (simp add: zero_less_mult_iff)
+    from d have d': "(1 + 1)*?d \<noteq> 0" by simp
+    from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)"  by (simp add: ring_simps)
+    have "?rhs \<longleftrightarrow> Ifm vs (- ?s / ((1 + 1)*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
+    also have "\<dots> \<longleftrightarrow> ?a* (- ?s / ((1 + 1)*?d))+ ?r < 0" by (simp add: r[of "- (?s / ((1 + 1)*?d))"])
+    also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a* (- ?s / ((1 + 1)*?d))+ ?r) < 0"
+      using d mult_less_cancel_left_disj[of "(1 + 1) * ?d" "?a* (- ?s / ((1 + 1)*?d))+ ?r" 0] d' d'' order_less_not_sym[OF d''] by simp
+    also have "\<dots> \<longleftrightarrow> - ?a*?s+  (1 + 1)*?d *?r < 0" 
+      using nonzero_mult_divide_cancel_left[OF d'] d
+      by (simp add: ring_simps diff_divide_distrib less_le del: left_distrib)
+    finally have ?thesis using c d nc nd 
+      apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+      apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+      using d order_less_not_sym[OF d] less_imp_neq[OF d]
+      by (simp add: ring_simps )  }
+  moreover
+  {assume c: "?c = 0" and d: "?d<0"  hence d': "(1 + 1)*?d \<noteq> 0" by simp
+    from d have d'': "(1 + 1)*?d < 0" by (simp add: mult_less_0_iff)
+    from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)"  by (simp add: ring_simps)
+    have "?rhs \<longleftrightarrow> Ifm vs (- ?s / ((1 + 1)*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
+    also have "\<dots> \<longleftrightarrow> ?a* (- ?s / ((1 + 1)*?d))+ ?r < 0" by (simp add: r[of "- (?s / ((1 + 1)*?d))"])
+    also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a* (- ?s / ((1 + 1)*?d))+ ?r) > 0"
+      using d order_less_not_sym[OF d''] less_imp_neq[OF d''] d'' mult_less_cancel_left_disj[of "(1 + 1) * ?d" 0 "?a* (- ?s / ((1 + 1)*?d))+ ?r"] by simp
+    also have "\<dots> \<longleftrightarrow> ?a*?s -  (1 + 1)*?d *?r < 0" 
+      using nonzero_mult_divide_cancel_left[OF d'] d order_less_not_sym[OF d''] less_imp_neq[OF d''] d''
+	by (simp add: ring_simps diff_divide_distrib del:  left_distrib)
+    finally have ?thesis using c d nc nd 
+      apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+      apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+      using d order_less_not_sym[OF d] less_imp_neq[OF d]
+      by (simp add: ring_simps )    }
+ultimately show ?thesis by blast
+qed
+
+definition "msubstle c t d s a r = 
+  evaldjf (split conj) 
+  [(let cd = c *\<^sub>p d in (lt (CP (~\<^sub>p cd)), Le (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+  (let cd = c *\<^sub>p d in (lt (CP cd), Le (Sub (Mul a (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+   (conj (lt (CP (~\<^sub>p c))) (Eq (CP d)) , Le (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+   (conj (lt (CP c)) (Eq (CP d)) , Le (Sub (Mul a t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+   (conj (lt (CP (~\<^sub>p d))) (Eq (CP c)) , Le (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+   (conj (lt (CP d)) (Eq (CP c)) , Le (Sub (Mul a s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+   (conj (Eq (CP c)) (Eq (CP d)) , Le r)]"
+
+lemma msubstle_nb: assumes lp: "islin (Le (CNP 0 a r))" and t: "tmbound0 t" and s: "tmbound0 s"
+  shows "bound0 (msubstle c t d s a r)"
+proof-
+  have th: "\<forall>x\<in> set [(let cd = c *\<^sub>p d in (lt (CP (~\<^sub>p cd)), Le (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+  (let cd = c *\<^sub>p d in (lt (CP cd), Le (Sub (Mul a (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+   (conj (lt (CP (~\<^sub>p c))) (Eq (CP d)) , Le (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+   (conj (lt (CP c)) (Eq (CP d)) , Le (Sub (Mul a t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+   (conj (lt (CP (~\<^sub>p d))) (Eq (CP c)) , Le (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+   (conj (lt (CP d)) (Eq (CP c)) , Le (Sub (Mul a s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+   (conj (Eq (CP c)) (Eq (CP d)) , Le r)]. bound0 (split conj x)"
+    using lp by (simp add: Let_def t s lt_nb )
+  from evaldjf_bound0[OF th] show ?thesis by (simp add: msubstle_def)
+qed
+
+lemma msubstle: assumes nc: "isnpoly c" and nd: "isnpoly d" and lp: "islin (Le (CNP 0 a r))" 
+  shows "Ifm vs (x#bs) (msubstle c t d s a r) \<longleftrightarrow> 
+  Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /(1 + 1))#bs) (Le (CNP 0 a r))" (is "?lhs = ?rhs")
+proof-
+  let ?Nt = "\<lambda>x t. Itm vs (x#bs) t"
+  let ?N = "\<lambda>p. Ipoly vs p"
+  let ?c = "?N c"
+  let ?d = "?N d"
+  let ?t = "?Nt x t"
+  let ?s = "?Nt x s"
+  let ?a = "?N a"
+  let ?r = "?Nt x r"
+  from lp have lin:"isnpoly a" "a \<noteq> 0\<^sub>p" "tmbound0 r" "allpolys isnpoly r" by simp_all
+  note r= tmbound0_I[OF lin(3), of vs _ bs x]
+  have cd_cs: "?c * ?d < 0 \<or> ?c * ?d > 0 \<or> (?c = 0 \<and> ?d = 0) \<or> (?c = 0 \<and> ?d < 0) \<or> (?c = 0 \<and> ?d > 0) \<or> (?c < 0 \<and> ?d = 0) \<or> (?c > 0 \<and> ?d = 0)" by auto
+  moreover
+  {assume c: "?c=0" and d: "?d=0"
+    hence ?thesis  using nc nd by (simp add: polyneg_norm polymul_norm lt r[of 0] msubstle_def Let_def evaldjf_ex)}
+  moreover
+  {assume dc: "?c*?d > 0" 
+    from mult_pos_pos[OF one_plus_one_pos dc] have dc': "(1 + 1)*?c *?d > 0" by simp
+    hence c:"?c \<noteq> 0" and d: "?d\<noteq> 0" by auto
+    from dc' have dc'': "\<not> (1 + 1)*?c *?d < 0" by simp
+    from add_frac_eq[OF c d, of "- ?t" "- ?s"]
+    have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)" 
+      by (simp add: ring_simps)
+    have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
+    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r <= 0" 
+      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
+    also have "\<dots> \<longleftrightarrow> ((1 + 1) * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r) <= 0"
+      
+      using dc' dc'' mult_le_cancel_left[of "(1 + 1) * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r" 0] by simp
+    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + (1 + 1)*?c*?d*?r <= 0" 
+      using nonzero_mult_divide_cancel_left[of "(1 + 1)*?c*?d"] c d
+      by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+    finally  have ?thesis using dc c d  nc nd dc'
+      apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm) 
+    apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+    by (simp add: ring_simps order_less_not_sym[OF dc])}
+  moreover
+  {assume dc: "?c*?d < 0" 
+
+    from dc one_plus_one_pos[where ?'a='a] have dc': "(1 + 1)*?c *?d < 0"
+      by (simp add: mult_less_0_iff field_simps add_neg_neg add_pos_pos)
+    hence c:"?c \<noteq> 0" and d: "?d\<noteq> 0" by auto
+    from add_frac_eq[OF c d, of "- ?t" "- ?s"]
+    have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)" 
+      by (simp add: ring_simps)
+    have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
+    also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r <= 0" 
+      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
+
+    also have "\<dots> \<longleftrightarrow> ((1 + 1) * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r) >= 0"
+      
+      using dc' order_less_not_sym[OF dc'] mult_le_cancel_left[of "(1 + 1) * ?c * ?d" 0 "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r"] by simp
+    also have "\<dots> \<longleftrightarrow> ?a * ((?d * ?t + ?c* ?s )) - (1 + 1)*?c*?d*?r <= 0" 
+      using nonzero_mult_divide_cancel_left[of "(1 + 1)*?c*?d"] c d
+      by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+    finally  have ?thesis using dc c d  nc nd
+      apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm) 
+      apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+      by (simp add: ring_simps order_less_not_sym[OF dc]) }
+  moreover
+  {assume c: "?c > 0" and d: "?d=0"  
+    from c have c'': "(1 + 1)*?c > 0" by (simp add: zero_less_mult_iff)
+    from c have c': "(1 + 1)*?c \<noteq> 0" by simp
+    from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)"  by (simp add: ring_simps)
+    have "?rhs \<longleftrightarrow> Ifm vs (- ?t / ((1 + 1)*?c) # bs) (Le (CNP 0 a r))" by (simp only: th)
+    also have "\<dots> \<longleftrightarrow> ?a* (- ?t / ((1 + 1)*?c))+ ?r <= 0" by (simp add: r[of "- (?t / ((1 + 1)*?c))"])
+    also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a* (- ?t / ((1 + 1)*?c))+ ?r) <= 0"
+      using c mult_le_cancel_left[of "(1 + 1) * ?c" "?a* (- ?t / ((1 + 1)*?c))+ ?r" 0] c' c'' order_less_not_sym[OF c''] by simp
+    also have "\<dots> \<longleftrightarrow> - ?a*?t+  (1 + 1)*?c *?r <= 0" 
+      using nonzero_mult_divide_cancel_left[OF c'] c
+      by (simp add: ring_simps diff_divide_distrib less_le del: left_distrib)
+    finally have ?thesis using c d nc nd 
+      apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+      apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+      using c order_less_not_sym[OF c] less_imp_neq[OF c]
+      by (simp add: ring_simps )  }
+  moreover
+  {assume c: "?c < 0" and d: "?d=0"  hence c': "(1 + 1)*?c \<noteq> 0" by simp
+    from c have c'': "(1 + 1)*?c < 0" by (simp add: mult_less_0_iff)
+    from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)"  by (simp add: ring_simps)
+    have "?rhs \<longleftrightarrow> Ifm vs (- ?t / ((1 + 1)*?c) # bs) (Le (CNP 0 a r))" by (simp only: th)
+    also have "\<dots> \<longleftrightarrow> ?a* (- ?t / ((1 + 1)*?c))+ ?r <= 0" by (simp add: r[of "- (?t / ((1 + 1)*?c))"])
+    also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a* (- ?t / ((1 + 1)*?c))+ ?r) >= 0"
+      using c order_less_not_sym[OF c''] less_imp_neq[OF c''] c'' mult_le_cancel_left[of "(1 + 1) * ?c" 0 "?a* (- ?t / ((1 + 1)*?c))+ ?r"] by simp
+    also have "\<dots> \<longleftrightarrow> ?a*?t -  (1 + 1)*?c *?r <= 0" 
+      using nonzero_mult_divide_cancel_left[OF c'] c order_less_not_sym[OF c''] less_imp_neq[OF c''] c''
+	by (simp add: ring_simps diff_divide_distrib del:  left_distrib)
+    finally have ?thesis using c d nc nd 
+      apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+      apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+      using c order_less_not_sym[OF c] less_imp_neq[OF c]
+      by (simp add: ring_simps )    }
+  moreover
+  moreover
+  {assume c: "?c = 0" and d: "?d>0"  
+    from d have d'': "(1 + 1)*?d > 0" by (simp add: zero_less_mult_iff)
+    from d have d': "(1 + 1)*?d \<noteq> 0" by simp
+    from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)"  by (simp add: ring_simps)
+    have "?rhs \<longleftrightarrow> Ifm vs (- ?s / ((1 + 1)*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
+    also have "\<dots> \<longleftrightarrow> ?a* (- ?s / ((1 + 1)*?d))+ ?r <= 0" by (simp add: r[of "- (?s / ((1 + 1)*?d))"])
+    also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a* (- ?s / ((1 + 1)*?d))+ ?r) <= 0"
+      using d mult_le_cancel_left[of "(1 + 1) * ?d" "?a* (- ?s / ((1 + 1)*?d))+ ?r" 0] d' d'' order_less_not_sym[OF d''] by simp
+    also have "\<dots> \<longleftrightarrow> - ?a*?s+  (1 + 1)*?d *?r <= 0" 
+      using nonzero_mult_divide_cancel_left[OF d'] d
+      by (simp add: ring_simps diff_divide_distrib less_le del: left_distrib)
+    finally have ?thesis using c d nc nd 
+      apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+      apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+      using d order_less_not_sym[OF d] less_imp_neq[OF d]
+      by (simp add: ring_simps )  }
+  moreover
+  {assume c: "?c = 0" and d: "?d<0"  hence d': "(1 + 1)*?d \<noteq> 0" by simp
+    from d have d'': "(1 + 1)*?d < 0" by (simp add: mult_less_0_iff)
+    from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)"  by (simp add: ring_simps)
+    have "?rhs \<longleftrightarrow> Ifm vs (- ?s / ((1 + 1)*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
+    also have "\<dots> \<longleftrightarrow> ?a* (- ?s / ((1 + 1)*?d))+ ?r <= 0" by (simp add: r[of "- (?s / ((1 + 1)*?d))"])
+    also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a* (- ?s / ((1 + 1)*?d))+ ?r) >= 0"
+      using d order_less_not_sym[OF d''] less_imp_neq[OF d''] d'' mult_le_cancel_left[of "(1 + 1) * ?d" 0 "?a* (- ?s / ((1 + 1)*?d))+ ?r"] by simp
+    also have "\<dots> \<longleftrightarrow> ?a*?s -  (1 + 1)*?d *?r <= 0" 
+      using nonzero_mult_divide_cancel_left[OF d'] d order_less_not_sym[OF d''] less_imp_neq[OF d''] d''
+	by (simp add: ring_simps diff_divide_distrib del:  left_distrib)
+    finally have ?thesis using c d nc nd 
+      apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+      apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+      using d order_less_not_sym[OF d] less_imp_neq[OF d]
+      by (simp add: ring_simps )    }
+ultimately show ?thesis by blast
+qed
+
+
+fun msubst :: "fm \<Rightarrow> (poly \<times> tm) \<times> (poly \<times> tm) \<Rightarrow> fm" where
+  "msubst (And p q) ((c,t), (d,s)) = conj (msubst p ((c,t),(d,s))) (msubst q ((c,t),(d,s)))"
+| "msubst (Or p q) ((c,t), (d,s)) = disj (msubst p ((c,t),(d,s))) (msubst q ((c,t), (d,s)))"
+| "msubst (Eq (CNP 0 a r)) ((c,t),(d,s)) = msubsteq c t d s a r"
+| "msubst (NEq (CNP 0 a r)) ((c,t),(d,s)) = msubstneq c t d s a r"
+| "msubst (Lt (CNP 0 a r)) ((c,t),(d,s)) = msubstlt c t d s a r"
+| "msubst (Le (CNP 0 a r)) ((c,t),(d,s)) = msubstle c t d s a r"
+| "msubst p ((c,t),(d,s)) = p"
+
+lemma msubst_I: assumes lp: "islin p" and nc: "isnpoly c" and nd: "isnpoly d"
+  shows "Ifm vs (x#bs) (msubst p ((c,t),(d,s))) = Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /(1 + 1))#bs) p"
+  using lp
+by (induct p rule: islin.induct, auto simp add: tmbound0_I[where b="(- (Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>) + - (Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>)) /(1 + 1)" and b'=x and bs = bs and vs=vs] bound0_I[where b="(- (Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>) + - (Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>)) /(1 + 1)" and b'=x and bs = bs and vs=vs] msubsteq msubstneq msubstlt[OF nc nd] msubstle[OF nc nd])
+
+lemma msubst_nb: assumes lp: "islin p" and t: "tmbound0 t" and s: "tmbound0 s"
+  shows "bound0 (msubst p ((c,t),(d,s)))"
+  using lp t s
+  by (induct p rule: islin.induct, auto simp add: msubsteq_nb msubstneq_nb msubstlt_nb msubstle_nb)
+
+lemma fr_eq_msubst: 
+  assumes lp: "islin p"
+  shows "(\<exists> x. Ifm vs (x#bs) p) = ((Ifm vs (x#bs) (minusinf p)) \<or> (Ifm vs (x#bs) (plusinf p)) \<or> (\<exists> (c,t) \<in> set (uset p). \<exists> (d,s) \<in> set (uset p). Ifm vs (x#bs) (msubst p ((c,t),(d,s)))))"
+  (is "(\<exists> x. ?I x p) = (?M \<or> ?P \<or> ?F)" is "?E = ?D")
+proof-
+from uset_l[OF lp] have th: "\<forall>(c, s)\<in>set (uset p). isnpoly c \<and> tmbound0 s" by blast
+{fix c t d s assume ctU: "(c,t) \<in>set (uset p)" and dsU: "(d,s) \<in>set (uset p)" 
+  and pts: "Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / (1+1) # bs) p"
+  from th[rule_format, OF ctU] th[rule_format, OF dsU] have norm:"isnpoly c" "isnpoly d" by simp_all
+  from msubst_I[OF lp norm, of vs x bs t s] pts
+  have "Ifm vs (x # bs) (msubst p ((c, t), d, s))" ..}
+moreover
+{fix c t d s assume ctU: "(c,t) \<in>set (uset p)" and dsU: "(d,s) \<in>set (uset p)" 
+  and pts: "Ifm vs (x # bs) (msubst p ((c, t), d, s))"
+  from th[rule_format, OF ctU] th[rule_format, OF dsU] have norm:"isnpoly c" "isnpoly d" by simp_all
+  from msubst_I[OF lp norm, of vs x bs t s] pts
+  have "Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / (1+1) # bs) p" ..}
+ultimately have th': "(\<exists> (c,t) \<in> set (uset p). \<exists> (d,s) \<in> set (uset p). Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / (1+1) # bs) p) \<longleftrightarrow> ?F" by blast
+from fr_eq[OF lp, of vs bs x, simplified th'] show ?thesis .
+qed 
+
+text {* Rest of the implementation *}
+
+consts alluopairs:: "'a list \<Rightarrow> ('a \<times> 'a) list"
+primrec
+  "alluopairs [] = []"
+  "alluopairs (x#xs) = (map (Pair x) (x#xs))@(alluopairs xs)"
+
+lemma alluopairs_set1: "set (alluopairs xs) \<le> {(x,y). x\<in> set xs \<and> y\<in> set xs}"
+by (induct xs, auto)
+
+lemma alluopairs_set:
+  "\<lbrakk>x\<in> set xs ; y \<in> set xs\<rbrakk> \<Longrightarrow> (x,y) \<in> set (alluopairs xs) \<or> (y,x) \<in> set (alluopairs xs) "
+by (induct xs, auto)
+
+lemma alluopairs_ex:
+  assumes Pc: "\<forall> x \<in> set xs. \<forall>y\<in> set xs. P x y = P y x"
+  shows "(\<exists> x \<in> set xs. \<exists> y \<in> set xs. P x y) = (\<exists> (x,y) \<in> set (alluopairs xs). P x y)"
+proof
+  assume "\<exists>x\<in>set xs. \<exists>y\<in>set xs. P x y"
+  then obtain x y where x: "x \<in> set xs" and y:"y \<in> set xs" and P: "P x y"  by blast
+  from alluopairs_set[OF x y] P Pc x y show"\<exists>(x, y)\<in>set (alluopairs xs). P x y" 
+    by auto
+next
+  assume "\<exists>(x, y)\<in>set (alluopairs xs). P x y"
+  then obtain "x" and "y"  where xy:"(x,y) \<in> set (alluopairs xs)" and P: "P x y" by blast+
+  from xy have "x \<in> set xs \<and> y\<in> set xs" using alluopairs_set1 by blast
+  with P show "\<exists>x\<in>set xs. \<exists>y\<in>set xs. P x y" by blast
+qed
+
+lemma nth_pos2: "0 < n \<Longrightarrow> (x#xs) ! n = xs ! (n - 1)"
+using Nat.gr0_conv_Suc
+by clarsimp
+
+lemma filter_length: "length (List.filter P xs) < Suc (length xs)"
+  apply (induct xs, auto) done
+
+consts remdps:: "'a list \<Rightarrow> 'a list"
+
+recdef remdps "measure size"
+  "remdps [] = []"
+  "remdps (x#xs) = (x#(remdps (List.filter (\<lambda> y. y \<noteq> x) xs)))"
+(hints simp add: filter_length[rule_format])
+
+lemma remdps_set[simp]: "set (remdps xs) = set xs"
+  by (induct xs rule: remdps.induct, auto)
+
+lemma simpfm_lin:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "qfree p \<Longrightarrow> islin (simpfm p)"
+  by (induct p rule: simpfm.induct, auto simp add: conj_lin disj_lin)
+
+definition 
+  "ferrack p \<equiv> let q = simpfm p ; mp = minusinf q ; pp = plusinf q
+  in if (mp = T \<or> pp = T) then T 
+     else (let U = alluopairs (remdps (uset  q))
+           in decr0 (disj mp (disj pp (evaldjf (simpfm o (msubst q)) U ))))"
+
+lemma ferrack: 
+  assumes qf: "qfree p"
+  shows "qfree (ferrack p) \<and> ((Ifm vs bs (ferrack p)) = (Ifm vs bs (E p)))"
+  (is "_ \<and> (?rhs = ?lhs)")
+proof-
+  let ?I = "\<lambda> x p. Ifm vs (x#bs) p"
+  let ?N = "\<lambda> t. Ipoly vs t"
+  let ?Nt = "\<lambda>x t. Itm vs (x#bs) t"
+  let ?q = "simpfm p" 
+  let ?U = "remdps(uset ?q)"
+  let ?Up = "alluopairs ?U"
+  let ?mp = "minusinf ?q"
+  let ?pp = "plusinf ?q"
+  let ?I = "\<lambda>p. Ifm vs (x#bs) p"
+  from simpfm_lin[OF qf] simpfm_qf[OF qf] have lq: "islin ?q" and q_qf: "qfree ?q" .
+  from minusinf_nb[OF lq] plusinf_nb[OF lq] have mp_nb: "bound0 ?mp" and pp_nb: "bound0 ?pp" .
+  from bound0_qf[OF mp_nb] bound0_qf[OF pp_nb] have mp_qf: "qfree ?mp" and pp_qf: "qfree ?pp" .
+  from uset_l[OF lq] have U_l: "\<forall>(c, s)\<in>set ?U. isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s"
+    by simp
+  {fix c t d s assume ctU: "(c,t) \<in> set ?U" and dsU: "(d,s) \<in> set ?U"
+    from U_l ctU dsU have norm: "isnpoly c" "isnpoly d" by auto
+    from msubst_I[OF lq norm, of vs x bs t s] msubst_I[OF lq norm(2,1), of vs x bs s t]
+    have "?I (msubst ?q ((c,t),(d,s))) = ?I (msubst ?q ((d,s),(c,t)))" by (simp add: ring_simps)}
+  hence th0: "\<forall>x \<in> set ?U. \<forall>y \<in> set ?U. ?I (msubst ?q (x, y)) \<longleftrightarrow> ?I (msubst ?q (y, x))" by clarsimp
+  {fix x assume xUp: "x \<in> set ?Up" 
+    then  obtain c t d s where ctU: "(c,t) \<in> set ?U" and dsU: "(d,s) \<in> set ?U" 
+      and x: "x = ((c,t),(d,s))" using alluopairs_set1[of ?U] by auto  
+    from U_l[rule_format, OF ctU] U_l[rule_format, OF dsU] 
+    have nbs: "tmbound0 t" "tmbound0 s" by simp_all
+    from simpfm_bound0[OF msubst_nb[OF lq nbs, of c d]] 
+    have "bound0 ((simpfm o (msubst (simpfm p))) x)" using x by simp}
+  with evaldjf_bound0[of ?Up "(simpfm o (msubst (simpfm p)))"]
+  have "bound0 (evaldjf (simpfm o (msubst (simpfm p))) ?Up)" by blast
+  with mp_nb pp_nb 
+  have th1: "bound0 (disj ?mp (disj ?pp (evaldjf (simpfm o (msubst ?q)) ?Up )))" by (simp add: disj_nb)
+  from decr0_qf[OF th1] have thqf: "qfree (ferrack p)" by (simp add: ferrack_def Let_def)
+  have "?lhs \<longleftrightarrow> (\<exists>x. Ifm vs (x#bs) ?q)" by simp
+  also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> (\<exists>(c, t)\<in>set ?U. \<exists>(d, s)\<in>set ?U. ?I (msubst (simpfm p) ((c, t), d, s)))" using fr_eq_msubst[OF lq, of vs bs x] by simp
+  also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> (\<exists> (x,y) \<in> set ?Up. ?I ((simpfm o (msubst ?q)) (x,y)))" using alluopairs_ex[OF th0] by simp
+  also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> ?I (evaldjf (simpfm o (msubst ?q)) ?Up)" 
+    by (simp add: evaldjf_ex)
+  also have "\<dots> \<longleftrightarrow> ?I (disj ?mp (disj ?pp (evaldjf (simpfm o (msubst ?q)) ?Up)))" by simp
+  also have "\<dots> \<longleftrightarrow> ?rhs" using decr0[OF th1, of vs x bs]
+    apply (simp add: ferrack_def Let_def)
+    by (cases "?mp = T \<or> ?pp = T", auto)
+  finally show ?thesis using thqf by blast
+qed
+
+definition "frpar p = simpfm (qelim p ferrack)"
+lemma frpar: "qfree (frpar p) \<and> (Ifm vs bs (frpar p) \<longleftrightarrow> Ifm vs bs p)"
+proof-
+  from ferrack have th: "\<forall>bs p. qfree p \<longrightarrow> qfree (ferrack p) \<and> Ifm vs bs (ferrack p) = Ifm vs bs (E p)" by blast
+  from qelim[OF th, of p bs] show ?thesis  unfolding frpar_def by auto
+qed
+
+declare polyadd.simps[code]
+lemma [simp,code]: "polyadd (CN c n p, CN c' n' p') = 
+    (if n < n' then CN (polyadd(c,CN c' n' p')) n p
+     else if n'<n then CN (polyadd(CN c n p, c')) n' p'
+     else (let cc' = polyadd (c,c') ; 
+               pp' = polyadd (p,p')
+           in (if pp' = 0\<^sub>p then cc' else CN cc' n pp')))"
+  by (simp add: Let_def stupid)
+
+
+
+(*
+lemmas [code func] = polysub_def
+lemmas [code func del] = Zero_nat_def
+code_gen  "frpar" in SML to FRParTest
+*)
+
+section{* Second implemenation: Case splits not local *}
+
+lemma fr_eq2:  assumes lp: "islin p"
+  shows "(\<exists> x. Ifm vs (x#bs) p) \<longleftrightarrow> 
+   ((Ifm vs (x#bs) (minusinf p)) \<or> (Ifm vs (x#bs) (plusinf p)) \<or> 
+    (Ifm vs (0#bs) p) \<or> 
+    (\<exists> (n,t) \<in> set (uset p). Ipoly vs n \<noteq> 0 \<and> Ifm vs ((- Itm vs (x#bs) t /  (Ipoly vs n * (1 + 1)))#bs) p) \<or> 
+    (\<exists> (n,t) \<in> set (uset p). \<exists> (m,s) \<in> set (uset p). Ipoly vs n \<noteq> 0 \<and> Ipoly vs m \<noteq> 0 \<and> Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs n + - Itm vs (x#bs) s / Ipoly vs m) /(1 + 1))#bs) p))"
+  (is "(\<exists> x. ?I x p) = (?M \<or> ?P \<or> ?Z \<or> ?U \<or> ?F)" is "?E = ?D")
+proof
+  assume px: "\<exists> x. ?I x p"
+  have "?M \<or> ?P \<or> (\<not> ?M \<and> \<not> ?P)" by blast
+  moreover {assume "?M \<or> ?P" hence "?D" by blast}
+  moreover {assume nmi: "\<not> ?M" and npi: "\<not> ?P"
+    from inf_uset[OF lp nmi npi, OF px] 
+    obtain c t d s where ct: "(c,t) \<in> set (uset p)" "(d,s) \<in> set (uset p)" "?I ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / ((1\<Colon>'a) + (1\<Colon>'a))) p"
+      by auto
+    let ?c = "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>"
+    let ?d = "\<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>"
+    let ?s = "Itm vs (x # bs) s"
+    let ?t = "Itm vs (x # bs) t"
+    have eq2: "\<And>(x::'a). x + x = (1 + 1) * x"
+      by  (simp add: ring_simps)
+    {assume "?c = 0 \<and> ?d = 0"
+      with ct have ?D by simp}
+    moreover
+    {assume z: "?c = 0" "?d \<noteq> 0"
+      from z have ?D using ct by auto}
+    moreover
+    {assume z: "?c \<noteq> 0" "?d = 0"
+      with ct have ?D by auto }
+    moreover
+    {assume z: "?c \<noteq> 0" "?d \<noteq> 0"
+      from z have ?F using ct
+	apply - apply (rule bexI[where x = "(c,t)"], simp_all)
+	by (rule bexI[where x = "(d,s)"], simp_all)
+      hence ?D by blast}
+    ultimately have ?D by auto}
+  ultimately show "?D" by blast
+next
+  assume "?D" 
+  moreover {assume m:"?M" from minusinf_ex[OF lp m] have "?E" .}
+  moreover {assume p: "?P" from plusinf_ex[OF lp p] have "?E" . }
+  moreover {assume f:"?F" hence "?E" by blast}
+  ultimately show "?E" by blast
+qed
+
+definition "msubsteq2 c t a b = Eq (Add (Mul a t) (Mul c b))"
+definition "msubstltpos c t a b = Lt (Add (Mul a t) (Mul c b))"
+definition "msubstlepos c t a b = Le (Add (Mul a t) (Mul c b))"
+definition "msubstltneg c t a b = Lt (Neg (Add (Mul a t) (Mul c b)))"
+definition "msubstleneg c t a b = Le (Neg (Add (Mul a t) (Mul c b)))"
+
+lemma msubsteq2: 
+  assumes nz: "Ipoly vs c \<noteq> 0" and l: "islin (Eq (CNP 0 a b))"
+  shows "Ifm vs (x#bs) (msubsteq2 c t a b) = Ifm vs (((Itm vs (x#bs) t /  Ipoly vs c ))#bs) (Eq (CNP 0 a b))" (is "?lhs = ?rhs")
+  using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" , symmetric]
+  by (simp add: msubsteq2_def field_simps)
+
+lemma msubstltpos: 
+  assumes nz: "Ipoly vs c > 0" and l: "islin (Lt (CNP 0 a b))"
+  shows "Ifm vs (x#bs) (msubstltpos c t a b) = Ifm vs (((Itm vs (x#bs) t /  Ipoly vs c ))#bs) (Lt (CNP 0 a b))" (is "?lhs = ?rhs")
+  using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" , symmetric]
+  by (simp add: msubstltpos_def field_simps)
+
+lemma msubstlepos: 
+  assumes nz: "Ipoly vs c > 0" and l: "islin (Le (CNP 0 a b))"
+  shows "Ifm vs (x#bs) (msubstlepos c t a b) = Ifm vs (((Itm vs (x#bs) t /  Ipoly vs c ))#bs) (Le (CNP 0 a b))" (is "?lhs = ?rhs")
+  using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" , symmetric]
+  by (simp add: msubstlepos_def field_simps)
+
+lemma msubstltneg: 
+  assumes nz: "Ipoly vs c < 0" and l: "islin (Lt (CNP 0 a b))"
+  shows "Ifm vs (x#bs) (msubstltneg c t a b) = Ifm vs (((Itm vs (x#bs) t /  Ipoly vs c ))#bs) (Lt (CNP 0 a b))" (is "?lhs = ?rhs")
+  using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" , symmetric]
+  by (simp add: msubstltneg_def field_simps del: minus_add_distrib)
+
+lemma msubstleneg: 
+  assumes nz: "Ipoly vs c < 0" and l: "islin (Le (CNP 0 a b))"
+  shows "Ifm vs (x#bs) (msubstleneg c t a b) = Ifm vs (((Itm vs (x#bs) t /  Ipoly vs c ))#bs) (Le (CNP 0 a b))" (is "?lhs = ?rhs")
+  using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" , symmetric]
+  by (simp add: msubstleneg_def field_simps del: minus_add_distrib)
+
+fun msubstpos :: "fm \<Rightarrow> poly \<Rightarrow> tm \<Rightarrow> fm" where
+  "msubstpos (And p q) c t = And (msubstpos p c t) (msubstpos q c t)"
+| "msubstpos (Or p q) c t = Or (msubstpos p c t) (msubstpos q c t)"
+| "msubstpos (Eq (CNP 0 a r)) c t = msubsteq2 c t a r"
+| "msubstpos (NEq (CNP 0 a r)) c t = NOT (msubsteq2 c t a r)"
+| "msubstpos (Lt (CNP 0 a r)) c t = msubstltpos c t a r"
+| "msubstpos (Le (CNP 0 a r)) c t = msubstlepos c t a r"
+| "msubstpos p c t = p"
+    
+lemma msubstpos_I: 
+  assumes lp: "islin p" and pos: "Ipoly vs c > 0"
+  shows "Ifm vs (x#bs) (msubstpos p c t) = Ifm vs (Itm vs (x#bs) t /  Ipoly vs c #bs) p"
+  using lp pos
+  by (induct p rule: islin.induct, auto simp add: msubsteq2 msubstltpos[OF pos] msubstlepos[OF pos] tmbound0_I[of _ vs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" bs x] bound0_I[of _ vs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" bs x] field_simps)
+
+fun msubstneg :: "fm \<Rightarrow> poly \<Rightarrow> tm \<Rightarrow> fm" where
+  "msubstneg (And p q) c t = And (msubstneg p c t) (msubstneg q c t)"
+| "msubstneg (Or p q) c t = Or (msubstneg p c t) (msubstneg q c t)"
+| "msubstneg (Eq (CNP 0 a r)) c t = msubsteq2 c t a r"
+| "msubstneg (NEq (CNP 0 a r)) c t = NOT (msubsteq2 c t a r)"
+| "msubstneg (Lt (CNP 0 a r)) c t = msubstltneg c t a r"
+| "msubstneg (Le (CNP 0 a r)) c t = msubstleneg c t a r"
+| "msubstneg p c t = p"
+
+lemma msubstneg_I: 
+  assumes lp: "islin p" and pos: "Ipoly vs c < 0"
+  shows "Ifm vs (x#bs) (msubstneg p c t) = Ifm vs (Itm vs (x#bs) t /  Ipoly vs c #bs) p"
+  using lp pos
+  by (induct p rule: islin.induct, auto simp add: msubsteq2 msubstltneg[OF pos] msubstleneg[OF pos] tmbound0_I[of _ vs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" bs x] bound0_I[of _ vs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" bs x] field_simps)
+
+
+definition "msubst2 p c t = disj (conj (lt (CP (polyneg c))) (simpfm (msubstpos p c t))) (conj (lt (CP c)) (simpfm (msubstneg p c t)))"
+
+lemma msubst2: assumes lp: "islin p" and nc: "isnpoly c" and nz: "Ipoly vs c \<noteq> 0"
+  shows "Ifm vs (x#bs) (msubst2 p c t) = Ifm vs (Itm vs (x#bs) t /  Ipoly vs c #bs) p"
+proof-
+  let ?c = "Ipoly vs c"
+  from nc have anc: "allpolys isnpoly (CP c)" "allpolys isnpoly (CP (~\<^sub>p c))" 
+    by (simp_all add: polyneg_norm)
+  from nz have "?c > 0 \<or> ?c < 0" by arith
+  moreover
+  {assume c: "?c < 0"
+    from c msubstneg_I[OF lp c, of x bs t] lt[OF anc(1), of vs "x#bs"] lt[OF anc(2), of vs "x#bs"]
+    have ?thesis by (auto simp add: msubst2_def)}
+  moreover
+  {assume c: "?c > 0"
+    from c msubstpos_I[OF lp c, of x bs t] lt[OF anc(1), of vs "x#bs"] lt[OF anc(2), of vs "x#bs"]
+    have ?thesis by (auto simp add: msubst2_def)}
+  ultimately show ?thesis by blast
+qed
+
+term msubsteq2
+lemma msubsteq2_nb: "tmbound0 t \<Longrightarrow> islin (Eq (CNP 0 a r)) \<Longrightarrow> bound0 (msubsteq2 c t a r)"
+  by (simp add: msubsteq2_def)
+
+lemma msubstltpos_nb: "tmbound0 t \<Longrightarrow> islin (Lt (CNP 0 a r)) \<Longrightarrow> bound0 (msubstltpos c t a r)"
+  by (simp add: msubstltpos_def)
+lemma msubstltneg_nb: "tmbound0 t \<Longrightarrow> islin (Lt (CNP 0 a r)) \<Longrightarrow> bound0 (msubstltneg c t a r)"
+  by (simp add: msubstltneg_def)
+
+lemma msubstlepos_nb: "tmbound0 t \<Longrightarrow> islin (Le (CNP 0 a r)) \<Longrightarrow> bound0 (msubstlepos c t a r)"
+  by (simp add: msubstlepos_def)
+lemma msubstleneg_nb: "tmbound0 t \<Longrightarrow> islin (Le (CNP 0 a r)) \<Longrightarrow> bound0 (msubstleneg c t a r)"
+  by (simp add: msubstleneg_def)
+
+lemma msubstpos_nb: assumes lp: "islin p" and tnb: "tmbound0 t"
+  shows "bound0 (msubstpos p c t)"
+using lp tnb
+by (induct p c t rule: msubstpos.induct, auto simp add: msubsteq2_nb msubstltpos_nb msubstlepos_nb)
+
+lemma msubstneg_nb: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})" and lp: "islin p" and tnb: "tmbound0 t"
+  shows "bound0 (msubstneg p c t)"
+using lp tnb
+by (induct p c t rule: msubstneg.induct, auto simp add: msubsteq2_nb msubstltneg_nb msubstleneg_nb)
+
+lemma msubst2_nb: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})" and lp: "islin p" and tnb: "tmbound0 t"
+  shows "bound0 (msubst2 p c t)"
+using lp tnb
+by (simp add: msubst2_def msubstneg_nb msubstpos_nb conj_nb disj_nb lt_nb simpfm_bound0)
+    
+lemma of_int2: "of_int 2 = 1 + 1"
+proof-
+  have "(2::int) = 1 + 1" by simp
+  hence "of_int 2 = of_int (1 + 1)" by simp
+  thus ?thesis unfolding of_int_add by simp
+qed
+
+lemma of_int_minus2: "of_int (-2) = - (1 + 1)"
+proof-
+  have th: "(-2::int) = - 2" by simp
+  show ?thesis unfolding th by (simp only: of_int_minus of_int2)
+qed
+
+
+lemma islin_qf: "islin p \<Longrightarrow> qfree p"
+  by (induct p rule: islin.induct, auto simp add: bound0_qf)
+lemma fr_eq_msubst2: 
+  assumes lp: "islin p"
+  shows "(\<exists> x. Ifm vs (x#bs) p) \<longleftrightarrow> ((Ifm vs (x#bs) (minusinf p)) \<or> (Ifm vs (x#bs) (plusinf p)) \<or> Ifm vs (x#bs) (subst0 (CP 0\<^sub>p) p) \<or> (\<exists>(n, t)\<in>set (uset p). Ifm vs (x# bs) (msubst2 p (n *\<^sub>p (C (-2,1))) t)) \<or> (\<exists> (c,t) \<in> set (uset p). \<exists> (d,s) \<in> set (uset p). Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))))"
+  (is "(\<exists> x. ?I x p) = (?M \<or> ?P \<or> ?Pz \<or> ?PU \<or> ?F)" is "?E = ?D")
+proof-
+  from uset_l[OF lp] have th: "\<forall>(c, s)\<in>set (uset p). isnpoly c \<and> tmbound0 s" by blast
+  let ?I = "\<lambda>p. Ifm vs (x#bs) p"
+  have n2: "isnpoly (C (-2,1))" by (simp add: isnpoly_def)
+  note eq0 = subst0[OF islin_qf[OF lp], of vs x bs "CP 0\<^sub>p", simplified]
+  
+  have eq1: "(\<exists>(n, t)\<in>set (uset p). ?I (msubst2 p (n *\<^sub>p (C (-2,1))) t)) \<longleftrightarrow> (\<exists>(n, t)\<in>set (uset p). \<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> Ifm vs (- Itm vs (x # bs) t / (\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> * (1 + 1)) # bs) p)"
+  proof-
+    {fix n t assume H: "(n, t)\<in>set (uset p)" "?I(msubst2 p (n *\<^sub>p C (-2, 1)) t)"
+      from H(1) th have "isnpoly n" by blast
+      hence nn: "isnpoly (n *\<^sub>p (C (-2,1)))" by (simp_all add: polymul_norm n2)
+      have nn': "allpolys isnpoly (CP (~\<^sub>p (n *\<^sub>p C (-2, 1))))"
+	by (simp add: polyneg_norm nn)
+      hence nn2: "\<lparr>n *\<^sub>p(C (-2,1)) \<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>n \<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" using H(2) nn' nn 
+	by (auto simp add: msubst2_def lt zero_less_mult_iff mult_less_0_iff)
+      from msubst2[OF lp nn nn2(1), of x bs t]
+      have "\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> Ifm vs (- Itm vs (x # bs) t / (\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> * (1 + 1)) # bs) p"
+	using H(2) nn2 by (simp add: of_int_minus2 del: minus_add_distrib)}
+    moreover
+    {fix n t assume H: "(n, t)\<in>set (uset p)" "\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "Ifm vs (- Itm vs (x # bs) t / (\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> * (1 + 1)) # bs) p"
+      from H(1) th have "isnpoly n" by blast
+      hence nn: "isnpoly (n *\<^sub>p (C (-2,1)))" "\<lparr>n *\<^sub>p(C (-2,1)) \<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
+	using H(2) by (simp_all add: polymul_norm n2)
+      from msubst2[OF lp nn, of x bs t] have "?I (msubst2 p (n *\<^sub>p (C (-2,1))) t)" using H(2,3) by (simp add: of_int_minus2 del: minus_add_distrib)}
+    ultimately show ?thesis by blast
+  qed
+  have eq2: "(\<exists> (c,t) \<in> set (uset p). \<exists> (d,s) \<in> set (uset p). Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))) \<longleftrightarrow> (\<exists>(n, t)\<in>set (uset p).
+     \<exists>(m, s)\<in>set (uset p). \<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> \<lparr>m\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> Ifm vs ((- Itm vs (x # bs) t / \<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>m\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / (1 + 1) # bs) p)" 
+  proof-
+    {fix c t d s assume H: "(c,t) \<in> set (uset p)" "(d,s) \<in> set (uset p)" 
+     "Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))"
+      from H(1,2) th have "isnpoly c" "isnpoly d" by blast+
+      hence nn: "isnpoly (C (-2, 1) *\<^sub>p c*\<^sub>p d)" 
+	by (simp_all add: polymul_norm n2)
+      have stupid: "allpolys isnpoly (CP (~\<^sub>p (C (-2, 1) *\<^sub>p c *\<^sub>p d)))" "allpolys isnpoly (CP ((C (-2, 1) *\<^sub>p c *\<^sub>p d)))"
+	by (simp_all add: polyneg_norm nn)
+      have nn': "\<lparr>(C (-2, 1) *\<^sub>p c*\<^sub>p d)\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
+	using H(3) by (auto simp add: msubst2_def lt[OF stupid(1)]  lt[OF stupid(2)] zero_less_mult_iff mult_less_0_iff)
+      from msubst2[OF lp nn nn'(1), of x bs ] H(3) nn'
+      have "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / (1 + 1) # bs) p" 
+	apply (simp add: add_divide_distrib of_int_minus2 del: minus_add_distrib)
+	by (simp add: mult_commute)}
+    moreover
+    {fix c t d s assume H: "(c,t) \<in> set (uset p)" "(d,s) \<in> set (uset p)" 
+      "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / (1 + 1) # bs) p"
+     from H(1,2) th have "isnpoly c" "isnpoly d" by blast+
+      hence nn: "isnpoly (C (-2, 1) *\<^sub>p c*\<^sub>p d)" "\<lparr>(C (-2, 1) *\<^sub>p c*\<^sub>p d)\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
+	using H(3,4) by (simp_all add: polymul_norm n2)
+      from msubst2[OF lp nn, of x bs ] H(3,4,5) 
+      have "Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))" apply (simp add: add_divide_distrib of_int_minus2 del: minus_add_distrib)by (simp add: mult_commute)}
+    ultimately show ?thesis by blast
+  qed
+  from fr_eq2[OF lp, of vs bs x] show ?thesis
+    unfolding eq0 eq1 eq2 by blast  
+qed
+
+definition 
+"ferrack2 p \<equiv> let q = simpfm p ; mp = minusinf q ; pp = plusinf q
+ in if (mp = T \<or> pp = T) then T 
+  else (let U = remdps (uset  q)
+    in decr0 (list_disj [mp, pp, simpfm (subst0 (CP 0\<^sub>p) q), evaldjf (\<lambda>(c,t). msubst2 q (c *\<^sub>p C (-2, 1)) t) U, 
+   evaldjf (\<lambda>((b,a),(d,c)). msubst2 q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) (alluopairs U)]))"
+
+definition "frpar2 p = simpfm (qelim (prep p) ferrack2)"
+
+lemma ferrack2: assumes qf: "qfree p"
+  shows "qfree (ferrack2 p) \<and> ((Ifm vs bs (ferrack2 p)) = (Ifm vs bs (E p)))"
+  (is "_ \<and> (?rhs = ?lhs)")
+proof-
+  let ?J = "\<lambda> x p. Ifm vs (x#bs) p"
+  let ?N = "\<lambda> t. Ipoly vs t"
+  let ?Nt = "\<lambda>x t. Itm vs (x#bs) t"
+  let ?q = "simpfm p" 
+  let ?qz = "subst0 (CP 0\<^sub>p) ?q"
+  let ?U = "remdps(uset ?q)"
+  let ?Up = "alluopairs ?U"
+  let ?mp = "minusinf ?q"
+  let ?pp = "plusinf ?q"
+  let ?I = "\<lambda>p. Ifm vs (x#bs) p"
+  from simpfm_lin[OF qf] simpfm_qf[OF qf] have lq: "islin ?q" and q_qf: "qfree ?q" .
+  from minusinf_nb[OF lq] plusinf_nb[OF lq] have mp_nb: "bound0 ?mp" and pp_nb: "bound0 ?pp" .
+  from bound0_qf[OF mp_nb] bound0_qf[OF pp_nb] have mp_qf: "qfree ?mp" and pp_qf: "qfree ?pp" .
+  from uset_l[OF lq] have U_l: "\<forall>(c, s)\<in>set ?U. isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s"
+    by simp
+  have bnd0: "\<forall>x \<in> set ?U. bound0 ((\<lambda>(c,t). msubst2 ?q (c *\<^sub>p C (-2, 1)) t) x)" 
+  proof-
+    {fix c t assume ct: "(c,t) \<in> set ?U"
+      hence tnb: "tmbound0 t" using U_l by blast
+      from msubst2_nb[OF lq tnb]
+      have "bound0 ((\<lambda>(c,t). msubst2 ?q (c *\<^sub>p C (-2, 1)) t) (c,t))" by simp}
+    thus ?thesis by auto
+  qed
+  have bnd1: "\<forall>x \<in> set ?Up. bound0 ((\<lambda>((b,a),(d,c)). msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) x)" 
+  proof-
+    {fix b a d c assume badc: "((b,a),(d,c)) \<in> set ?Up"
+      from badc U_l alluopairs_set1[of ?U] 
+      have nb: "tmbound0 (Add (Mul d a) (Mul b c))" by auto
+      from msubst2_nb[OF lq nb] have "bound0 ((\<lambda>((b,a),(d,c)). msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) ((b,a),(d,c)))" by simp}
+    thus ?thesis by auto
+  qed
+  have stupid: "bound0 F" by simp
+  let ?R = "list_disj [?mp, ?pp, simpfm (subst0 (CP 0\<^sub>p) ?q), evaldjf (\<lambda>(c,t). msubst2 ?q (c *\<^sub>p C (-2, 1)) t) ?U, 
+   evaldjf (\<lambda>((b,a),(d,c)). msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) (alluopairs ?U)]"
+  from subst0_nb[of "CP 0\<^sub>p" ?q] q_qf evaldjf_bound0[OF bnd1] evaldjf_bound0[OF bnd0] mp_nb pp_nb stupid
+  have nb: "bound0 ?R "
+    by (simp add: list_disj_def disj_nb0 simpfm_bound0)
+  let ?s = "\<lambda>((b,a),(d,c)). msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))"
+
+  {fix b a d c assume baU: "(b,a) \<in> set ?U" and dcU: "(d,c) \<in> set ?U"
+    from U_l baU dcU have norm: "isnpoly b" "isnpoly d" "isnpoly (C (-2, 1))" 
+      by auto (simp add: isnpoly_def)
+    have norm2: "isnpoly (C (-2, 1) *\<^sub>p b*\<^sub>p d)" "isnpoly (C (-2, 1) *\<^sub>p d*\<^sub>p b)"
+      using norm by (simp_all add: polymul_norm)
+    have stupid: "allpolys isnpoly (CP (C (-2, 1) *\<^sub>p b*\<^sub>p d))" "allpolys isnpoly (CP (C (-2, 1) *\<^sub>p d*\<^sub>p b))" "allpolys isnpoly (CP (~\<^sub>p(C (-2, 1) *\<^sub>p b*\<^sub>p d)))" "allpolys isnpoly (CP (~\<^sub>p(C (-2, 1) *\<^sub>p d*\<^sub>p b)))"
+      by (simp_all add: polyneg_norm norm2)
+    have "?I (msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) = ?I (msubst2 ?q (C (-2, 1) *\<^sub>p d*\<^sub>p b) (Add (Mul b c) (Mul d a)))" (is "?lhs \<longleftrightarrow> ?rhs")
+    proof
+      assume H: ?lhs
+      hence z: "\<lparr>C (-2, 1) *\<^sub>p b *\<^sub>p d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>C (-2, 1) *\<^sub>p d *\<^sub>p b\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" 
+	by (auto simp add: msubst2_def lt[OF stupid(3)] lt[OF stupid(1)] mult_less_0_iff zero_less_mult_iff)
+      from msubst2[OF lq norm2(1) z(1), of x bs] 
+	msubst2[OF lq norm2(2) z(2), of x bs] H 
+      show ?rhs by (simp add: ring_simps)
+    next
+      assume H: ?rhs
+      hence z: "\<lparr>C (-2, 1) *\<^sub>p b *\<^sub>p d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>C (-2, 1) *\<^sub>p d *\<^sub>p b\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" 
+	by (auto simp add: msubst2_def lt[OF stupid(4)] lt[OF stupid(2)] mult_less_0_iff zero_less_mult_iff)
+      from msubst2[OF lq norm2(1) z(1), of x bs] 
+	msubst2[OF lq norm2(2) z(2), of x bs] H 
+      show ?lhs by (simp add: ring_simps)
+    qed}
+  hence th0: "\<forall>x \<in> set ?U. \<forall>y \<in> set ?U. ?I (?s (x, y)) \<longleftrightarrow> ?I (?s (y, x))"
+    by clarsimp
+
+  have "?lhs \<longleftrightarrow> (\<exists>x. Ifm vs (x#bs) ?q)" by simp
+  also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> ?I (subst0 (CP 0\<^sub>p) ?q) \<or> (\<exists>(n,t) \<in> set ?U. ?I (msubst2 ?q (n *\<^sub>p C (-2, 1)) t)) \<or> (\<exists>(b, a)\<in>set ?U. \<exists>(d, c)\<in>set ?U. ?I (msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))))"
+    using fr_eq_msubst2[OF lq, of vs bs x] by simp
+  also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> ?I (subst0 (CP 0\<^sub>p) ?q) \<or> (\<exists>(n,t) \<in> set ?U. ?I (msubst2 ?q (n *\<^sub>p C (-2, 1)) t)) \<or> (\<exists> x\<in>set ?U. \<exists> y \<in>set ?U. ?I (?s (x,y)))"
+    by (simp add: split_def)
+  also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> ?I (subst0 (CP 0\<^sub>p) ?q) \<or> (\<exists>(n,t) \<in> set ?U. ?I (msubst2 ?q (n *\<^sub>p C (-2, 1)) t)) \<or> (\<exists> (x,y) \<in> set ?Up. ?I (?s (x,y)))"
+    using alluopairs_ex[OF th0] by simp 
+  also have "\<dots> \<longleftrightarrow> ?I ?R" 
+    by (simp add: list_disj_def evaldjf_ex split_def)
+  also have "\<dots> \<longleftrightarrow> ?rhs"
+    unfolding ferrack2_def
+    apply (cases "?mp = T") 
+    apply (simp add: list_disj_def)
+    apply (cases "?pp = T") 
+    apply (simp add: list_disj_def)
+    by (simp_all add: Let_def decr0[OF nb])
+  finally show ?thesis using decr0_qf[OF nb]  
+    by (simp  add: ferrack2_def Let_def)
+qed
+
+lemma frpar2: "qfree (frpar2 p) \<and> (Ifm vs bs (frpar2 p) \<longleftrightarrow> Ifm vs bs p)"
+proof-
+  from ferrack2 have th: "\<forall>bs p. qfree p \<longrightarrow> qfree (ferrack2 p) \<and> Ifm vs bs (ferrack2 p) = Ifm vs bs (E p)" by blast
+  from qelim[OF th, of "prep p" bs] 
+show ?thesis  unfolding frpar2_def by (auto simp add: prep)
+qed
+
+code_module FRPar
+  contains 
+  frpar = "frpar"
+  frpar2 = "frpar2"
+  test = "%x . frpar (E(Lt (Mul 1\<^sub>p (Bound 0))))"
+
+ML{* 
+
+structure ReflectedFRPar = 
+struct
+
+val bT = HOLogic.boolT;
+fun num rT x = HOLogic.mk_number rT x;
+fun rrelT rT = [rT,rT] ---> rT;
+fun rrT rT = [rT, rT] ---> bT;
+fun divt rT = Const(@{const_name "HOL.divide"},rrelT rT);
+fun timest rT = Const(@{const_name "HOL.times"},rrelT rT);
+fun plust rT = Const(@{const_name "HOL.plus"},rrelT rT);
+fun minust rT = Const(@{const_name "HOL.minus"},rrelT rT);
+fun uminust rT = Const(@{const_name "HOL.uminus"}, rT --> rT);
+fun powt rT = Const(@{const_name "power"}, [rT,@{typ "nat"}] ---> rT);
+val brT = [bT, bT] ---> bT;
+val nott = @{term "Not"};
+val conjt = @{term "op &"};
+val disjt = @{term "op |"};
+val impt = @{term "op -->"};
+val ifft = @{term "op = :: bool => _"}
+fun llt rT = Const(@{const_name "HOL.less"},rrT rT);
+fun lle rT = Const(@{const_name "HOL.less"},rrT rT);
+fun eqt rT = Const("op =",rrT rT);
+fun rz rT = Const(@{const_name "HOL.zero"},rT);
+
+fun dest_nat t = case t of
+  Const ("Suc",_)$t' => 1 + dest_nat t'
+| _ => (snd o HOLogic.dest_number) t;
+
+fun num_of_term m t = 
+ case t of
+   Const(@{const_name "uminus"},_)$t => FRPar.Neg (num_of_term m t)
+ | Const(@{const_name "HOL.plus"},_)$a$b => FRPar.Add (num_of_term m a, num_of_term m b)
+ | Const(@{const_name "HOL.minus"},_)$a$b => FRPar.Sub (num_of_term m a, num_of_term m b)
+ | Const(@{const_name "HOL.times"},_)$a$b => FRPar.Mul (num_of_term m a, num_of_term m b)
+ | Const(@{const_name "power"},_)$a$n => FRPar.Pw (num_of_term m a, dest_nat n)
+ | Const(@{const_name "HOL.divide"},_)$a$b => FRPar.C (HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd)
+ | _ => (FRPar.C (HOLogic.dest_number t |> snd,1) 
+         handle TERM _ => FRPar.Bound (AList.lookup (op aconv) m t |> valOf));
+
+fun tm_of_term m m' t = 
+ case t of
+   Const(@{const_name "uminus"},_)$t => FRPar.tm_Neg (tm_of_term m m' t)
+ | Const(@{const_name "HOL.plus"},_)$a$b => FRPar.tm_Add (tm_of_term m m' a, tm_of_term m m' b)
+ | Const(@{const_name "HOL.minus"},_)$a$b => FRPar.tm_Sub (tm_of_term m m' a, tm_of_term m m' b)
+ | Const(@{const_name "HOL.times"},_)$a$b => FRPar.tm_Mul (num_of_term m' a, tm_of_term m m' b)
+ | _ => (FRPar.CP (num_of_term m' t) 
+         handle TERM _ => FRPar.tm_Bound (AList.lookup (op aconv) m t |> valOf)
+              | Option => FRPar.tm_Bound (AList.lookup (op aconv) m t |> valOf));
+
+fun term_of_num T m t = 
+ case t of
+  FRPar.C (a,b) => (if b = 1 then num T a else if b=0 then (rz T) 
+                                        else (divt T) $ num T a $ num T b)
+| FRPar.Bound i => AList.lookup (op = : int*int -> bool) m i |> valOf
+| FRPar.Add(a,b) => (plust T)$(term_of_num T m a)$(term_of_num T m b)
+| FRPar.Mul(a,b) => (timest T)$(term_of_num T m a)$(term_of_num T m b)
+| FRPar.Sub(a,b) => (minust T)$(term_of_num T m a)$(term_of_num T m b)
+| FRPar.Neg a => (uminust T)$(term_of_num T m a)
+| FRPar.Pw(a,n) => (powt T)$(term_of_num T m t)$(HOLogic.mk_number HOLogic.natT n)
+| FRPar.CN(c,n,p) => term_of_num T m (FRPar.Add(c,FRPar.Mul(FRPar.Bound n, p)))
+| _ => error "term_of_num: Unknown term";
+
+fun term_of_tm T m m' t = 
+ case t of
+  FRPar.CP p => term_of_num T m' p
+| FRPar.tm_Bound i => AList.lookup (op = : int*int -> bool) m i |> valOf
+| FRPar.tm_Add(a,b) => (plust T)$(term_of_tm T m m' a)$(term_of_tm T m m' b)
+| FRPar.tm_Mul(a,b) => (timest T)$(term_of_num T m' a)$(term_of_tm T m m' b)
+| FRPar.tm_Sub(a,b) => (minust T)$(term_of_tm T m m' a)$(term_of_tm T m m' b)
+| FRPar.tm_Neg a => (uminust T)$(term_of_tm T m m' a)
+| FRPar.CNP(n,c,p) => term_of_tm T m m' (FRPar.tm_Add(FRPar.tm_Mul(c, FRPar.tm_Bound n), p))
+| _ => error "term_of_tm: Unknown term";
+
+fun fm_of_term m m' fm = 
+ case fm of
+    Const("True",_) => FRPar.T
+  | Const("False",_) => FRPar.F
+  | Const("Not",_)$p => FRPar.NOT (fm_of_term m m' p)
+  | Const("op &",_)$p$q => FRPar.And(fm_of_term m m' p, fm_of_term m m' q)
+  | Const("op |",_)$p$q => FRPar.Or(fm_of_term m m' p, fm_of_term m m' q)
+  | Const("op -->",_)$p$q => FRPar.Imp(fm_of_term m m' p, fm_of_term m m' q)
+  | Const("op =",ty)$p$q => 
+       if domain_type ty = bT then FRPar.Iff(fm_of_term m m' p, fm_of_term m m' q)
+       else FRPar.Eq (FRPar.tm_Sub(tm_of_term m m' p, tm_of_term m m' q))
+  | Const(@{const_name "HOL.less"},_)$p$q => 
+        FRPar.Lt (FRPar.tm_Sub(tm_of_term m m' p, tm_of_term m m' q))
+  | Const(@{const_name "HOL.less_eq"},_)$p$q => 
+        FRPar.Le (FRPar.tm_Sub(tm_of_term m m' p, tm_of_term m m' q))
+  | Const("Ex",_)$Abs(xn,xT,p) => 
+     let val (xn', p') =  variant_abs (xn,xT,p)
+         val x = Free(xn',xT)
+         fun incr i = i + 1
+         val m0 = (x,0):: (map (apsnd incr) m)
+      in FRPar.E (fm_of_term m0 m' p') end
+  | Const("All",_)$Abs(xn,xT,p) => 
+     let val (xn', p') =  variant_abs (xn,xT,p)
+         val x = Free(xn',xT)
+         fun incr i = i + 1
+         val m0 = (x,0):: (map (apsnd incr) m)
+      in FRPar.A (fm_of_term m0 m' p') end
+  | _ => error "fm_of_term";
+
+
+fun term_of_fm T m m' t = 
+  case t of
+    FRPar.T => Const("True",bT)
+  | FRPar.F => Const("False",bT)
+  | FRPar.NOT p => nott $ (term_of_fm T m m' p)
+  | FRPar.And (p,q) => conjt $ (term_of_fm T m m' p) $ (term_of_fm T m m' q)
+  | FRPar.Or (p,q) => disjt $ (term_of_fm T m m' p) $ (term_of_fm T m m' q)
+  | FRPar.Imp (p,q) => impt $ (term_of_fm T m m' p) $ (term_of_fm T m m' q)
+  | FRPar.Iff (p,q) => ifft $ (term_of_fm T m m' p) $ (term_of_fm T m m' q)
+  | FRPar.Lt p => (llt T) $ (term_of_tm T m m' p) $ (rz T)
+  | FRPar.Le p => (lle T) $ (term_of_tm T m m' p) $ (rz T)
+  | FRPar.Eq p => (eqt T) $ (term_of_tm T m m' p) $ (rz T)
+  | FRPar.NEq p => nott $ ((eqt T) $ (term_of_tm T m m' p) $ (rz T))
+  | _ => error "term_of_fm: quantifiers!!!!???";
+
+fun frpar_oracle (T,m, m', fm) = 
+ let 
+   val t = HOLogic.dest_Trueprop fm
+   val im = 0 upto (length m - 1)
+   val im' = 0 upto (length m' - 1)   
+ in HOLogic.mk_Trueprop (HOLogic.mk_eq(t, term_of_fm T (im ~~ m) (im' ~~ m')  
+                                                     (FRPar.frpar (fm_of_term (m ~~ im) (m' ~~ im') t))))
+ end;
+
+fun frpar_oracle2 (T,m, m', fm) = 
+ let 
+   val t = HOLogic.dest_Trueprop fm
+   val im = 0 upto (length m - 1)
+   val im' = 0 upto (length m' - 1)   
+ in HOLogic.mk_Trueprop (HOLogic.mk_eq(t, term_of_fm T (im ~~ m) (im' ~~ m')  
+                                                     (FRPar.frpar2 (fm_of_term (m ~~ im) (m' ~~ im') t))))
+ end;
+
+end;
+
+
+*}
+
+oracle frpar_oracle = {* fn (ty, ts, ts', ct) => 
+ let 
+  val thy = Thm.theory_of_cterm ct
+ in cterm_of thy (ReflectedFRPar.frpar_oracle (ty,ts, ts', term_of ct))
+ end *}
+
+oracle frpar_oracle2 = {* fn (ty, ts, ts', ct) => 
+ let 
+  val thy = Thm.theory_of_cterm ct
+ in cterm_of thy (ReflectedFRPar.frpar_oracle2 (ty,ts, ts', term_of ct))
+ end *}
+
+ML{* 
+structure FRParTac = 
+struct
+
+fun frpar_tac T ps ctxt i = 
+ (ObjectLogic.full_atomize_tac i) 
+ THEN (fn st =>
+  let
+    val g = List.nth (cprems_of st, i - 1)
+    val thy = ProofContext.theory_of ctxt
+    val fs = subtract (op aconv) (map Free (Term.add_frees (term_of g) [])) ps
+    val th = frpar_oracle (T, fs,ps, (* Pattern.eta_long [] *)g)
+  in rtac (th RS iffD2) i st end);
+
+fun frpar2_tac T ps ctxt i = 
+ (ObjectLogic.full_atomize_tac i) 
+ THEN (fn st =>
+  let
+    val g = List.nth (cprems_of st, i - 1)
+    val thy = ProofContext.theory_of ctxt
+    val fs = subtract (op aconv) (map Free (Term.add_frees (term_of g) [])) ps
+    val th = frpar_oracle2 (T, fs,ps, (* Pattern.eta_long [] *)g)
+  in rtac (th RS iffD2) i st end);
+
+end;
+
+*}
+
+method_setup frpar = {*
+let
+ fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ()
+ fun simple_keyword k = Scan.lift (Args.$$$ k) >> K ()
+ val parsN = "pars"
+ val typN = "type"
+ val any_keyword = keyword parsN || keyword typN
+ val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat
+ val cterms = thms >> map Drule.dest_term;
+ val terms = Scan.repeat (Scan.unless any_keyword Args.term)
+ val typ = Scan.unless any_keyword Args.typ
+in
+ (keyword typN |-- typ) -- (keyword parsN |-- terms) >>
+  (fn (T,ps) => fn ctxt => SIMPLE_METHOD' (FRParTac.frpar_tac T ps ctxt))
+end
+*} "Parametric QE for linear Arithmetic over fields, Version 1"
+
+method_setup frpar2 = {*
+let
+ fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ()
+ fun simple_keyword k = Scan.lift (Args.$$$ k) >> K ()
+ val parsN = "pars"
+ val typN = "type"
+ val any_keyword = keyword parsN || keyword typN
+ val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat
+ val cterms = thms >> map Drule.dest_term;
+ val terms = Scan.repeat (Scan.unless any_keyword Args.term)
+ val typ = Scan.unless any_keyword Args.typ
+in
+ (keyword typN |-- typ) -- (keyword parsN |-- terms) >>
+  (fn (T,ps) => fn ctxt => SIMPLE_METHOD' (FRParTac.frpar2_tac T ps ctxt))
+end
+*} "Parametric QE for linear Arithmetic over fields, Version 2"
+
+
+lemma "\<exists>(x::'a::{division_by_zero,ordered_field,number_ring}). y \<noteq> -1 \<longrightarrow> (y + 1)*x < 0"
+  apply (frpar type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "y::'a::{division_by_zero,ordered_field,number_ring}")
+  apply (simp add: ring_simps)
+  apply (rule spec[where x=y])
+  apply (frpar type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "z::'a::{division_by_zero,ordered_field,number_ring}")
+  by simp
+
+text{* Collins/Jones Problem *}
+(*
+lemma "\<exists>(r::'a::{division_by_zero,ordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0"
+proof-
+  have "(\<exists>(r::'a::{division_by_zero,ordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0) \<longleftrightarrow> (\<exists>(r::'a::{division_by_zero,ordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < 2 *(a^2 + b^2) - (3*(a^2 + b^2)) * r + (2*a)*r \<and> 2*(a^2 + b^2) - (3*(a^2 + b^2) - 4*a + 1)*r - 2*a < 0)" (is "?lhs \<longleftrightarrow> ?rhs")
+by (simp add: ring_simps)
+have "?rhs"
+
+  apply (frpar type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "a::'a::{division_by_zero,ordered_field,number_ring}" "b::'a::{division_by_zero,ordered_field,number_ring}")
+  apply (simp add: ring_simps)
+oops
+*)
+(*
+lemma "ALL (x::'a::{division_by_zero,ordered_field,number_ring}) y. (1 - t)*x \<le> (1+t)*y \<and> (1 - t)*y \<le> (1+t)*x --> 0 \<le> y"
+apply (frpar type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "t::'a::{division_by_zero,ordered_field,number_ring}")
+oops
+*)
+
+lemma "\<exists>(x::'a::{division_by_zero,ordered_field,number_ring}). y \<noteq> -1 \<longrightarrow> (y + 1)*x < 0"
+  apply (frpar2 type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "y::'a::{division_by_zero,ordered_field,number_ring}")
+  apply (simp add: ring_simps)
+  apply (rule spec[where x=y])
+  apply (frpar2 type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "z::'a::{division_by_zero,ordered_field,number_ring}")
+  by simp
+
+text{* Collins/Jones Problem *}
+
+(*
+lemma "\<exists>(r::'a::{division_by_zero,ordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0"
+proof-
+  have "(\<exists>(r::'a::{division_by_zero,ordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0) \<longleftrightarrow> (\<exists>(r::'a::{division_by_zero,ordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < 2 *(a^2 + b^2) - (3*(a^2 + b^2)) * r + (2*a)*r \<and> 2*(a^2 + b^2) - (3*(a^2 + b^2) - 4*a + 1)*r - 2*a < 0)" (is "?lhs \<longleftrightarrow> ?rhs")
+by (simp add: ring_simps)
+have "?rhs"
+  apply (frpar2 type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "a::'a::{division_by_zero,ordered_field,number_ring}" "b::'a::{division_by_zero,ordered_field,number_ring}")
+  apply simp
+oops
+*)
+
+(*
+lemma "ALL (x::'a::{division_by_zero,ordered_field,number_ring}) y. (1 - t)*x \<le> (1+t)*y \<and> (1 - t)*y \<le> (1+t)*x --> 0 \<le> y"
+apply (frpar2 type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "t::'a::{division_by_zero,ordered_field,number_ring}")
+apply (simp add: field_simps linorder_neq_iff[symmetric])
+apply ferrack
+oops
+*)
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Decision_Procs/Polynomial_List.thy	Mon Oct 26 09:03:57 2009 +0100
@@ -0,0 +1,783 @@
+(*  Title:       HOL/Decision_Procs/Polynomial_List.thy
+    Author:      Amine Chaieb
+*)
+
+header{*Univariate Polynomials as Lists *}
+
+theory Polynomial_List
+imports Main
+begin
+
+text{* Application of polynomial as a real function. *}
+
+consts poly :: "'a list => 'a  => ('a::{comm_ring})"
+primrec
+  poly_Nil:  "poly [] x = 0"
+  poly_Cons: "poly (h#t) x = h + x * poly t x"
+
+
+subsection{*Arithmetic Operations on Polynomials*}
+
+text{*addition*}
+consts padd :: "['a list, 'a list] => ('a::comm_ring_1) list"  (infixl "+++" 65)
+primrec
+  padd_Nil:  "[] +++ l2 = l2"
+  padd_Cons: "(h#t) +++ l2 = (if l2 = [] then h#t
+                            else (h + hd l2)#(t +++ tl l2))"
+
+text{*Multiplication by a constant*}
+consts cmult :: "['a :: comm_ring_1, 'a list] => 'a list"  (infixl "%*" 70)
+primrec
+   cmult_Nil:  "c %* [] = []"
+   cmult_Cons: "c %* (h#t) = (c * h)#(c %* t)"
+
+text{*Multiplication by a polynomial*}
+consts pmult :: "['a list, 'a list] => ('a::comm_ring_1) list"  (infixl "***" 70)
+primrec
+   pmult_Nil:  "[] *** l2 = []"
+   pmult_Cons: "(h#t) *** l2 = (if t = [] then h %* l2
+                              else (h %* l2) +++ ((0) # (t *** l2)))"
+
+text{*Repeated multiplication by a polynomial*}
+consts mulexp :: "[nat, 'a list, 'a  list] => ('a ::comm_ring_1) list"
+primrec
+   mulexp_zero:  "mulexp 0 p q = q"
+   mulexp_Suc:   "mulexp (Suc n) p q = p *** mulexp n p q"
+
+text{*Exponential*}
+consts pexp :: "['a list, nat] => ('a::comm_ring_1) list"  (infixl "%^" 80)
+primrec
+   pexp_0:   "p %^ 0 = [1]"
+   pexp_Suc: "p %^ (Suc n) = p *** (p %^ n)"
+
+text{*Quotient related value of dividing a polynomial by x + a*}
+(* Useful for divisor properties in inductive proofs *)
+consts "pquot" :: "['a list, 'a::field] => 'a list"
+primrec
+   pquot_Nil:  "pquot [] a= []"
+   pquot_Cons: "pquot (h#t) a = (if t = [] then [h]
+                   else (inverse(a) * (h - hd( pquot t a)))#(pquot t a))"
+
+
+text{*normalization of polynomials (remove extra 0 coeff)*}
+consts pnormalize :: "('a::comm_ring_1) list => 'a list"
+primrec
+   pnormalize_Nil:  "pnormalize [] = []"
+   pnormalize_Cons: "pnormalize (h#p) = (if ( (pnormalize p) = [])
+                                     then (if (h = 0) then [] else [h])
+                                     else (h#(pnormalize p)))"
+
+definition "pnormal p = ((pnormalize p = p) \<and> p \<noteq> [])"
+definition "nonconstant p = (pnormal p \<and> (\<forall>x. p \<noteq> [x]))"
+text{*Other definitions*}
+
+definition
+  poly_minus :: "'a list => ('a :: comm_ring_1) list"      ("-- _" [80] 80) where
+  "-- p = (- 1) %* p"
+
+definition
+  divides :: "[('a::comm_ring_1) list, 'a list] => bool"  (infixl "divides" 70) where
+  "p1 divides p2 = (\<exists>q. poly p2 = poly(p1 *** q))"
+
+definition
+  order :: "('a::comm_ring_1) => 'a list => nat" where
+    --{*order of a polynomial*}
+  "order a p = (SOME n. ([-a, 1] %^ n) divides p &
+                      ~ (([-a, 1] %^ (Suc n)) divides p))"
+
+definition
+  degree :: "('a::comm_ring_1) list => nat" where
+     --{*degree of a polynomial*}
+  "degree p = length (pnormalize p) - 1"
+
+definition
+  rsquarefree :: "('a::comm_ring_1) list => bool" where
+     --{*squarefree polynomials --- NB with respect to real roots only.*}
+  "rsquarefree p = (poly p \<noteq> poly [] &
+                     (\<forall>a. (order a p = 0) | (order a p = 1)))"
+
+lemma padd_Nil2: "p +++ [] = p"
+by (induct p) auto
+declare padd_Nil2 [simp]
+
+lemma padd_Cons_Cons: "(h1 # p1) +++ (h2 # p2) = (h1 + h2) # (p1 +++ p2)"
+by auto
+
+lemma pminus_Nil: "-- [] = []"
+by (simp add: poly_minus_def)
+declare pminus_Nil [simp]
+
+lemma pmult_singleton: "[h1] *** p1 = h1 %* p1"
+by simp
+
+lemma poly_ident_mult: "1 %* t = t"
+by (induct "t", auto)
+declare poly_ident_mult [simp]
+
+lemma poly_simple_add_Cons: "[a] +++ ((0)#t) = (a#t)"
+by simp
+declare poly_simple_add_Cons [simp]
+
+text{*Handy general properties*}
+
+lemma padd_commut: "b +++ a = a +++ b"
+apply (subgoal_tac "\<forall>a. b +++ a = a +++ b")
+apply (induct_tac [2] "b", auto)
+apply (rule padd_Cons [THEN ssubst])
+apply (case_tac "aa", auto)
+done
+
+lemma padd_assoc [rule_format]: "\<forall>b c. (a +++ b) +++ c = a +++ (b +++ c)"
+apply (induct "a", simp, clarify)
+apply (case_tac b, simp_all)
+done
+
+lemma poly_cmult_distr [rule_format]:
+     "\<forall>q. a %* ( p +++ q) = (a %* p +++ a %* q)"
+apply (induct "p", simp, clarify) 
+apply (case_tac "q")
+apply (simp_all add: right_distrib)
+done
+
+lemma pmult_by_x[simp]: "[0, 1] *** t = ((0)#t)"
+apply (induct "t", simp)
+by (auto simp add: mult_zero_left poly_ident_mult padd_commut)
+
+
+text{*properties of evaluation of polynomials.*}
+
+lemma poly_add: "poly (p1 +++ p2) x = poly p1 x + poly p2 x"
+apply (subgoal_tac "\<forall>p2. poly (p1 +++ p2) x = poly (p1) x + poly (p2) x")
+apply (induct_tac [2] "p1", auto)
+apply (case_tac "p2")
+apply (auto simp add: right_distrib)
+done
+
+lemma poly_cmult: "poly (c %* p) x = c * poly p x"
+apply (induct "p") 
+apply (case_tac [2] "x=0")
+apply (auto simp add: right_distrib mult_ac)
+done
+
+lemma poly_minus: "poly (-- p) x = - (poly p x)"
+apply (simp add: poly_minus_def)
+apply (auto simp add: poly_cmult)
+done
+
+lemma poly_mult: "poly (p1 *** p2) x = poly p1 x * poly p2 x"
+apply (subgoal_tac "\<forall>p2. poly (p1 *** p2) x = poly p1 x * poly p2 x")
+apply (simp (no_asm_simp))
+apply (induct "p1")
+apply (auto simp add: poly_cmult)
+apply (case_tac p1)
+apply (auto simp add: poly_cmult poly_add left_distrib right_distrib mult_ac)
+done
+
+lemma poly_exp: "poly (p %^ n) (x::'a::comm_ring_1) = (poly p x) ^ n"
+apply (induct "n")
+apply (auto simp add: poly_cmult poly_mult power_Suc)
+done
+
+text{*More Polynomial Evaluation Lemmas*}
+
+lemma poly_add_rzero: "poly (a +++ []) x = poly a x"
+by simp
+declare poly_add_rzero [simp]
+
+lemma poly_mult_assoc: "poly ((a *** b) *** c) x = poly (a *** (b *** c)) x"
+  by (simp add: poly_mult mult_assoc)
+
+lemma poly_mult_Nil2: "poly (p *** []) x = 0"
+by (induct "p", auto)
+declare poly_mult_Nil2 [simp]
+
+lemma poly_exp_add: "poly (p %^ (n + d)) x = poly( p %^ n *** p %^ d) x"
+apply (induct "n")
+apply (auto simp add: poly_mult mult_assoc)
+done
+
+subsection{*Key Property: if @{term "f(a) = 0"} then @{term "(x - a)"} divides
+ @{term "p(x)"} *}
+
+lemma lemma_poly_linear_rem: "\<forall>h. \<exists>q r. h#t = [r] +++ [-a, 1] *** q"
+apply (induct "t", safe)
+apply (rule_tac x = "[]" in exI)
+apply (rule_tac x = h in exI, simp)
+apply (drule_tac x = aa in spec, safe)
+apply (rule_tac x = "r#q" in exI)
+apply (rule_tac x = "a*r + h" in exI)
+apply (case_tac "q", auto)
+done
+
+lemma poly_linear_rem: "\<exists>q r. h#t = [r] +++ [-a, 1] *** q"
+by (cut_tac t = t and a = a in lemma_poly_linear_rem, auto)
+
+
+lemma poly_linear_divides: "(poly p a = 0) = ((p = []) | (\<exists>q. p = [-a, 1] *** q))"
+apply (auto simp add: poly_add poly_cmult right_distrib)
+apply (case_tac "p", simp) 
+apply (cut_tac h = aa and t = list and a = a in poly_linear_rem, safe)
+apply (case_tac "q", auto)
+apply (drule_tac x = "[]" in spec, simp)
+apply (auto simp add: poly_add poly_cmult add_assoc)
+apply (drule_tac x = "aa#lista" in spec, auto)
+done
+
+lemma lemma_poly_length_mult: "\<forall>h k a. length (k %* p +++  (h # (a %* p))) = Suc (length p)"
+by (induct "p", auto)
+declare lemma_poly_length_mult [simp]
+
+lemma lemma_poly_length_mult2: "\<forall>h k. length (k %* p +++  (h # p)) = Suc (length p)"
+by (induct "p", auto)
+declare lemma_poly_length_mult2 [simp]
+
+lemma poly_length_mult: "length([-a,1] *** q) = Suc (length q)"
+by auto
+declare poly_length_mult [simp]
+
+
+subsection{*Polynomial length*}
+
+lemma poly_cmult_length: "length (a %* p) = length p"
+by (induct "p", auto)
+declare poly_cmult_length [simp]
+
+lemma poly_add_length [rule_format]:
+     "\<forall>p2. length (p1 +++ p2) =
+             (if (length p1 < length p2) then length p2 else length p1)"
+apply (induct "p1", simp_all)
+apply arith
+done
+
+lemma poly_root_mult_length: "length([a,b] *** p) = Suc (length p)"
+by (simp add: poly_cmult_length poly_add_length)
+declare poly_root_mult_length [simp]
+
+lemma poly_mult_not_eq_poly_Nil: "(poly (p *** q) x \<noteq> poly [] x) =
+      (poly p x \<noteq> poly [] x & poly q x \<noteq> poly [] (x::'a::idom))"
+apply (auto simp add: poly_mult)
+done
+declare poly_mult_not_eq_poly_Nil [simp]
+
+lemma poly_mult_eq_zero_disj: "(poly (p *** q) (x::'a::idom) = 0) = (poly p x = 0 | poly q x = 0)"
+by (auto simp add: poly_mult)
+
+text{*Normalisation Properties*}
+
+lemma poly_normalized_nil: "(pnormalize p = []) --> (poly p x = 0)"
+by (induct "p", auto)
+
+text{*A nontrivial polynomial of degree n has no more than n roots*}
+
+lemma poly_roots_index_lemma0 [rule_format]:
+   "\<forall>p x. poly p x \<noteq> poly [] x & length p = n
+    --> (\<exists>i. \<forall>x. (poly p x = (0::'a::idom)) --> (\<exists>m. (m \<le> n & x = i m)))"
+apply (induct "n", safe)
+apply (rule ccontr)
+apply (subgoal_tac "\<exists>a. poly p a = 0", safe)
+apply (drule poly_linear_divides [THEN iffD1], safe)
+apply (drule_tac x = q in spec)
+apply (drule_tac x = x in spec)
+apply (simp del: poly_Nil pmult_Cons)
+apply (erule exE)
+apply (drule_tac x = "%m. if m = Suc n then a else i m" in spec, safe)
+apply (drule poly_mult_eq_zero_disj [THEN iffD1], safe)
+apply (drule_tac x = "Suc (length q)" in spec)
+apply (auto simp add: ring_simps)
+apply (drule_tac x = xa in spec)
+apply (clarsimp simp add: ring_simps)
+apply (drule_tac x = m in spec)
+apply (auto simp add:ring_simps)
+done
+lemmas poly_roots_index_lemma1 = conjI [THEN poly_roots_index_lemma0, standard]
+
+lemma poly_roots_index_length0: "poly p (x::'a::idom) \<noteq> poly [] x ==>
+      \<exists>i. \<forall>x. (poly p x = 0) --> (\<exists>n. n \<le> length p & x = i n)"
+by (blast intro: poly_roots_index_lemma1)
+
+lemma poly_roots_finite_lemma: "poly p (x::'a::idom) \<noteq> poly [] x ==>
+      \<exists>N i. \<forall>x. (poly p x = 0) --> (\<exists>n. (n::nat) < N & x = i n)"
+apply (drule poly_roots_index_length0, safe)
+apply (rule_tac x = "Suc (length p)" in exI)
+apply (rule_tac x = i in exI) 
+apply (simp add: less_Suc_eq_le)
+done
+
+
+lemma real_finite_lemma:
+  assumes P: "\<forall>x. P x --> (\<exists>n. n < length j & x = j!n)"
+  shows "finite {(x::'a::idom). P x}"
+proof-
+  let ?M = "{x. P x}"
+  let ?N = "set j"
+  have "?M \<subseteq> ?N" using P by auto
+  thus ?thesis using finite_subset by auto
+qed
+
+lemma poly_roots_index_lemma [rule_format]:
+   "\<forall>p x. poly p x \<noteq> poly [] x & length p = n
+    --> (\<exists>i. \<forall>x. (poly p x = (0::'a::{idom})) --> x \<in> set i)"
+apply (induct "n", safe)
+apply (rule ccontr)
+apply (subgoal_tac "\<exists>a. poly p a = 0", safe)
+apply (drule poly_linear_divides [THEN iffD1], safe)
+apply (drule_tac x = q in spec)
+apply (drule_tac x = x in spec)
+apply (auto simp del: poly_Nil pmult_Cons)
+apply (drule_tac x = "a#i" in spec)
+apply (auto simp only: poly_mult List.list.size)
+apply (drule_tac x = xa in spec)
+apply (clarsimp simp add: ring_simps)
+done
+
+lemmas poly_roots_index_lemma2 = conjI [THEN poly_roots_index_lemma, standard]
+
+lemma poly_roots_index_length: "poly p (x::'a::idom) \<noteq> poly [] x ==>
+      \<exists>i. \<forall>x. (poly p x = 0) --> x \<in> set i"
+by (blast intro: poly_roots_index_lemma2)
+
+lemma poly_roots_finite_lemma': "poly p (x::'a::idom) \<noteq> poly [] x ==>
+      \<exists>i. \<forall>x. (poly p x = 0) --> x \<in> set i"
+by (drule poly_roots_index_length, safe)
+
+lemma UNIV_nat_infinite: "\<not> finite (UNIV :: nat set)"
+  unfolding finite_conv_nat_seg_image
+proof(auto simp add: expand_set_eq image_iff)
+  fix n::nat and f:: "nat \<Rightarrow> nat"
+  let ?N = "{i. i < n}"
+  let ?fN = "f ` ?N"
+  let ?y = "Max ?fN + 1"
+  from nat_seg_image_imp_finite[of "?fN" "f" n] 
+  have thfN: "finite ?fN" by simp
+  {assume "n =0" hence "\<exists>x. \<forall>xa<n. x \<noteq> f xa" by auto}
+  moreover
+  {assume nz: "n \<noteq> 0"
+    hence thne: "?fN \<noteq> {}" by (auto simp add: neq0_conv)
+    have "\<forall>x\<in> ?fN. Max ?fN \<ge> x" using nz Max_ge_iff[OF thfN thne] by auto
+    hence "\<forall>x\<in> ?fN. ?y > x" by (auto simp add: less_Suc_eq_le)
+    hence "?y \<notin> ?fN" by auto
+    hence "\<exists>x. \<forall>xa<n. x \<noteq> f xa" by auto }
+  ultimately show "\<exists>x. \<forall>xa<n. x \<noteq> f xa" by blast
+qed
+
+lemma UNIV_ring_char_0_infinte: "\<not> finite (UNIV:: ('a::ring_char_0) set)"
+proof
+  assume F: "finite (UNIV :: 'a set)"
+  have th0: "of_nat ` UNIV \<subseteq> (UNIV:: 'a set)" by simp
+  from finite_subset[OF th0 F] have th: "finite (of_nat ` UNIV :: 'a set)" .
+  have th': "inj_on (of_nat::nat \<Rightarrow> 'a) (UNIV)"
+    unfolding inj_on_def by auto
+  from finite_imageD[OF th th'] UNIV_nat_infinite 
+  show False by blast
+qed
+
+lemma poly_roots_finite: "(poly p \<noteq> poly []) = 
+  finite {x. poly p x = (0::'a::{idom, ring_char_0})}"
+proof
+  assume H: "poly p \<noteq> poly []"
+  show "finite {x. poly p x = (0::'a)}"
+    using H
+    apply -
+    apply (erule contrapos_np, rule ext)
+    apply (rule ccontr)
+    apply (clarify dest!: poly_roots_finite_lemma')
+    using finite_subset
+  proof-
+    fix x i
+    assume F: "\<not> finite {x. poly p x = (0\<Colon>'a)}" 
+      and P: "\<forall>x. poly p x = (0\<Colon>'a) \<longrightarrow> x \<in> set i"
+    let ?M= "{x. poly p x = (0\<Colon>'a)}"
+    from P have "?M \<subseteq> set i" by auto
+    with finite_subset F show False by auto
+  qed
+next
+  assume F: "finite {x. poly p x = (0\<Colon>'a)}"
+  show "poly p \<noteq> poly []" using F UNIV_ring_char_0_infinte by auto  
+qed
+
+text{*Entirety and Cancellation for polynomials*}
+
+lemma poly_entire_lemma: "[| poly (p:: ('a::{idom,ring_char_0}) list) \<noteq> poly [] ; poly q \<noteq> poly [] |]
+      ==>  poly (p *** q) \<noteq> poly []"
+by (auto simp add: poly_roots_finite poly_mult Collect_disj_eq)
+
+lemma poly_entire: "(poly (p *** q) = poly ([]::('a::{idom,ring_char_0}) list)) = ((poly p = poly []) | (poly q = poly []))"
+apply (auto intro: ext dest: fun_cong simp add: poly_entire_lemma poly_mult)
+apply (blast intro: ccontr dest: poly_entire_lemma poly_mult [THEN subst])
+done
+
+lemma poly_entire_neg: "(poly (p *** q) \<noteq> poly ([]::('a::{idom,ring_char_0}) list)) = ((poly p \<noteq> poly []) & (poly q \<noteq> poly []))"
+by (simp add: poly_entire)
+
+lemma fun_eq: " (f = g) = (\<forall>x. f x = g x)"
+by (auto intro!: ext)
+
+lemma poly_add_minus_zero_iff: "(poly (p +++ -- q) = poly []) = (poly p = poly q)"
+by (auto simp add: ring_simps poly_add poly_minus_def fun_eq poly_cmult)
+
+lemma poly_add_minus_mult_eq: "poly (p *** q +++ --(p *** r)) = poly (p *** (q +++ -- r))"
+by (auto simp add: poly_add poly_minus_def fun_eq poly_mult poly_cmult right_distrib)
+
+lemma poly_mult_left_cancel: "(poly (p *** q) = poly (p *** r)) = (poly p = poly ([]::('a::{idom, ring_char_0}) list) | poly q = poly r)"
+apply (rule_tac p1 = "p *** q" in poly_add_minus_zero_iff [THEN subst])
+apply (auto intro: ext simp add: poly_add_minus_mult_eq poly_entire poly_add_minus_zero_iff)
+done
+
+lemma poly_exp_eq_zero:
+     "(poly (p %^ n) = poly ([]::('a::idom) list)) = (poly p = poly [] & n \<noteq> 0)"
+apply (simp only: fun_eq add: all_simps [symmetric]) 
+apply (rule arg_cong [where f = All]) 
+apply (rule ext)
+apply (induct_tac "n")
+apply (auto simp add: poly_mult)
+done
+declare poly_exp_eq_zero [simp]
+
+lemma poly_prime_eq_zero: "poly [a,(1::'a::comm_ring_1)] \<noteq> poly []"
+apply (simp add: fun_eq)
+apply (rule_tac x = "1 - a" in exI, simp)
+done
+declare poly_prime_eq_zero [simp]
+
+lemma poly_exp_prime_eq_zero: "(poly ([a, (1::'a::idom)] %^ n) \<noteq> poly [])"
+by auto
+declare poly_exp_prime_eq_zero [simp]
+
+text{*A more constructive notion of polynomials being trivial*}
+
+lemma poly_zero_lemma': "poly (h # t) = poly [] ==> h = (0::'a::{idom,ring_char_0}) & poly t = poly []"
+apply(simp add: fun_eq)
+apply (case_tac "h = 0")
+apply (drule_tac [2] x = 0 in spec, auto) 
+apply (case_tac "poly t = poly []", simp) 
+proof-
+  fix x
+  assume H: "\<forall>x. x = (0\<Colon>'a) \<or> poly t x = (0\<Colon>'a)"  and pnz: "poly t \<noteq> poly []"
+  let ?S = "{x. poly t x = 0}"
+  from H have "\<forall>x. x \<noteq>0 \<longrightarrow> poly t x = 0" by blast
+  hence th: "?S \<supseteq> UNIV - {0}" by auto
+  from poly_roots_finite pnz have th': "finite ?S" by blast
+  from finite_subset[OF th th'] UNIV_ring_char_0_infinte[where ?'a = 'a]
+  show "poly t x = (0\<Colon>'a)" by simp
+  qed
+
+lemma poly_zero: "(poly p = poly []) = list_all (%c. c = (0::'a::{idom,ring_char_0})) p"
+apply (induct "p", simp)
+apply (rule iffI)
+apply (drule poly_zero_lemma', auto)
+done
+
+
+
+text{*Basics of divisibility.*}
+
+lemma poly_primes: "([a, (1::'a::idom)] divides (p *** q)) = ([a, 1] divides p | [a, 1] divides q)"
+apply (auto simp add: divides_def fun_eq poly_mult poly_add poly_cmult left_distrib [symmetric])
+apply (drule_tac x = "-a" in spec)
+apply (auto simp add: poly_linear_divides poly_add poly_cmult left_distrib [symmetric])
+apply (rule_tac x = "qa *** q" in exI)
+apply (rule_tac [2] x = "p *** qa" in exI)
+apply (auto simp add: poly_add poly_mult poly_cmult mult_ac)
+done
+
+lemma poly_divides_refl: "p divides p"
+apply (simp add: divides_def)
+apply (rule_tac x = "[1]" in exI)
+apply (auto simp add: poly_mult fun_eq)
+done
+declare poly_divides_refl [simp]
+
+lemma poly_divides_trans: "[| p divides q; q divides r |] ==> p divides r"
+apply (simp add: divides_def, safe)
+apply (rule_tac x = "qa *** qaa" in exI)
+apply (auto simp add: poly_mult fun_eq mult_assoc)
+done
+
+lemma poly_divides_exp: "m \<le> n ==> (p %^ m) divides (p %^ n)"
+apply (auto simp add: le_iff_add)
+apply (induct_tac k)
+apply (rule_tac [2] poly_divides_trans)
+apply (auto simp add: divides_def)
+apply (rule_tac x = p in exI)
+apply (auto simp add: poly_mult fun_eq mult_ac)
+done
+
+lemma poly_exp_divides: "[| (p %^ n) divides q;  m\<le>n |] ==> (p %^ m) divides q"
+by (blast intro: poly_divides_exp poly_divides_trans)
+
+lemma poly_divides_add:
+   "[| p divides q; p divides r |] ==> p divides (q +++ r)"
+apply (simp add: divides_def, auto)
+apply (rule_tac x = "qa +++ qaa" in exI)
+apply (auto simp add: poly_add fun_eq poly_mult right_distrib)
+done
+
+lemma poly_divides_diff:
+   "[| p divides q; p divides (q +++ r) |] ==> p divides r"
+apply (simp add: divides_def, auto)
+apply (rule_tac x = "qaa +++ -- qa" in exI)
+apply (auto simp add: poly_add fun_eq poly_mult poly_minus right_diff_distrib algebra_simps)
+done
+
+lemma poly_divides_diff2: "[| p divides r; p divides (q +++ r) |] ==> p divides q"
+apply (erule poly_divides_diff)
+apply (auto simp add: poly_add fun_eq poly_mult divides_def add_ac)
+done
+
+lemma poly_divides_zero: "poly p = poly [] ==> q divides p"
+apply (simp add: divides_def)
+apply (rule exI[where x="[]"])
+apply (auto simp add: fun_eq poly_mult)
+done
+
+lemma poly_divides_zero2: "q divides []"
+apply (simp add: divides_def)
+apply (rule_tac x = "[]" in exI)
+apply (auto simp add: fun_eq)
+done
+declare poly_divides_zero2 [simp]
+
+text{*At last, we can consider the order of a root.*}
+
+
+lemma poly_order_exists_lemma [rule_format]:
+     "\<forall>p. length p = d --> poly p \<noteq> poly [] 
+             --> (\<exists>n q. p = mulexp n [-a, (1::'a::{idom,ring_char_0})] q & poly q a \<noteq> 0)"
+apply (induct "d")
+apply (simp add: fun_eq, safe)
+apply (case_tac "poly p a = 0")
+apply (drule_tac poly_linear_divides [THEN iffD1], safe)
+apply (drule_tac x = q in spec)
+apply (drule_tac poly_entire_neg [THEN iffD1], safe, force) 
+apply (rule_tac x = "Suc n" in exI)
+apply (rule_tac x = qa in exI)
+apply (simp del: pmult_Cons)
+apply (rule_tac x = 0 in exI, force) 
+done
+
+(* FIXME: Tidy up *)
+lemma poly_order_exists:
+     "[| length p = d; poly p \<noteq> poly [] |]
+      ==> \<exists>n. ([-a, 1] %^ n) divides p &
+                ~(([-a, (1::'a::{idom,ring_char_0})] %^ (Suc n)) divides p)"
+apply (drule poly_order_exists_lemma [where a=a], assumption, clarify)  
+apply (rule_tac x = n in exI, safe)
+apply (unfold divides_def)
+apply (rule_tac x = q in exI)
+apply (induct_tac "n", simp)
+apply (simp (no_asm_simp) add: poly_add poly_cmult poly_mult right_distrib mult_ac)
+apply safe
+apply (subgoal_tac "poly (mulexp n [- a, 1] q) \<noteq> poly ([- a, 1] %^ Suc n *** qa)") 
+apply simp 
+apply (induct_tac "n")
+apply (simp del: pmult_Cons pexp_Suc)
+apply (erule_tac Q = "poly q a = 0" in contrapos_np)
+apply (simp add: poly_add poly_cmult)
+apply (rule pexp_Suc [THEN ssubst])
+apply (rule ccontr)
+apply (simp add: poly_mult_left_cancel poly_mult_assoc del: pmult_Cons pexp_Suc)
+done
+
+lemma poly_one_divides: "[1] divides p"
+by (simp add: divides_def, auto)
+declare poly_one_divides [simp]
+
+lemma poly_order: "poly p \<noteq> poly []
+      ==> EX! n. ([-a, (1::'a::{idom,ring_char_0})] %^ n) divides p &
+                 ~(([-a, 1] %^ (Suc n)) divides p)"
+apply (auto intro: poly_order_exists simp add: less_linear simp del: pmult_Cons pexp_Suc)
+apply (cut_tac x = y and y = n in less_linear)
+apply (drule_tac m = n in poly_exp_divides)
+apply (auto dest: Suc_le_eq [THEN iffD2, THEN [2] poly_exp_divides]
+            simp del: pmult_Cons pexp_Suc)
+done
+
+text{*Order*}
+
+lemma some1_equalityD: "[| n = (@n. P n); EX! n. P n |] ==> P n"
+by (blast intro: someI2)
+
+lemma order:
+      "(([-a, (1::'a::{idom,ring_char_0})] %^ n) divides p &
+        ~(([-a, 1] %^ (Suc n)) divides p)) =
+        ((n = order a p) & ~(poly p = poly []))"
+apply (unfold order_def)
+apply (rule iffI)
+apply (blast dest: poly_divides_zero intro!: some1_equality [symmetric] poly_order)
+apply (blast intro!: poly_order [THEN [2] some1_equalityD])
+done
+
+lemma order2: "[| poly p \<noteq> poly [] |]
+      ==> ([-a, (1::'a::{idom,ring_char_0})] %^ (order a p)) divides p &
+              ~(([-a, 1] %^ (Suc(order a p))) divides p)"
+by (simp add: order del: pexp_Suc)
+
+lemma order_unique: "[| poly p \<noteq> poly []; ([-a, 1] %^ n) divides p;
+         ~(([-a, (1::'a::{idom,ring_char_0})] %^ (Suc n)) divides p)
+      |] ==> (n = order a p)"
+by (insert order [of a n p], auto) 
+
+lemma order_unique_lemma: "(poly p \<noteq> poly [] & ([-a, 1] %^ n) divides p &
+         ~(([-a, (1::'a::{idom,ring_char_0})] %^ (Suc n)) divides p))
+      ==> (n = order a p)"
+by (blast intro: order_unique)
+
+lemma order_poly: "poly p = poly q ==> order a p = order a q"
+by (auto simp add: fun_eq divides_def poly_mult order_def)
+
+lemma pexp_one: "p %^ (Suc 0) = p"
+apply (induct "p")
+apply (auto simp add: numeral_1_eq_1)
+done
+declare pexp_one [simp]
+
+lemma lemma_order_root [rule_format]:
+     "\<forall>p a. 0 < n & [- a, 1] %^ n divides p & ~ [- a, 1] %^ (Suc n) divides p
+             --> poly p a = 0"
+apply (induct "n", blast)
+apply (auto simp add: divides_def poly_mult simp del: pmult_Cons)
+done
+
+lemma order_root: "(poly p a = (0::'a::{idom,ring_char_0})) = ((poly p = poly []) | order a p \<noteq> 0)"
+apply (case_tac "poly p = poly []", auto)
+apply (simp add: poly_linear_divides del: pmult_Cons, safe)
+apply (drule_tac [!] a = a in order2)
+apply (rule ccontr)
+apply (simp add: divides_def poly_mult fun_eq del: pmult_Cons, blast)
+using neq0_conv
+apply (blast intro: lemma_order_root)
+done
+
+lemma order_divides: "(([-a, 1::'a::{idom,ring_char_0}] %^ n) divides p) = ((poly p = poly []) | n \<le> order a p)"
+apply (case_tac "poly p = poly []", auto)
+apply (simp add: divides_def fun_eq poly_mult)
+apply (rule_tac x = "[]" in exI)
+apply (auto dest!: order2 [where a=a]
+	    intro: poly_exp_divides simp del: pexp_Suc)
+done
+
+lemma order_decomp:
+     "poly p \<noteq> poly []
+      ==> \<exists>q. (poly p = poly (([-a, 1] %^ (order a p)) *** q)) &
+                ~([-a, 1::'a::{idom,ring_char_0}] divides q)"
+apply (unfold divides_def)
+apply (drule order2 [where a = a])
+apply (simp add: divides_def del: pexp_Suc pmult_Cons, safe)
+apply (rule_tac x = q in exI, safe)
+apply (drule_tac x = qa in spec)
+apply (auto simp add: poly_mult fun_eq poly_exp mult_ac simp del: pmult_Cons)
+done
+
+text{*Important composition properties of orders.*}
+
+lemma order_mult: "poly (p *** q) \<noteq> poly []
+      ==> order a (p *** q) = order a p + order (a::'a::{idom,ring_char_0}) q"
+apply (cut_tac a = a and p = "p***q" and n = "order a p + order a q" in order)
+apply (auto simp add: poly_entire simp del: pmult_Cons)
+apply (drule_tac a = a in order2)+
+apply safe
+apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons, safe)
+apply (rule_tac x = "qa *** qaa" in exI)
+apply (simp add: poly_mult mult_ac del: pmult_Cons)
+apply (drule_tac a = a in order_decomp)+
+apply safe
+apply (subgoal_tac "[-a,1] divides (qa *** qaa) ")
+apply (simp add: poly_primes del: pmult_Cons)
+apply (auto simp add: divides_def simp del: pmult_Cons)
+apply (rule_tac x = qb in exI)
+apply (subgoal_tac "poly ([-a, 1] %^ (order a p) *** (qa *** qaa)) = poly ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))")
+apply (drule poly_mult_left_cancel [THEN iffD1], force)
+apply (subgoal_tac "poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** (qa *** qaa))) = poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))) ")
+apply (drule poly_mult_left_cancel [THEN iffD1], force)
+apply (simp add: fun_eq poly_exp_add poly_mult mult_ac del: pmult_Cons)
+done
+
+
+
+lemma order_root2: "poly p \<noteq> poly [] ==> (poly p a = 0) = (order (a::'a::{idom,ring_char_0}) p \<noteq> 0)"
+by (rule order_root [THEN ssubst], auto)
+
+
+lemma pmult_one: "[1] *** p = p"
+by auto
+declare pmult_one [simp]
+
+lemma poly_Nil_zero: "poly [] = poly [0]"
+by (simp add: fun_eq)
+
+lemma rsquarefree_decomp:
+     "[| rsquarefree p; poly p a = (0::'a::{idom,ring_char_0}) |]
+      ==> \<exists>q. (poly p = poly ([-a, 1] *** q)) & poly q a \<noteq> 0"
+apply (simp add: rsquarefree_def, safe)
+apply (frule_tac a = a in order_decomp)
+apply (drule_tac x = a in spec)
+apply (drule_tac a = a in order_root2 [symmetric])
+apply (auto simp del: pmult_Cons)
+apply (rule_tac x = q in exI, safe)
+apply (simp add: poly_mult fun_eq)
+apply (drule_tac p1 = q in poly_linear_divides [THEN iffD1])
+apply (simp add: divides_def del: pmult_Cons, safe)
+apply (drule_tac x = "[]" in spec)
+apply (auto simp add: fun_eq)
+done
+
+
+text{*Normalization of a polynomial.*}
+
+lemma poly_normalize: "poly (pnormalize p) = poly p"
+apply (induct "p")
+apply (auto simp add: fun_eq)
+done
+declare poly_normalize [simp]
+
+
+text{*The degree of a polynomial.*}
+
+lemma lemma_degree_zero:
+     "list_all (%c. c = 0) p \<longleftrightarrow>  pnormalize p = []"
+by (induct "p", auto)
+
+lemma degree_zero: "(poly p = poly ([]:: (('a::{idom,ring_char_0}) list))) \<Longrightarrow> (degree p = 0)"
+apply (simp add: degree_def)
+apply (case_tac "pnormalize p = []")
+apply (auto simp add: poly_zero lemma_degree_zero )
+done
+
+lemma pnormalize_sing: "(pnormalize [x] = [x]) \<longleftrightarrow> x \<noteq> 0" by simp
+lemma pnormalize_pair: "y \<noteq> 0 \<longleftrightarrow> (pnormalize [x, y] = [x, y])" by simp
+lemma pnormal_cons: "pnormal p \<Longrightarrow> pnormal (c#p)" 
+  unfolding pnormal_def by simp
+lemma pnormal_tail: "p\<noteq>[] \<Longrightarrow> pnormal (c#p) \<Longrightarrow> pnormal p"
+  unfolding pnormal_def 
+  apply (cases "pnormalize p = []", auto)
+  by (cases "c = 0", auto)
+lemma pnormal_last_nonzero: "pnormal p ==> last p \<noteq> 0"
+  apply (induct p, auto simp add: pnormal_def)
+  apply (case_tac "pnormalize p = []", auto)
+  by (case_tac "a=0", auto)
+lemma  pnormal_length: "pnormal p \<Longrightarrow> 0 < length p"
+  unfolding pnormal_def length_greater_0_conv by blast
+lemma pnormal_last_length: "\<lbrakk>0 < length p ; last p \<noteq> 0\<rbrakk> \<Longrightarrow> pnormal p"
+  apply (induct p, auto)
+  apply (case_tac "p = []", auto)
+  apply (simp add: pnormal_def)
+  by (rule pnormal_cons, auto)
+lemma pnormal_id: "pnormal p \<longleftrightarrow> (0 < length p \<and> last p \<noteq> 0)"
+  using pnormal_last_length pnormal_length pnormal_last_nonzero by blast
+
+text{*Tidier versions of finiteness of roots.*}
+
+lemma poly_roots_finite_set: "poly p \<noteq> poly [] ==> finite {x::'a::{idom,ring_char_0}. poly p x = 0}"
+unfolding poly_roots_finite .
+
+text{*bound for polynomial.*}
+
+lemma poly_mono: "abs(x) \<le> k ==> abs(poly p (x::'a::{ordered_idom})) \<le> poly (map abs p) k"
+apply (induct "p", auto)
+apply (rule_tac y = "abs a + abs (x * poly p x)" in order_trans)
+apply (rule abs_triangle_ineq)
+apply (auto intro!: mult_mono simp add: abs_mult)
+done
+
+lemma poly_Sing: "poly [c] x = c" by simp
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy	Mon Oct 26 09:03:57 2009 +0100
@@ -0,0 +1,1743 @@
+(*  Title:      HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy
+    Author:     Amine Chaieb
+*)
+
+header {* Implementation and verification of mutivariate polynomials Library *}
+
+
+theory Reflected_Multivariate_Polynomial
+imports Parity Abstract_Rat Efficient_Nat List Polynomial_List
+begin
+
+  (* Impelementation *)
+
+subsection{* Datatype of polynomial expressions *} 
+
+datatype poly = C Num| Bound nat| Add poly poly|Sub poly poly
+  | Mul poly poly| Neg poly| Pw poly nat| CN poly nat poly
+
+ML{* @{term "Add"}*}
+syntax "_poly0" :: "poly" ("0\<^sub>p")
+translations "0\<^sub>p" \<rightleftharpoons> "C (0\<^sub>N)"
+syntax "_poly" :: "int \<Rightarrow> poly" ("_\<^sub>p")
+translations "i\<^sub>p" \<rightleftharpoons> "C (i\<^sub>N)"
+
+subsection{* Boundedness, substitution and all that *}
+consts polysize:: "poly \<Rightarrow> nat"
+primrec
+  "polysize (C c) = 1"
+  "polysize (Bound n) = 1"
+  "polysize (Neg p) = 1 + polysize p"
+  "polysize (Add p q) = 1 + polysize p + polysize q"
+  "polysize (Sub p q) = 1 + polysize p + polysize q"
+  "polysize (Mul p q) = 1 + polysize p + polysize q"
+  "polysize (Pw p n) = 1 + polysize p"
+  "polysize (CN c n p) = 4 + polysize c + polysize p"
+
+consts 
+  polybound0:: "poly \<Rightarrow> bool" (* a poly is INDEPENDENT of Bound 0 *)
+  polysubst0:: "poly \<Rightarrow> poly \<Rightarrow> poly" (* substitute a poly into a poly for Bound 0 *)
+primrec
+  "polybound0 (C c) = True"
+  "polybound0 (Bound n) = (n>0)"
+  "polybound0 (Neg a) = polybound0 a"
+  "polybound0 (Add a b) = (polybound0 a \<and> polybound0 b)"
+  "polybound0 (Sub a b) = (polybound0 a \<and> polybound0 b)" 
+  "polybound0 (Mul a b) = (polybound0 a \<and> polybound0 b)"
+  "polybound0 (Pw p n) = (polybound0 p)"
+  "polybound0 (CN c n p) = (n \<noteq> 0 \<and> polybound0 c \<and> polybound0 p)"
+primrec
+  "polysubst0 t (C c) = (C c)"
+  "polysubst0 t (Bound n) = (if n=0 then t else Bound n)"
+  "polysubst0 t (Neg a) = Neg (polysubst0 t a)"
+  "polysubst0 t (Add a b) = Add (polysubst0 t a) (polysubst0 t b)"
+  "polysubst0 t (Sub a b) = Sub (polysubst0 t a) (polysubst0 t b)" 
+  "polysubst0 t (Mul a b) = Mul (polysubst0 t a) (polysubst0 t b)"
+  "polysubst0 t (Pw p n) = Pw (polysubst0 t p) n"
+  "polysubst0 t (CN c n p) = (if n=0 then Add (polysubst0 t c) (Mul t (polysubst0 t p))
+                             else CN (polysubst0 t c) n (polysubst0 t p))"
+
+consts 
+  decrpoly:: "poly \<Rightarrow> poly" 
+recdef decrpoly "measure polysize"
+  "decrpoly (Bound n) = Bound (n - 1)"
+  "decrpoly (Neg a) = Neg (decrpoly a)"
+  "decrpoly (Add a b) = Add (decrpoly a) (decrpoly b)"
+  "decrpoly (Sub a b) = Sub (decrpoly a) (decrpoly b)"
+  "decrpoly (Mul a b) = Mul (decrpoly a) (decrpoly b)"
+  "decrpoly (Pw p n) = Pw (decrpoly p) n"
+  "decrpoly (CN c n p) = CN (decrpoly c) (n - 1) (decrpoly p)"
+  "decrpoly a = a"
+
+subsection{* Degrees and heads and coefficients *}
+
+consts degree:: "poly \<Rightarrow> nat"
+recdef degree "measure size"
+  "degree (CN c 0 p) = 1 + degree p"
+  "degree p = 0"
+consts head:: "poly \<Rightarrow> poly"
+
+recdef head "measure size"
+  "head (CN c 0 p) = head p"
+  "head p = p"
+  (* More general notions of degree and head *)
+consts degreen:: "poly \<Rightarrow> nat \<Rightarrow> nat"
+recdef degreen "measure size"
+  "degreen (CN c n p) = (\<lambda>m. if n=m then 1 + degreen p n else 0)"
+  "degreen p = (\<lambda>m. 0)"
+
+consts headn:: "poly \<Rightarrow> nat \<Rightarrow> poly"
+recdef headn "measure size"
+  "headn (CN c n p) = (\<lambda>m. if n \<le> m then headn p m else CN c n p)"
+  "headn p = (\<lambda>m. p)"
+
+consts coefficients:: "poly \<Rightarrow> poly list"
+recdef coefficients "measure size"
+  "coefficients (CN c 0 p) = c#(coefficients p)"
+  "coefficients p = [p]"
+
+consts isconstant:: "poly \<Rightarrow> bool"
+recdef isconstant "measure size"
+  "isconstant (CN c 0 p) = False"
+  "isconstant p = True"
+
+consts behead:: "poly \<Rightarrow> poly"
+recdef behead "measure size"
+  "behead (CN c 0 p) = (let p' = behead p in if p' = 0\<^sub>p then c else CN c 0 p')"
+  "behead p = 0\<^sub>p"
+
+consts headconst:: "poly \<Rightarrow> Num"
+recdef headconst "measure size"
+  "headconst (CN c n p) = headconst p"
+  "headconst (C n) = n"
+
+subsection{* Operations for normalization *}
+consts 
+  polyadd :: "poly\<times>poly \<Rightarrow> poly"
+  polyneg :: "poly \<Rightarrow> poly" ("~\<^sub>p")
+  polysub :: "poly\<times>poly \<Rightarrow> poly"
+  polymul :: "poly\<times>poly \<Rightarrow> poly"
+  polypow :: "nat \<Rightarrow> poly \<Rightarrow> poly"
+syntax "_polyadd" :: "poly \<Rightarrow> poly \<Rightarrow> poly" (infixl "+\<^sub>p" 60)
+translations "a +\<^sub>p b" \<rightleftharpoons> "polyadd (a,b)"  
+syntax "_polymul" :: "poly \<Rightarrow> poly \<Rightarrow> poly" (infixl "*\<^sub>p" 60)
+translations "a *\<^sub>p b" \<rightleftharpoons> "polymul (a,b)"  
+syntax "_polysub" :: "poly \<Rightarrow> poly \<Rightarrow> poly" (infixl "-\<^sub>p" 60)
+translations "a -\<^sub>p b" \<rightleftharpoons> "polysub (a,b)"  
+syntax "_polypow" :: "nat \<Rightarrow> poly \<Rightarrow> poly" (infixl "^\<^sub>p" 60)
+translations "a ^\<^sub>p k" \<rightleftharpoons> "polypow k a" 
+
+recdef polyadd "measure (\<lambda> (a,b). polysize a + polysize b)"
+  "polyadd (C c, C c') = C (c+\<^sub>Nc')"
+  "polyadd (C c, CN c' n' p') = CN (polyadd (C c, c')) n' p'"
+  "polyadd (CN c n p, C c') = CN (polyadd (c, C c')) n p"
+stupid:  "polyadd (CN c n p, CN c' n' p') = 
+    (if n < n' then CN (polyadd(c,CN c' n' p')) n p
+     else if n'<n then CN (polyadd(CN c n p, c')) n' p'
+     else (let cc' = polyadd (c,c') ; 
+               pp' = polyadd (p,p')
+           in (if pp' = 0\<^sub>p then cc' else CN cc' n pp')))"
+  "polyadd (a, b) = Add a b"
+(hints recdef_simp add: Let_def measure_def split_def inv_image_def)
+
+(*
+declare stupid [simp del, code del]
+
+lemma [simp,code]: "polyadd (CN c n p, CN c' n' p') = 
+    (if n < n' then CN (polyadd(c,CN c' n' p')) n p
+     else if n'<n then CN (polyadd(CN c n p, c')) n' p'
+     else (let cc' = polyadd (c,c') ; 
+               pp' = polyadd (p,p')
+           in (if pp' = 0\<^sub>p then cc' else CN cc' n pp')))"
+  by (simp add: Let_def stupid)
+*)
+
+recdef polyneg "measure size"
+  "polyneg (C c) = C (~\<^sub>N c)"
+  "polyneg (CN c n p) = CN (polyneg c) n (polyneg p)"
+  "polyneg a = Neg a"
+
+defs polysub_def[code]: "polysub \<equiv> \<lambda> (p,q). polyadd (p,polyneg q)"
+
+recdef polymul "measure (\<lambda>(a,b). size a + size b)"
+  "polymul(C c, C c') = C (c*\<^sub>Nc')"
+  "polymul(C c, CN c' n' p') = 
+      (if c = 0\<^sub>N then 0\<^sub>p else CN (polymul(C c,c')) n' (polymul(C c, p')))"
+  "polymul(CN c n p, C c') = 
+      (if c' = 0\<^sub>N  then 0\<^sub>p else CN (polymul(c,C c')) n (polymul(p, C c')))"
+  "polymul(CN c n p, CN c' n' p') = 
+  (if n<n' then CN (polymul(c,CN c' n' p')) n (polymul(p,CN c' n' p'))
+  else if n' < n 
+  then CN (polymul(CN c n p,c')) n' (polymul(CN c n p,p'))
+  else polyadd(polymul(CN c n p, c'),CN 0\<^sub>p n' (polymul(CN c n p, p'))))"
+  "polymul (a,b) = Mul a b"
+recdef polypow "measure id"
+  "polypow 0 = (\<lambda>p. 1\<^sub>p)"
+  "polypow n = (\<lambda>p. let q = polypow (n div 2) p ; d = polymul(q,q) in 
+                    if even n then d else polymul(p,d))"
+
+consts polynate :: "poly \<Rightarrow> poly"
+recdef polynate "measure polysize"
+  "polynate (Bound n) = CN 0\<^sub>p n 1\<^sub>p"
+  "polynate (Add p q) = (polynate p +\<^sub>p polynate q)"
+  "polynate (Sub p q) = (polynate p -\<^sub>p polynate q)"
+  "polynate (Mul p q) = (polynate p *\<^sub>p polynate q)"
+  "polynate (Neg p) = (~\<^sub>p (polynate p))"
+  "polynate (Pw p n) = ((polynate p) ^\<^sub>p n)"
+  "polynate (CN c n p) = polynate (Add c (Mul (Bound n) p))"
+  "polynate (C c) = C (normNum c)"
+
+fun poly_cmul :: "Num \<Rightarrow> poly \<Rightarrow> poly" where
+  "poly_cmul y (C x) = C (y *\<^sub>N x)"
+| "poly_cmul y (CN c n p) = CN (poly_cmul y c) n (poly_cmul y p)"
+| "poly_cmul y p = C y *\<^sub>p p"
+
+constdefs monic:: "poly \<Rightarrow> (poly \<times> bool)"
+  "monic p \<equiv> (let h = headconst p in if h = 0\<^sub>N then (p,False) else ((C (Ninv h)) *\<^sub>p p, 0>\<^sub>N h))"
+
+subsection{* Pseudo-division *}
+
+constdefs shift1:: "poly \<Rightarrow> poly"
+  "shift1 p \<equiv> CN 0\<^sub>p 0 p"
+consts funpow :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a"
+
+primrec
+  "funpow 0 f x = x"
+  "funpow (Suc n) f x = funpow n f (f x)"
+function (tailrec) polydivide_aux :: "(poly \<times> nat \<times> poly \<times> nat \<times> poly) \<Rightarrow> (nat \<times> poly)"
+  where
+  "polydivide_aux (a,n,p,k,s) = 
+  (if s = 0\<^sub>p then (k,s)
+  else (let b = head s; m = degree s in
+  (if m < n then (k,s) else 
+  (let p'= funpow (m - n) shift1 p in 
+  (if a = b then polydivide_aux (a,n,p,k,s -\<^sub>p p') 
+  else polydivide_aux (a,n,p,Suc k, (a *\<^sub>p s) -\<^sub>p (b *\<^sub>p p')))))))"
+  by pat_completeness auto
+
+
+constdefs polydivide:: "poly \<Rightarrow> poly \<Rightarrow> (nat \<times> poly)"
+  "polydivide s p \<equiv> polydivide_aux (head p,degree p,p,0, s)"
+
+fun poly_deriv_aux :: "nat \<Rightarrow> poly \<Rightarrow> poly" where
+  "poly_deriv_aux n (CN c 0 p) = CN (poly_cmul ((int n)\<^sub>N) c) 0 (poly_deriv_aux (n + 1) p)"
+| "poly_deriv_aux n p = poly_cmul ((int n)\<^sub>N) p"
+
+fun poly_deriv :: "poly \<Rightarrow> poly" where
+  "poly_deriv (CN c 0 p) = poly_deriv_aux 1 p"
+| "poly_deriv p = 0\<^sub>p"
+
+  (* Verification *)
+lemma nth_pos2[simp]: "0 < n \<Longrightarrow> (x#xs) ! n = xs ! (n - 1)"
+using Nat.gr0_conv_Suc
+by clarsimp
+
+subsection{* Semantics of the polynomial representation *}
+
+consts Ipoly :: "'a list \<Rightarrow> poly \<Rightarrow> 'a::{ring_char_0,power,division_by_zero,field}"
+primrec
+  "Ipoly bs (C c) = INum c"
+  "Ipoly bs (Bound n) = bs!n"
+  "Ipoly bs (Neg a) = - Ipoly bs a"
+  "Ipoly bs (Add a b) = Ipoly bs a + Ipoly bs b"
+  "Ipoly bs (Sub a b) = Ipoly bs a - Ipoly bs b"
+  "Ipoly bs (Mul a b) = Ipoly bs a * Ipoly bs b"
+  "Ipoly bs (Pw t n) = (Ipoly bs t) ^ n"
+  "Ipoly bs (CN c n p) = (Ipoly bs c) + (bs!n)*(Ipoly bs p)"
+syntax "_Ipoly" :: "poly \<Rightarrow> 'a list \<Rightarrow>'a::{ring_char_0,power,division_by_zero,field}" ("\<lparr>_\<rparr>\<^sub>p\<^bsup>_\<^esup>")
+translations "\<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup>" \<rightleftharpoons> "Ipoly bs p"  
+
+lemma Ipoly_CInt: "Ipoly bs (C (i,1)) = of_int i" 
+  by (simp add: INum_def)
+lemma Ipoly_CRat: "Ipoly bs (C (i, j)) = of_int i / of_int j" 
+  by (simp  add: INum_def)
+
+lemmas RIpoly_eqs = Ipoly.simps(2-7) Ipoly_CInt Ipoly_CRat
+
+subsection {* Normal form and normalization *}
+
+consts isnpolyh:: "poly \<Rightarrow> nat \<Rightarrow> bool"
+recdef isnpolyh "measure size"
+  "isnpolyh (C c) = (\<lambda>k. isnormNum c)"
+  "isnpolyh (CN c n p) = (\<lambda>k. n\<ge> k \<and> (isnpolyh c (Suc n)) \<and> (isnpolyh p n) \<and> (p \<noteq> 0\<^sub>p))"
+  "isnpolyh p = (\<lambda>k. False)"
+
+lemma isnpolyh_mono: "\<lbrakk>n' \<le> n ; isnpolyh p n\<rbrakk> \<Longrightarrow> isnpolyh p n'"
+by (induct p rule: isnpolyh.induct, auto)
+
+constdefs isnpoly:: "poly \<Rightarrow> bool"
+  "isnpoly p \<equiv> isnpolyh p 0"
+
+text{* polyadd preserves normal forms *}
+
+lemma polyadd_normh: "\<lbrakk>isnpolyh p n0 ; isnpolyh q n1\<rbrakk> 
+      \<Longrightarrow> isnpolyh (polyadd(p,q)) (min n0 n1)"
+proof(induct p q arbitrary: n0 n1 rule: polyadd.induct)
+  case (2 a b c' n' p' n0 n1)
+  from prems have  th1: "isnpolyh (C (a,b)) (Suc n')" by simp 
+  from prems(3) have th2: "isnpolyh c' (Suc n')"  and nplen1: "n' \<ge> n1" by simp_all
+  with isnpolyh_mono have cp: "isnpolyh c' (Suc n')" by simp
+  with prems(1)[OF th1 th2] have th3:"isnpolyh (C (a,b) +\<^sub>p c') (Suc n')" by simp
+  from nplen1 have n01len1: "min n0 n1 \<le> n'" by simp 
+  thus ?case using prems th3 by simp
+next
+  case (3 c' n' p' a b n1 n0)
+  from prems have  th1: "isnpolyh (C (a,b)) (Suc n')" by simp 
+  from prems(2) have th2: "isnpolyh c' (Suc n')"  and nplen1: "n' \<ge> n1" by simp_all
+  with isnpolyh_mono have cp: "isnpolyh c' (Suc n')" by simp
+  with prems(1)[OF th2 th1] have th3:"isnpolyh (c' +\<^sub>p C (a,b)) (Suc n')" by simp
+  from nplen1 have n01len1: "min n0 n1 \<le> n'" by simp 
+  thus ?case using prems th3 by simp
+next
+  case (4 c n p c' n' p' n0 n1)
+  hence nc: "isnpolyh c (Suc n)" and np: "isnpolyh p n" by simp_all
+  from prems have nc': "isnpolyh c' (Suc n')" and np': "isnpolyh p' n'" by simp_all 
+  from prems have ngen0: "n \<ge> n0" by simp
+  from prems have n'gen1: "n' \<ge> n1" by simp 
+  have "n < n' \<or> n' < n \<or> n = n'" by auto
+  moreover {assume eq: "n = n'" hence eq': "\<not> n' < n \<and> \<not> n < n'" by simp
+    with prems(2)[rule_format, OF eq' nc nc'] 
+    have ncc':"isnpolyh (c +\<^sub>p c') (Suc n)" by auto
+    hence ncc'n01: "isnpolyh (c +\<^sub>p c') (min n0 n1)"
+      using isnpolyh_mono[where n'="min n0 n1" and n="Suc n"] ngen0 n'gen1 by auto
+    from eq prems(1)[rule_format, OF eq' np np'] have npp': "isnpolyh (p +\<^sub>p p') n" by simp
+    have minle: "min n0 n1 \<le> n'" using ngen0 n'gen1 eq by simp
+    from minle npp' ncc'n01 prems ngen0 n'gen1 ncc' have ?case by (simp add: Let_def)}
+  moreover {assume lt: "n < n'"
+    have "min n0 n1 \<le> n0" by simp
+    with prems have th1:"min n0 n1 \<le> n" by auto 
+    from prems have th21: "isnpolyh c (Suc n)" by simp
+    from prems have th22: "isnpolyh (CN c' n' p') n'" by simp
+    from lt have th23: "min (Suc n) n' = Suc n" by arith
+    from prems(4)[rule_format, OF lt th21 th22]
+    have "isnpolyh (polyadd (c, CN c' n' p')) (Suc n)" using th23 by simp
+    with prems th1 have ?case by simp } 
+  moreover {assume gt: "n' < n" hence gt': "n' < n \<and> \<not> n < n'" by simp
+    have "min n0 n1 \<le> n1"  by simp
+    with prems have th1:"min n0 n1 \<le> n'" by auto
+    from prems have th21: "isnpolyh c' (Suc n')" by simp_all
+    from prems have th22: "isnpolyh (CN c n p) n" by simp
+    from gt have th23: "min n (Suc n') = Suc n'" by arith
+    from prems(3)[rule_format, OF  gt' th22 th21]
+    have "isnpolyh (polyadd (CN c n p,c')) (Suc n')" using th23 by simp
+    with prems th1 have ?case by simp}
+      ultimately show ?case by blast
+qed auto
+
+lemma polyadd[simp]: "Ipoly bs (polyadd (p,q)) = (Ipoly bs p) + (Ipoly bs q)"
+by (induct p q rule: polyadd.induct, auto simp add: Let_def ring_simps right_distrib[symmetric] simp del: right_distrib)
+
+lemma polyadd_norm: "\<lbrakk> isnpoly p ; isnpoly q\<rbrakk> \<Longrightarrow> isnpoly (polyadd(p,q))"
+  using polyadd_normh[of "p" "0" "q" "0"] isnpoly_def by simp
+
+text{* The degree of addition and other general lemmas needed for the normal form of polymul*}
+
+lemma polyadd_different_degreen: 
+  "\<lbrakk>isnpolyh p n0 ; isnpolyh q n1; degreen p m \<noteq> degreen q m ; m \<le> min n0 n1\<rbrakk> \<Longrightarrow> 
+  degreen (polyadd(p,q)) m = max (degreen p m) (degreen q m)"
+proof (induct p q arbitrary: m n0 n1 rule: polyadd.induct)
+  case (4 c n p c' n' p' m n0 n1)
+  thus ?case 
+    apply (cases "n' < n", simp_all add: Let_def)
+    apply (cases "n = n'", simp_all)
+    apply (cases "n' = m", simp_all add: Let_def)
+    by (erule allE[where x="m"], erule allE[where x="Suc m"], 
+           erule allE[where x="m"], erule allE[where x="Suc m"], 
+           clarsimp,erule allE[where x="m"],erule allE[where x="Suc m"], simp)
+qed simp_all 
+
+lemma headnz[simp]: "\<lbrakk>isnpolyh p n ; p \<noteq> 0\<^sub>p\<rbrakk> \<Longrightarrow> headn p m \<noteq> 0\<^sub>p"
+  by (induct p arbitrary: n rule: headn.induct, auto)
+lemma degree_isnpolyh_Suc[simp]: "isnpolyh p (Suc n) \<Longrightarrow> degree p = 0"
+  by (induct p arbitrary: n rule: degree.induct, auto)
+lemma degreen_0[simp]: "isnpolyh p n \<Longrightarrow> m < n \<Longrightarrow> degreen p m = 0"
+  by (induct p arbitrary: n rule: degreen.induct, auto)
+
+lemma degree_isnpolyh_Suc': "n > 0 \<Longrightarrow> isnpolyh p n \<Longrightarrow> degree p = 0"
+  by (induct p arbitrary: n rule: degree.induct, auto)
+
+lemma degree_npolyhCN[simp]: "isnpolyh (CN c n p) n0 \<Longrightarrow> degree c = 0"
+  using degree_isnpolyh_Suc by auto
+lemma degreen_npolyhCN[simp]: "isnpolyh (CN c n p) n0 \<Longrightarrow> degreen c n = 0"
+  using degreen_0 by auto
+
+
+lemma degreen_polyadd:
+  assumes np: "isnpolyh p n0" and nq: "isnpolyh q n1" and m: "m \<le> max n0 n1"
+  shows "degreen (p +\<^sub>p q) m \<le> max (degreen p m) (degreen q m)"
+  using np nq m
+proof (induct p q arbitrary: n0 n1 m rule: polyadd.induct)
+  case (2 c c' n' p' n0 n1) thus ?case  by (cases n', simp_all)
+next
+  case (3 c n p c' n0 n1) thus ?case by (cases n, auto)
+next
+  case (4 c n p c' n' p' n0 n1 m) 
+  thus ?case 
+    apply (cases "n < n'", simp_all add: Let_def)
+    apply (cases "n' < n", simp_all)
+    apply (erule allE[where x="n"],erule allE[where x="Suc n"],clarify)
+    apply (erule allE[where x="n'"],erule allE[where x="Suc n'"],clarify)
+    by (erule allE[where x="m"],erule allE[where x="m"], auto)
+qed auto
+
+
+lemma polyadd_eq_const_degreen: "\<lbrakk> isnpolyh p n0 ; isnpolyh q n1 ; polyadd (p,q) = C c\<rbrakk> 
+  \<Longrightarrow> degreen p m = degreen q m"
+proof (induct p q arbitrary: m n0 n1 c rule: polyadd.induct)
+  case (4 c n p c' n' p' m n0 n1 x) 
+  hence z: "CN c n p +\<^sub>p CN c' n' p' = C x" by simp
+  {assume nn': "n' < n" hence ?case using prems by simp}
+  moreover 
+  {assume nn':"\<not> n' < n" hence "n < n' \<or> n = n'" by arith
+    moreover {assume "n < n'" with prems have ?case by simp }
+    moreover {assume eq: "n = n'" hence ?case using prems 
+	by (cases "p +\<^sub>p p' = 0\<^sub>p", auto simp add: Let_def) }
+    ultimately have ?case by blast}
+  ultimately show ?case by blast
+qed simp_all
+
+lemma polymul_properties:
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  and np: "isnpolyh p n0" and nq: "isnpolyh q n1" and m: "m \<le> min n0 n1"
+  shows "isnpolyh (p *\<^sub>p q) (min n0 n1)" 
+  and "(p *\<^sub>p q = 0\<^sub>p) = (p = 0\<^sub>p \<or> q = 0\<^sub>p)" 
+  and "degreen (p *\<^sub>p q) m = (if (p = 0\<^sub>p \<or> q = 0\<^sub>p) then 0 
+                             else degreen p m + degreen q m)"
+  using np nq m
+proof(induct p q arbitrary: n0 n1 m rule: polymul.induct)
+  case (2 a b c' n' p') 
+  let ?c = "(a,b)"
+  { case (1 n0 n1) 
+    hence n: "isnpolyh (C ?c) n'" "isnpolyh c' (Suc n')" "isnpolyh p' n'" "isnormNum ?c" 
+      "isnpolyh (CN c' n' p') n1"
+      by simp_all
+    {assume "?c = 0\<^sub>N" hence ?case by auto}
+      moreover {assume cnz: "?c \<noteq> 0\<^sub>N" 
+	from "2.hyps"(1)[rule_format,where xb="n'",  OF cnz n(1) n(3)] 
+	  "2.hyps"(2)[rule_format, where x="Suc n'" 
+	  and xa="Suc n'" and xb = "n'", OF cnz ] cnz n have ?case
+	  by (auto simp add: min_def)}
+      ultimately show ?case by blast
+  next
+    case (2 n0 n1) thus ?case by auto 
+  next
+    case (3 n0 n1) thus ?case  using "2.hyps" by auto } 
+next
+  case (3 c n p a b){
+    let ?c' = "(a,b)"
+    case (1 n0 n1) 
+    hence n: "isnpolyh (C ?c') n" "isnpolyh c (Suc n)" "isnpolyh p n" "isnormNum ?c'" 
+      "isnpolyh (CN c n p) n0"
+      by simp_all
+    {assume "?c' = 0\<^sub>N" hence ?case by auto}
+      moreover {assume cnz: "?c' \<noteq> 0\<^sub>N"
+	from "3.hyps"(1)[rule_format,where xb="n",  OF cnz n(3) n(1)] 
+	  "3.hyps"(2)[rule_format, where x="Suc n" 
+	  and xa="Suc n" and xb = "n", OF cnz ] cnz n have ?case
+	  by (auto simp add: min_def)}
+      ultimately show ?case by blast
+  next
+    case (2 n0 n1) thus ?case apply auto done
+  next
+    case (3 n0 n1) thus ?case  using "3.hyps" by auto } 
+next
+  case (4 c n p c' n' p')
+  let ?cnp = "CN c n p" let ?cnp' = "CN c' n' p'"
+    {fix n0 n1
+      assume "isnpolyh ?cnp n0" and "isnpolyh ?cnp' n1"
+      hence cnp: "isnpolyh ?cnp n" and cnp': "isnpolyh ?cnp' n'"
+	and np: "isnpolyh p n" and nc: "isnpolyh c (Suc n)" 
+	and np': "isnpolyh p' n'" and nc': "isnpolyh c' (Suc n')"
+	and nn0: "n \<ge> n0" and nn1:"n' \<ge> n1"
+	by simp_all
+      have "n < n' \<or> n' < n \<or> n' = n" by auto
+      moreover
+      {assume nn': "n < n'"
+	with "4.hyps"(5)[rule_format, OF nn' np cnp', where xb ="n"] 
+	  "4.hyps"(6)[rule_format, OF nn' nc cnp', where xb="n"] nn' nn0 nn1 cnp
+	have "isnpolyh (?cnp *\<^sub>p ?cnp') (min n0 n1)"
+	  by (simp add: min_def) }
+      moreover
+
+      {assume nn': "n > n'" hence stupid: "n' < n \<and> \<not> n < n'" by arith
+	with "4.hyps"(3)[rule_format, OF stupid cnp np', where xb="n'"]
+	  "4.hyps"(4)[rule_format, OF stupid cnp nc', where xb="Suc n'"] 
+	  nn' nn0 nn1 cnp'
+	have "isnpolyh (?cnp *\<^sub>p ?cnp') (min n0 n1)"
+	  by (cases "Suc n' = n", simp_all add: min_def)}
+      moreover
+      {assume nn': "n' = n" hence stupid: "\<not> n' < n \<and> \<not> n < n'" by arith
+	from "4.hyps"(1)[rule_format, OF stupid cnp np', where xb="n"]
+	  "4.hyps"(2)[rule_format, OF stupid cnp nc', where xb="n"] nn' cnp cnp' nn1
+	
+	have "isnpolyh (?cnp *\<^sub>p ?cnp') (min n0 n1)"
+	  by simp (rule polyadd_normh,simp_all add: min_def isnpolyh_mono[OF nn0]) }
+      ultimately show "isnpolyh (?cnp *\<^sub>p ?cnp') (min n0 n1)" by blast }
+    note th = this
+    {fix n0 n1 m
+      assume np: "isnpolyh ?cnp n0" and np':"isnpolyh ?cnp' n1"
+      and m: "m \<le> min n0 n1"
+      let ?d = "degreen (?cnp *\<^sub>p ?cnp') m"
+      let ?d1 = "degreen ?cnp m"
+      let ?d2 = "degreen ?cnp' m"
+      let ?eq = "?d = (if ?cnp = 0\<^sub>p \<or> ?cnp' = 0\<^sub>p then 0  else ?d1 + ?d2)"
+      have "n'<n \<or> n < n' \<or> n' = n" by auto
+      moreover 
+      {assume "n' < n \<or> n < n'"
+	with "4.hyps" np np' m 
+	have ?eq apply (cases "n' < n", simp_all)
+	apply (erule allE[where x="n"],erule allE[where x="n"],auto) 
+	done }
+      moreover
+      {assume nn': "n' = n"  hence nn:"\<not> n' < n \<and> \<not> n < n'" by arith
+ 	from "4.hyps"(1)[rule_format, OF nn, where x="n" and xa ="n'" and xb="n"]
+	  "4.hyps"(2)[rule_format, OF nn, where x="n" and xa ="Suc n'" and xb="n"] 
+	  np np' nn'
+	have norm: "isnpolyh ?cnp n" "isnpolyh c' (Suc n)" "isnpolyh (?cnp *\<^sub>p c') n"
+	  "isnpolyh p' n" "isnpolyh (?cnp *\<^sub>p p') n" "isnpolyh (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n"
+	  "(?cnp *\<^sub>p c' = 0\<^sub>p) = (c' = 0\<^sub>p)" 
+	  "?cnp *\<^sub>p p' \<noteq> 0\<^sub>p" by (auto simp add: min_def)
+	{assume mn: "m = n" 
+	  from "4.hyps"(1)[rule_format, OF nn norm(1,4), where xb="n"]
+	    "4.hyps"(2)[rule_format, OF nn norm(1,2), where xb="n"] norm nn' mn
+	  have degs:  "degreen (?cnp *\<^sub>p c') n = 
+	    (if c'=0\<^sub>p then 0 else ?d1 + degreen c' n)"
+	    "degreen (?cnp *\<^sub>p p') n = ?d1  + degreen p' n" by (simp_all add: min_def)
+	  from degs norm
+	  have th1: "degreen(?cnp *\<^sub>p c') n < degreen (CN 0\<^sub>p n (?cnp *\<^sub>p p')) n" by simp
+	  hence neq: "degreen (?cnp *\<^sub>p c') n \<noteq> degreen (CN 0\<^sub>p n (?cnp *\<^sub>p p')) n"
+	    by simp
+	  have nmin: "n \<le> min n n" by (simp add: min_def)
+	  from polyadd_different_degreen[OF norm(3,6) neq nmin] th1
+	  have deg: "degreen (CN c n p *\<^sub>p c' +\<^sub>p CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n = degreen (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n" by simp 
+	  from "4.hyps"(1)[rule_format, OF nn norm(1,4), where xb="n"]
+	    "4.hyps"(2)[rule_format, OF nn norm(1,2), where xb="n"]
+	    mn norm m nn' deg
+	  have ?eq by simp}
+	moreover
+	{assume mn: "m \<noteq> n" hence mn': "m < n" using m np by auto
+	  from nn' m np have max1: "m \<le> max n n"  by simp 
+	  hence min1: "m \<le> min n n" by simp	
+	  hence min2: "m \<le> min n (Suc n)" by simp
+	  {assume "c' = 0\<^sub>p"
+	    from `c' = 0\<^sub>p` have ?eq
+	      using "4.hyps"(1)[rule_format, OF nn norm(1,4) min1]
+	    "4.hyps"(2)[rule_format, OF nn norm(1,2) min2] mn nn'
+	      apply simp
+	      done}
+	  moreover
+	  {assume cnz: "c' \<noteq> 0\<^sub>p"
+	    from "4.hyps"(1)[rule_format, OF nn norm(1,4) min1]
+	      "4.hyps"(2)[rule_format, OF nn norm(1,2) min2]
+	      degreen_polyadd[OF norm(3,6) max1]
+
+	    have "degreen (?cnp *\<^sub>p c' +\<^sub>p CN 0\<^sub>p n (?cnp *\<^sub>p p')) m 
+	      \<le> max (degreen (?cnp *\<^sub>p c') m) (degreen (CN 0\<^sub>p n (?cnp *\<^sub>p p')) m)"
+	      using mn nn' cnz np np' by simp
+	    with "4.hyps"(1)[rule_format, OF nn norm(1,4) min1]
+	      "4.hyps"(2)[rule_format, OF nn norm(1,2) min2]
+	      degreen_0[OF norm(3) mn'] have ?eq using nn' mn cnz np np' by clarsimp}
+	  ultimately have ?eq by blast }
+	ultimately have ?eq by blast}
+      ultimately show ?eq by blast}
+    note degth = this
+    { case (2 n0 n1)
+      hence np: "isnpolyh ?cnp n0" and np': "isnpolyh ?cnp' n1" 
+	and m: "m \<le> min n0 n1" by simp_all
+      hence mn: "m \<le> n" by simp
+      let ?c0p = "CN 0\<^sub>p n (?cnp *\<^sub>p p')"
+      {assume C: "?cnp *\<^sub>p c' +\<^sub>p ?c0p = 0\<^sub>p" "n' = n"
+	hence nn: "\<not>n' < n \<and> \<not> n<n'" by simp
+	from "4.hyps"(1) [rule_format, OF nn, where x="n" and xa = "n" and xb="n"] 
+	  "4.hyps"(2) [rule_format, OF nn, where x="n" and xa = "Suc n" and xb="n"] 
+	  np np' C(2) mn
+	have norm: "isnpolyh ?cnp n" "isnpolyh c' (Suc n)" "isnpolyh (?cnp *\<^sub>p c') n"
+	  "isnpolyh p' n" "isnpolyh (?cnp *\<^sub>p p') n" "isnpolyh (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n"
+	  "(?cnp *\<^sub>p c' = 0\<^sub>p) = (c' = 0\<^sub>p)" 
+	  "?cnp *\<^sub>p p' \<noteq> 0\<^sub>p" 
+	  "degreen (?cnp *\<^sub>p c') n = (if c'=0\<^sub>p then 0 else degreen ?cnp n + degreen c' n)"
+	    "degreen (?cnp *\<^sub>p p') n = degreen ?cnp n + degreen p' n"
+	  by (simp_all add: min_def)
+	    
+	  from norm have cn: "isnpolyh (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n" by simp
+	  have degneq: "degreen (?cnp *\<^sub>p c') n < degreen (CN 0\<^sub>p n (?cnp *\<^sub>p p')) n" 
+	    using norm by simp
+	from polyadd_eq_const_degreen[OF norm(3) cn C(1), where m="n"]  degneq
+	have "False" by simp }
+      thus ?case using "4.hyps" by clarsimp}
+qed auto
+
+lemma polymul[simp]: "Ipoly bs (p *\<^sub>p q) = (Ipoly bs p) * (Ipoly bs q)"
+by(induct p q rule: polymul.induct, auto simp add: ring_simps)
+
+lemma polymul_normh: 
+    assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "\<lbrakk>isnpolyh p n0 ; isnpolyh q n1\<rbrakk> \<Longrightarrow> isnpolyh (p *\<^sub>p q) (min n0 n1)"
+  using polymul_properties(1)  by blast
+lemma polymul_eq0_iff: 
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "\<lbrakk> isnpolyh p n0 ; isnpolyh q n1\<rbrakk> \<Longrightarrow> (p *\<^sub>p q = 0\<^sub>p) = (p = 0\<^sub>p \<or> q = 0\<^sub>p) "
+  using polymul_properties(2)  by blast
+lemma polymul_degreen:  
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "\<lbrakk> isnpolyh p n0 ; isnpolyh q n1 ; m \<le> min n0 n1\<rbrakk> \<Longrightarrow> degreen (p *\<^sub>p q) m = (if (p = 0\<^sub>p \<or> q = 0\<^sub>p) then 0 else degreen p m + degreen q m)"
+  using polymul_properties(3) by blast
+lemma polymul_norm:   
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "\<lbrakk> isnpoly p; isnpoly q\<rbrakk> \<Longrightarrow> isnpoly (polymul (p,q))"
+  using polymul_normh[of "p" "0" "q" "0"] isnpoly_def by simp
+
+lemma headconst_zero: "isnpolyh p n0 \<Longrightarrow> headconst p = 0\<^sub>N \<longleftrightarrow> p = 0\<^sub>p"
+  by (induct p arbitrary: n0 rule: headconst.induct, auto)
+
+lemma headconst_isnormNum: "isnpolyh p n0 \<Longrightarrow> isnormNum (headconst p)"
+  by (induct p arbitrary: n0, auto)
+
+lemma monic_eqI: assumes np: "isnpolyh p n0" 
+  shows "INum (headconst p) * Ipoly bs (fst (monic p)) = (Ipoly bs p ::'a::{ring_char_0,power,division_by_zero,field})"
+  unfolding monic_def Let_def
+proof(cases "headconst p = 0\<^sub>N", simp_all add: headconst_zero[OF np])
+  let ?h = "headconst p"
+  assume pz: "p \<noteq> 0\<^sub>p"
+  {assume hz: "INum ?h = (0::'a)"
+    from headconst_isnormNum[OF np] have norm: "isnormNum ?h" "isnormNum 0\<^sub>N" by simp_all
+    from isnormNum_unique[where ?'a = 'a, OF norm] hz have "?h = 0\<^sub>N" by simp
+    with headconst_zero[OF np] have "p =0\<^sub>p" by blast with pz have "False" by blast}
+  thus "INum (headconst p) = (0::'a) \<longrightarrow> \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = 0" by blast
+qed
+
+
+ 
+
+text{* polyneg is a negation and preserves normal form *}
+lemma polyneg[simp]: "Ipoly bs (polyneg p) = - Ipoly bs p"
+by (induct p rule: polyneg.induct, auto)
+
+lemma polyneg0: "isnpolyh p n \<Longrightarrow> ((~\<^sub>p p) = 0\<^sub>p) = (p = 0\<^sub>p)"
+  by (induct p arbitrary: n rule: polyneg.induct, auto simp add: Nneg_def)
+lemma polyneg_polyneg: "isnpolyh p n0 \<Longrightarrow> ~\<^sub>p (~\<^sub>p p) = p"
+  by (induct p arbitrary: n0 rule: polyneg.induct, auto)
+lemma polyneg_normh: "\<And>n. isnpolyh p n \<Longrightarrow> isnpolyh (polyneg p) n "
+by (induct p rule: polyneg.induct, auto simp add: polyneg0)
+
+lemma polyneg_norm: "isnpoly p \<Longrightarrow> isnpoly (polyneg p)"
+  using isnpoly_def polyneg_normh by simp
+
+
+text{* polysub is a substraction and preserves normalform *}
+lemma polysub[simp]: "Ipoly bs (polysub (p,q)) = (Ipoly bs p) - (Ipoly bs q)"
+by (simp add: polysub_def polyneg polyadd)
+lemma polysub_normh: "\<And> n0 n1. \<lbrakk> isnpolyh p n0 ; isnpolyh q n1\<rbrakk> \<Longrightarrow> isnpolyh (polysub(p,q)) (min n0 n1)"
+by (simp add: polysub_def polyneg_normh polyadd_normh)
+
+lemma polysub_norm: "\<lbrakk> isnpoly p; isnpoly q\<rbrakk> \<Longrightarrow> isnpoly (polysub(p,q))"
+  using polyadd_norm polyneg_norm by (simp add: polysub_def) 
+lemma polysub_same_0[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "isnpolyh p n0 \<Longrightarrow> polysub (p, p) = 0\<^sub>p"
+unfolding polysub_def split_def fst_conv snd_conv
+by (induct p arbitrary: n0,auto simp add: Let_def Nsub0[simplified Nsub_def])
+
+lemma polysub_0: 
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "\<lbrakk> isnpolyh p n0 ; isnpolyh q n1\<rbrakk> \<Longrightarrow> (p -\<^sub>p q = 0\<^sub>p) = (p = q)"
+  unfolding polysub_def split_def fst_conv snd_conv
+  apply (induct p q arbitrary: n0 n1 rule:polyadd.induct, simp_all add: Nsub0[simplified Nsub_def])
+  apply (clarsimp simp add: Let_def)
+  apply (case_tac "n < n'", simp_all)
+  apply (case_tac "n' < n", simp_all)
+  apply (erule impE)+
+  apply (rule_tac x="Suc n" in exI, simp)
+  apply (rule_tac x="n" in exI, simp)
+  apply (erule impE)+
+  apply (rule_tac x="n" in exI, simp)
+  apply (rule_tac x="Suc n" in exI, simp)
+  apply (erule impE)+
+  apply (rule_tac x="Suc n" in exI, simp)
+  apply (rule_tac x="n" in exI, simp)
+  apply (erule impE)+
+  apply (rule_tac x="Suc n" in exI, simp)
+  apply clarsimp
+  done
+
+text{* polypow is a power function and preserves normal forms *}
+lemma polypow[simp]: "Ipoly bs (polypow n p) = ((Ipoly bs p :: 'a::{ring_char_0,division_by_zero,field})) ^ n"
+proof(induct n rule: polypow.induct)
+  case 1 thus ?case by simp
+next
+  case (2 n)
+  let ?q = "polypow ((Suc n) div 2) p"
+  let ?d = "polymul(?q,?q)"
+  have "odd (Suc n) \<or> even (Suc n)" by simp
+  moreover 
+  {assume odd: "odd (Suc n)"
+    have th: "(Suc (Suc (Suc (0\<Colon>nat)) * (Suc n div Suc (Suc (0\<Colon>nat))))) = Suc n div 2 + Suc n div 2 + 1" by arith
+    from odd have "Ipoly bs (p ^\<^sub>p Suc n) = Ipoly bs (polymul(p, ?d))" by (simp add: Let_def)
+    also have "\<dots> = (Ipoly bs p) * (Ipoly bs p)^(Suc n div 2)*(Ipoly bs p)^(Suc n div 2)"
+      using "2.hyps" by simp
+    also have "\<dots> = (Ipoly bs p) ^ (Suc n div 2 + Suc n div 2 + 1)"
+      apply (simp only: power_add power_one_right) by simp
+    also have "\<dots> = (Ipoly bs p) ^ (Suc (Suc (Suc (0\<Colon>nat)) * (Suc n div Suc (Suc (0\<Colon>nat)))))"
+      by (simp only: th)
+    finally have ?case 
+    using odd_nat_div_two_times_two_plus_one[OF odd, symmetric] by simp  }
+  moreover 
+  {assume even: "even (Suc n)"
+    have th: "(Suc (Suc (0\<Colon>nat))) * (Suc n div Suc (Suc (0\<Colon>nat))) = Suc n div 2 + Suc n div 2" by arith
+    from even have "Ipoly bs (p ^\<^sub>p Suc n) = Ipoly bs ?d" by (simp add: Let_def)
+    also have "\<dots> = (Ipoly bs p) ^ (Suc n div 2 + Suc n div 2)"
+      using "2.hyps" apply (simp only: power_add) by simp
+    finally have ?case using even_nat_div_two_times_two[OF even] by (simp only: th)}
+  ultimately show ?case by blast
+qed
+
+lemma polypow_normh: 
+    assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "isnpolyh p n \<Longrightarrow> isnpolyh (polypow k p) n"
+proof (induct k arbitrary: n rule: polypow.induct)
+  case (2 k n)
+  let ?q = "polypow (Suc k div 2) p"
+  let ?d = "polymul (?q,?q)"
+  from prems have th1:"isnpolyh ?q n" and th2: "isnpolyh p n" by blast+
+  from polymul_normh[OF th1 th1] have dn: "isnpolyh ?d n" by simp
+  from polymul_normh[OF th2 dn] have on: "isnpolyh (polymul(p,?d)) n" by simp
+  from dn on show ?case by (simp add: Let_def)
+qed auto 
+
+lemma polypow_norm:   
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "isnpoly p \<Longrightarrow> isnpoly (polypow k p)"
+  by (simp add: polypow_normh isnpoly_def)
+
+text{* Finally the whole normalization*}
+
+lemma polynate[simp]: "Ipoly bs (polynate p) = (Ipoly bs p :: 'a ::{ring_char_0,division_by_zero,field})"
+by (induct p rule:polynate.induct, auto)
+
+lemma polynate_norm[simp]: 
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "isnpoly (polynate p)"
+  by (induct p rule: polynate.induct, simp_all add: polyadd_norm polymul_norm polysub_norm polyneg_norm polypow_norm) (simp_all add: isnpoly_def)
+
+text{* shift1 *}
+
+
+lemma shift1: "Ipoly bs (shift1 p) = Ipoly bs (Mul (Bound 0) p)"
+by (simp add: shift1_def polymul)
+
+lemma shift1_isnpoly: 
+  assumes pn: "isnpoly p" and pnz: "p \<noteq> 0\<^sub>p" shows "isnpoly (shift1 p) "
+  using pn pnz by (simp add: shift1_def isnpoly_def )
+
+lemma shift1_nz[simp]:"shift1 p \<noteq> 0\<^sub>p"
+  by (simp add: shift1_def)
+lemma funpow_shift1_isnpoly: 
+  "\<lbrakk> isnpoly p ; p \<noteq> 0\<^sub>p\<rbrakk> \<Longrightarrow> isnpoly (funpow n shift1 p)"
+  by (induct n arbitrary: p, auto simp add: shift1_isnpoly)
+
+lemma funpow_isnpolyh: 
+  assumes f: "\<And> p. isnpolyh p n \<Longrightarrow> isnpolyh (f p) n "and np: "isnpolyh p n"
+  shows "isnpolyh (funpow k f p) n"
+  using f np by (induct k arbitrary: p, auto)
+
+lemma funpow_shift1: "(Ipoly bs (funpow n shift1 p) :: 'a :: {ring_char_0,division_by_zero,field}) = Ipoly bs (Mul (Pw (Bound 0) n) p)"
+  by (induct n arbitrary: p, simp_all add: shift1_isnpoly shift1 power_Suc )
+
+lemma shift1_isnpolyh: "isnpolyh p n0 \<Longrightarrow> p\<noteq> 0\<^sub>p \<Longrightarrow> isnpolyh (shift1 p) 0"
+  using isnpolyh_mono[where n="n0" and n'="0" and p="p"] by (simp add: shift1_def)
+
+lemma funpow_shift1_1: 
+  "(Ipoly bs (funpow n shift1 p) :: 'a :: {ring_char_0,division_by_zero,field}) = Ipoly bs (funpow n shift1 1\<^sub>p *\<^sub>p p)"
+  by (simp add: funpow_shift1)
+
+lemma poly_cmul[simp]: "Ipoly bs (poly_cmul c p) = Ipoly bs (Mul (C c) p)"
+by (induct p  arbitrary: n0 rule: poly_cmul.induct, auto simp add: ring_simps)
+
+lemma behead:
+  assumes np: "isnpolyh p n"
+  shows "Ipoly bs (Add (Mul (head p) (Pw (Bound 0) (degree p))) (behead p)) = (Ipoly bs p :: 'a :: {ring_char_0,division_by_zero,field})"
+  using np
+proof (induct p arbitrary: n rule: behead.induct)
+  case (1 c p n) hence pn: "isnpolyh p n" by simp
+  from prems(2)[OF pn] 
+  have th:"Ipoly bs (Add (Mul (head p) (Pw (Bound 0) (degree p))) (behead p)) = Ipoly bs p" . 
+  then show ?case using "1.hyps" apply (simp add: Let_def,cases "behead p = 0\<^sub>p")
+    by (simp_all add: th[symmetric] ring_simps power_Suc)
+qed (auto simp add: Let_def)
+
+lemma behead_isnpolyh:
+  assumes np: "isnpolyh p n" shows "isnpolyh (behead p) n"
+  using np by (induct p rule: behead.induct, auto simp add: Let_def isnpolyh_mono)
+
+subsection{* Miscilanious lemmas about indexes, decrementation, substitution  etc ... *}
+lemma isnpolyh_polybound0: "isnpolyh p (Suc n) \<Longrightarrow> polybound0 p"
+proof(induct p arbitrary: n rule: polybound0.induct, auto)
+  case (goal1 c n p n')
+  hence "n = Suc (n - 1)" by simp
+  hence "isnpolyh p (Suc (n - 1))"  using `isnpolyh p n` by simp
+  with prems(2) show ?case by simp
+qed
+
+lemma isconstant_polybound0: "isnpolyh p n0 \<Longrightarrow> isconstant p \<longleftrightarrow> polybound0 p"
+by (induct p arbitrary: n0 rule: isconstant.induct, auto simp add: isnpolyh_polybound0)
+
+lemma decrpoly_zero[simp]: "decrpoly p = 0\<^sub>p \<longleftrightarrow> p = 0\<^sub>p" by (induct p, auto)
+
+lemma decrpoly_normh: "isnpolyh p n0 \<Longrightarrow> polybound0 p \<Longrightarrow> isnpolyh (decrpoly p) (n0 - 1)"
+  apply (induct p arbitrary: n0, auto)
+  apply (atomize)
+  apply (erule_tac x = "Suc nat" in allE)
+  apply auto
+  done
+
+lemma head_polybound0: "isnpolyh p n0 \<Longrightarrow> polybound0 (head p)"
+ by (induct p  arbitrary: n0 rule: head.induct, auto intro: isnpolyh_polybound0)
+
+lemma polybound0_I:
+  assumes nb: "polybound0 a"
+  shows "Ipoly (b#bs) a = Ipoly (b'#bs) a"
+using nb
+by (induct a rule: polybound0.induct) auto 
+lemma polysubst0_I:
+  shows "Ipoly (b#bs) (polysubst0 a t) = Ipoly ((Ipoly (b#bs) a)#bs) t"
+  by (induct t) simp_all
+
+lemma polysubst0_I':
+  assumes nb: "polybound0 a"
+  shows "Ipoly (b#bs) (polysubst0 a t) = Ipoly ((Ipoly (b'#bs) a)#bs) t"
+  by (induct t) (simp_all add: polybound0_I[OF nb, where b="b" and b'="b'"])
+
+lemma decrpoly: assumes nb: "polybound0 t"
+  shows "Ipoly (x#bs) t = Ipoly bs (decrpoly t)"
+  using nb by (induct t rule: decrpoly.induct, simp_all)
+
+lemma polysubst0_polybound0: assumes nb: "polybound0 t"
+  shows "polybound0 (polysubst0 t a)"
+using nb by (induct a rule: polysubst0.induct, auto)
+
+lemma degree0_polybound0: "isnpolyh p n \<Longrightarrow> degree p = 0 \<Longrightarrow> polybound0 p"
+  by (induct p arbitrary: n rule: degree.induct, auto simp add: isnpolyh_polybound0)
+
+fun maxindex :: "poly \<Rightarrow> nat" where
+  "maxindex (Bound n) = n + 1"
+| "maxindex (CN c n p) = max  (n + 1) (max (maxindex c) (maxindex p))"
+| "maxindex (Add p q) = max (maxindex p) (maxindex q)"
+| "maxindex (Sub p q) = max (maxindex p) (maxindex q)"
+| "maxindex (Mul p q) = max (maxindex p) (maxindex q)"
+| "maxindex (Neg p) = maxindex p"
+| "maxindex (Pw p n) = maxindex p"
+| "maxindex (C x) = 0"
+
+definition wf_bs :: "'a list \<Rightarrow> poly \<Rightarrow> bool" where
+  "wf_bs bs p = (length bs \<ge> maxindex p)"
+
+lemma wf_bs_coefficients: "wf_bs bs p \<Longrightarrow> \<forall> c \<in> set (coefficients p). wf_bs bs c"
+proof(induct p rule: coefficients.induct)
+  case (1 c p) 
+  show ?case 
+  proof
+    fix x assume xc: "x \<in> set (coefficients (CN c 0 p))"
+    hence "x = c \<or> x \<in> set (coefficients p)" by simp
+    moreover 
+    {assume "x = c" hence "wf_bs bs x" using "1.prems"  unfolding wf_bs_def by simp}
+    moreover 
+    {assume H: "x \<in> set (coefficients p)" 
+      from "1.prems" have "wf_bs bs p" unfolding wf_bs_def by simp
+      with "1.hyps" H have "wf_bs bs x" by blast }
+    ultimately  show "wf_bs bs x" by blast
+  qed
+qed simp_all
+
+lemma maxindex_coefficients: " \<forall>c\<in> set (coefficients p). maxindex c \<le> maxindex p"
+by (induct p rule: coefficients.induct, auto)
+
+lemma length_exists: "\<exists>xs. length xs = n" by (rule exI[where x="replicate n x"], simp)
+
+lemma wf_bs_I: "wf_bs bs p ==> Ipoly (bs@bs') p = Ipoly bs p"
+  unfolding wf_bs_def by (induct p, auto simp add: nth_append)
+
+lemma take_maxindex_wf: assumes wf: "wf_bs bs p" 
+  shows "Ipoly (take (maxindex p) bs) p = Ipoly bs p"
+proof-
+  let ?ip = "maxindex p"
+  let ?tbs = "take ?ip bs"
+  from wf have "length ?tbs = ?ip" unfolding wf_bs_def by simp
+  hence wf': "wf_bs ?tbs p" unfolding wf_bs_def by  simp
+  have eq: "bs = ?tbs @ (drop ?ip bs)" by simp
+  from wf_bs_I[OF wf', of "drop ?ip bs"] show ?thesis using eq by simp
+qed
+
+lemma decr_maxindex: "polybound0 p \<Longrightarrow> maxindex (decrpoly p) = maxindex p - 1"
+  by (induct p, auto)
+
+lemma wf_bs_insensitive: "length bs = length bs' \<Longrightarrow> wf_bs bs p = wf_bs bs' p"
+  unfolding wf_bs_def by simp
+
+lemma wf_bs_insensitive': "wf_bs (x#bs) p = wf_bs (y#bs) p"
+  unfolding wf_bs_def by simp
+
+
+
+lemma wf_bs_coefficients': "\<forall>c \<in> set (coefficients p). wf_bs bs c \<Longrightarrow> wf_bs (x#bs) p"
+by(induct p rule: coefficients.induct, auto simp add: wf_bs_def)
+lemma coefficients_Nil[simp]: "coefficients p \<noteq> []"
+  by (induct p rule: coefficients.induct, simp_all)
+
+
+lemma coefficients_head: "last (coefficients p) = head p"
+  by (induct p rule: coefficients.induct, auto)
+
+lemma wf_bs_decrpoly: "wf_bs bs (decrpoly p) \<Longrightarrow> wf_bs (x#bs) p"
+  unfolding wf_bs_def by (induct p rule: decrpoly.induct, auto)
+
+lemma length_le_list_ex: "length xs \<le> n \<Longrightarrow> \<exists> ys. length (xs @ ys) = n"
+  apply (rule exI[where x="replicate (n - length xs) z"])
+  by simp
+lemma isnpolyh_Suc_const:"isnpolyh p (Suc n) \<Longrightarrow> isconstant p"
+by (cases p, auto) (case_tac "nat", simp_all)
+
+lemma wf_bs_polyadd: "wf_bs bs p \<and> wf_bs bs q \<longrightarrow> wf_bs bs (p +\<^sub>p q)"
+  unfolding wf_bs_def 
+  apply (induct p q rule: polyadd.induct)
+  apply (auto simp add: Let_def)
+  done
+
+lemma wf_bs_polyul: "wf_bs bs p \<Longrightarrow> wf_bs bs q \<Longrightarrow> wf_bs bs (p *\<^sub>p q)"
+
+ unfolding wf_bs_def 
+  apply (induct p q arbitrary: bs rule: polymul.induct) 
+  apply (simp_all add: wf_bs_polyadd)
+  apply clarsimp
+  apply (rule wf_bs_polyadd[unfolded wf_bs_def, rule_format])
+  apply auto
+  done
+
+lemma wf_bs_polyneg: "wf_bs bs p \<Longrightarrow> wf_bs bs (~\<^sub>p p)"
+  unfolding wf_bs_def by (induct p rule: polyneg.induct, auto)
+
+lemma wf_bs_polysub: "wf_bs bs p \<Longrightarrow> wf_bs bs q \<Longrightarrow> wf_bs bs (p -\<^sub>p q)"
+  unfolding polysub_def split_def fst_conv snd_conv using wf_bs_polyadd wf_bs_polyneg by blast
+
+subsection{* Canonicity of polynomial representation, see lemma isnpolyh_unique*}
+
+definition "polypoly bs p = map (Ipoly bs) (coefficients p)"
+definition "polypoly' bs p = map ((Ipoly bs o decrpoly)) (coefficients p)"
+definition "poly_nate bs p = map ((Ipoly bs o decrpoly)) (coefficients (polynate p))"
+
+lemma coefficients_normh: "isnpolyh p n0 \<Longrightarrow> \<forall> q \<in> set (coefficients p). isnpolyh q n0"
+proof (induct p arbitrary: n0 rule: coefficients.induct)
+  case (1 c p n0)
+  have cp: "isnpolyh (CN c 0 p) n0" by fact
+  hence norm: "isnpolyh c 0" "isnpolyh p 0" "p \<noteq> 0\<^sub>p" "n0 = 0"
+    by (auto simp add: isnpolyh_mono[where n'=0])
+  from "1.hyps"[OF norm(2)] norm(1) norm(4)  show ?case by simp 
+qed auto
+
+lemma coefficients_isconst:
+  "isnpolyh p n \<Longrightarrow> \<forall>q\<in>set (coefficients p). isconstant q"
+  by (induct p arbitrary: n rule: coefficients.induct, 
+    auto simp add: isnpolyh_Suc_const)
+
+lemma polypoly_polypoly':
+  assumes np: "isnpolyh p n0"
+  shows "polypoly (x#bs) p = polypoly' bs p"
+proof-
+  let ?cf = "set (coefficients p)"
+  from coefficients_normh[OF np] have cn_norm: "\<forall> q\<in> ?cf. isnpolyh q n0" .
+  {fix q assume q: "q \<in> ?cf"
+    from q cn_norm have th: "isnpolyh q n0" by blast
+    from coefficients_isconst[OF np] q have "isconstant q" by blast
+    with isconstant_polybound0[OF th] have "polybound0 q" by blast}
+  hence "\<forall>q \<in> ?cf. polybound0 q" ..
+  hence "\<forall>q \<in> ?cf. Ipoly (x#bs) q = Ipoly bs (decrpoly q)"
+    using polybound0_I[where b=x and bs=bs and b'=y] decrpoly[where x=x and bs=bs]
+    by auto
+  
+  thus ?thesis unfolding polypoly_def polypoly'_def by simp 
+qed
+
+lemma polypoly_poly:
+  assumes np: "isnpolyh p n0" shows "Ipoly (x#bs) p = poly (polypoly (x#bs) p) x"
+  using np 
+by (induct p arbitrary: n0 bs rule: coefficients.induct, auto simp add: polypoly_def)
+
+lemma polypoly'_poly: 
+  assumes np: "isnpolyh p n0" shows "\<lparr>p\<rparr>\<^sub>p\<^bsup>x # bs\<^esup> = poly (polypoly' bs p) x"
+  using polypoly_poly[OF np, simplified polypoly_polypoly'[OF np]] .
+
+
+lemma polypoly_poly_polybound0:
+  assumes np: "isnpolyh p n0" and nb: "polybound0 p"
+  shows "polypoly bs p = [Ipoly bs p]"
+  using np nb unfolding polypoly_def 
+  by (cases p, auto, case_tac nat, auto)
+
+lemma head_isnpolyh: "isnpolyh p n0 \<Longrightarrow> isnpolyh (head p) n0" 
+  by (induct p rule: head.induct, auto)
+
+lemma headn_nz[simp]: "isnpolyh p n0 \<Longrightarrow> (headn p m = 0\<^sub>p) = (p = 0\<^sub>p)"
+  by (cases p,auto)
+
+lemma head_eq_headn0: "head p = headn p 0"
+  by (induct p rule: head.induct, simp_all)
+
+lemma head_nz[simp]: "isnpolyh p n0 \<Longrightarrow> (head p = 0\<^sub>p) = (p = 0\<^sub>p)"
+  by (simp add: head_eq_headn0)
+
+lemma isnpolyh_zero_iff: 
+  assumes nq: "isnpolyh p n0" and eq :"\<forall>bs. wf_bs bs p \<longrightarrow> \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a::{ring_char_0,power,division_by_zero,field})"
+  shows "p = 0\<^sub>p"
+using nq eq
+proof (induct n\<equiv>"maxindex p" arbitrary: p n0 rule: nat_less_induct)
+  fix n p n0
+  assume H: "\<forall>m<n. \<forall>p n0. isnpolyh p n0 \<longrightarrow>
+    (\<forall>bs. wf_bs bs p \<longrightarrow> \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a)) \<longrightarrow> m = maxindex p \<longrightarrow> p = 0\<^sub>p"
+    and np: "isnpolyh p n0" and zp: "\<forall>bs. wf_bs bs p \<longrightarrow> \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a)" and n: "n = maxindex p"
+  {assume nz: "n = 0"
+    then obtain c where "p = C c" using n np by (cases p, auto)
+    with zp np have "p = 0\<^sub>p" unfolding wf_bs_def by simp}
+  moreover
+  {assume nz: "n \<noteq> 0"
+    let ?h = "head p"
+    let ?hd = "decrpoly ?h"
+    let ?ihd = "maxindex ?hd"
+    from head_isnpolyh[OF np] head_polybound0[OF np] have h:"isnpolyh ?h n0" "polybound0 ?h" 
+      by simp_all
+    hence nhd: "isnpolyh ?hd (n0 - 1)" using decrpoly_normh by blast
+    
+    from maxindex_coefficients[of p] coefficients_head[of p, symmetric]
+    have mihn: "maxindex ?h \<le> n" unfolding n by auto
+    with decr_maxindex[OF h(2)] nz  have ihd_lt_n: "?ihd < n" by auto
+    {fix bs:: "'a list"  assume bs: "wf_bs bs ?hd"
+      let ?ts = "take ?ihd bs"
+      let ?rs = "drop ?ihd bs"
+      have ts: "wf_bs ?ts ?hd" using bs unfolding wf_bs_def by simp
+      have bs_ts_eq: "?ts@ ?rs = bs" by simp
+      from wf_bs_decrpoly[OF ts] have tsh: " \<forall>x. wf_bs (x#?ts) ?h" by simp
+      from ihd_lt_n have "ALL x. length (x#?ts) \<le> n" by simp
+      with length_le_list_ex obtain xs where xs:"length ((x#?ts) @ xs) = n" by blast
+      hence "\<forall> x. wf_bs ((x#?ts) @ xs) p" using n unfolding wf_bs_def by simp
+      with zp have "\<forall> x. Ipoly ((x#?ts) @ xs) p = 0" by blast
+      hence "\<forall> x. Ipoly (x#(?ts @ xs)) p = 0" by simp
+      with polypoly_poly[OF np, where ?'a = 'a] polypoly_polypoly'[OF np, where ?'a = 'a]
+      have "\<forall>x. poly (polypoly' (?ts @ xs) p) x = poly [] x"  by simp
+      hence "poly (polypoly' (?ts @ xs) p) = poly []" by (auto intro: ext) 
+      hence "\<forall> c \<in> set (coefficients p). Ipoly (?ts @ xs) (decrpoly c) = 0"
+	thm poly_zero
+	using poly_zero[where ?'a='a] by (simp add: polypoly'_def list_all_iff)
+      with coefficients_head[of p, symmetric]
+      have th0: "Ipoly (?ts @ xs) ?hd = 0" by simp
+      from bs have wf'': "wf_bs ?ts ?hd" unfolding wf_bs_def by simp
+      with th0 wf_bs_I[of ?ts ?hd xs] have "Ipoly ?ts ?hd = 0" by simp
+      with wf'' wf_bs_I[of ?ts ?hd ?rs] bs_ts_eq have "\<lparr>?hd\<rparr>\<^sub>p\<^bsup>bs\<^esup> = 0" by simp }
+    then have hdz: "\<forall>bs. wf_bs bs ?hd \<longrightarrow> \<lparr>?hd\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a)" by blast
+    
+    from H[rule_format, OF ihd_lt_n nhd] hdz have "?hd = 0\<^sub>p" by blast
+    hence "?h = 0\<^sub>p" by simp
+    with head_nz[OF np] have "p = 0\<^sub>p" by simp}
+  ultimately show "p = 0\<^sub>p" by blast
+qed
+
+lemma isnpolyh_unique:  
+  assumes np:"isnpolyh p n0" and nq: "isnpolyh q n1"
+  shows "(\<forall>bs. \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (\<lparr>q\<rparr>\<^sub>p\<^bsup>bs\<^esup> :: 'a::{ring_char_0,power,division_by_zero,field})) \<longleftrightarrow>  p = q"
+proof(auto)
+  assume H: "\<forall>bs. (\<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> ::'a)= \<lparr>q\<rparr>\<^sub>p\<^bsup>bs\<^esup>"
+  hence "\<forall>bs.\<lparr>p -\<^sub>p q\<rparr>\<^sub>p\<^bsup>bs\<^esup>= (0::'a)" by simp
+  hence "\<forall>bs. wf_bs bs (p -\<^sub>p q) \<longrightarrow> \<lparr>p -\<^sub>p q\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a)" 
+    using wf_bs_polysub[where p=p and q=q] by auto
+  with isnpolyh_zero_iff[OF polysub_normh[OF np nq]] polysub_0[OF np nq]
+  show "p = q" by blast
+qed
+
+
+text{* consequenses of unicity on the algorithms for polynomial normalization *}
+
+lemma polyadd_commute:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  and np: "isnpolyh p n0" and nq: "isnpolyh q n1" shows "p +\<^sub>p q = q +\<^sub>p p"
+  using isnpolyh_unique[OF polyadd_normh[OF np nq] polyadd_normh[OF nq np]] by simp
+
+lemma zero_normh: "isnpolyh 0\<^sub>p n" by simp
+lemma one_normh: "isnpolyh 1\<^sub>p n" by simp
+lemma polyadd_0[simp]: 
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  and np: "isnpolyh p n0" shows "p +\<^sub>p 0\<^sub>p = p" and "0\<^sub>p +\<^sub>p p = p"
+  using isnpolyh_unique[OF polyadd_normh[OF np zero_normh] np] 
+    isnpolyh_unique[OF polyadd_normh[OF zero_normh np] np] by simp_all
+
+lemma polymul_1[simp]: 
+    assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  and np: "isnpolyh p n0" shows "p *\<^sub>p 1\<^sub>p = p" and "1\<^sub>p *\<^sub>p p = p"
+  using isnpolyh_unique[OF polymul_normh[OF np one_normh] np] 
+    isnpolyh_unique[OF polymul_normh[OF one_normh np] np] by simp_all
+lemma polymul_0[simp]: 
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  and np: "isnpolyh p n0" shows "p *\<^sub>p 0\<^sub>p = 0\<^sub>p" and "0\<^sub>p *\<^sub>p p = 0\<^sub>p"
+  using isnpolyh_unique[OF polymul_normh[OF np zero_normh] zero_normh] 
+    isnpolyh_unique[OF polymul_normh[OF zero_normh np] zero_normh] by simp_all
+
+lemma polymul_commute: 
+    assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  and np:"isnpolyh p n0" and nq: "isnpolyh q n1"
+  shows "p *\<^sub>p q = q *\<^sub>p p"
+using isnpolyh_unique[OF polymul_normh[OF np nq] polymul_normh[OF nq np], where ?'a = "'a\<Colon>{ring_char_0,power,division_by_zero,field}"] by simp
+
+declare polyneg_polyneg[simp]
+  
+lemma isnpolyh_polynate_id[simp]: 
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  and np:"isnpolyh p n0" shows "polynate p = p"
+  using isnpolyh_unique[where ?'a= "'a::{ring_char_0,division_by_zero,field}", OF polynate_norm[of p, unfolded isnpoly_def] np] polynate[where ?'a = "'a::{ring_char_0,division_by_zero,field}"] by simp
+
+lemma polynate_idempotent[simp]: 
+    assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "polynate (polynate p) = polynate p"
+  using isnpolyh_polynate_id[OF polynate_norm[of p, unfolded isnpoly_def]] .
+
+lemma poly_nate_polypoly': "poly_nate bs p = polypoly' bs (polynate p)"
+  unfolding poly_nate_def polypoly'_def ..
+lemma poly_nate_poly: shows "poly (poly_nate bs p) = (\<lambda>x:: 'a ::{ring_char_0,division_by_zero,field}. \<lparr>p\<rparr>\<^sub>p\<^bsup>x # bs\<^esup>)"
+  using polypoly'_poly[OF polynate_norm[unfolded isnpoly_def], symmetric, of bs p]
+  unfolding poly_nate_polypoly' by (auto intro: ext)
+
+subsection{* heads, degrees and all that *}
+lemma degree_eq_degreen0: "degree p = degreen p 0"
+  by (induct p rule: degree.induct, simp_all)
+
+lemma degree_polyneg: assumes n: "isnpolyh p n"
+  shows "degree (polyneg p) = degree p"
+  using n
+  by (induct p arbitrary: n rule: polyneg.induct, simp_all) (case_tac na, auto)
+
+lemma degree_polyadd:
+  assumes np: "isnpolyh p n0" and nq: "isnpolyh q n1"
+  shows "degree (p +\<^sub>p q) \<le> max (degree p) (degree q)"
+using degreen_polyadd[OF np nq, where m= "0"] degree_eq_degreen0 by simp
+
+
+lemma degree_polysub: assumes np: "isnpolyh p n0" and nq: "isnpolyh q n1"
+  shows "degree (p -\<^sub>p q) \<le> max (degree p) (degree q)"
+proof-
+  from nq have nq': "isnpolyh (~\<^sub>p q) n1" using polyneg_normh by simp
+  from degree_polyadd[OF np nq'] show ?thesis by (simp add: polysub_def degree_polyneg[OF nq])
+qed
+
+lemma degree_polysub_samehead: 
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  and np: "isnpolyh p n0" and nq: "isnpolyh q n1" and h: "head p = head q" 
+  and d: "degree p = degree q"
+  shows "degree (p -\<^sub>p q) < degree p \<or> (p -\<^sub>p q = 0\<^sub>p)"
+unfolding polysub_def split_def fst_conv snd_conv
+using np nq h d
+proof(induct p q rule:polyadd.induct)
+  case (1 a b a' b') thus ?case by (simp add: Nsub_def Nsub0[simplified Nsub_def]) 
+next
+  case (2 a b c' n' p') 
+  let ?c = "(a,b)"
+  from prems have "degree (C ?c) = degree (CN c' n' p')" by simp
+  hence nz:"n' > 0" by (cases n', auto)
+  hence "head (CN c' n' p') = CN c' n' p'" by (cases n', auto)
+  with prems show ?case by simp
+next
+  case (3 c n p a' b') 
+  let ?c' = "(a',b')"
+  from prems have "degree (C ?c') = degree (CN c n p)" by simp
+  hence nz:"n > 0" by (cases n, auto)
+  hence "head (CN c n p) = CN c n p" by (cases n, auto)
+  with prems show ?case by simp
+next
+  case (4 c n p c' n' p')
+  hence H: "isnpolyh (CN c n p) n0" "isnpolyh (CN c' n' p') n1" 
+    "head (CN c n p) = head (CN c' n' p')" "degree (CN c n p) = degree (CN c' n' p')" by simp+
+  hence degc: "degree c = 0" and degc': "degree c' = 0" by simp_all  
+  hence degnc: "degree (~\<^sub>p c) = 0" and degnc': "degree (~\<^sub>p c') = 0" 
+    using H(1-2) degree_polyneg by auto
+  from H have cnh: "isnpolyh c (Suc n)" and c'nh: "isnpolyh c' (Suc n')"  by simp+
+  from degree_polysub[OF cnh c'nh, simplified polysub_def] degc degc' have degcmc': "degree (c +\<^sub>p  ~\<^sub>pc') = 0"  by simp
+  from H have pnh: "isnpolyh p n" and p'nh: "isnpolyh p' n'" by auto
+  have "n = n' \<or> n < n' \<or> n > n'" by arith
+  moreover
+  {assume nn': "n = n'"
+    have "n = 0 \<or> n >0" by arith
+    moreover {assume nz: "n = 0" hence ?case using prems by (auto simp add: Let_def degcmc')}
+    moreover {assume nz: "n > 0"
+      with nn' H(3) have  cc':"c = c'" and pp': "p = p'" by (cases n, auto)+
+      hence ?case using polysub_same_0[OF p'nh, simplified polysub_def split_def fst_conv snd_conv] polysub_same_0[OF c'nh, simplified polysub_def split_def fst_conv snd_conv] using nn' prems by (simp add: Let_def)}
+    ultimately have ?case by blast}
+  moreover
+  {assume nn': "n < n'" hence n'p: "n' > 0" by simp 
+    hence headcnp':"head (CN c' n' p') = CN c' n' p'"  by (cases n', simp_all)
+    have degcnp': "degree (CN c' n' p') = 0" and degcnpeq: "degree (CN c n p) = degree (CN c' n' p')" using prems by (cases n', simp_all)
+    hence "n > 0" by (cases n, simp_all)
+    hence headcnp: "head (CN c n p) = CN c n p" by (cases n, auto)
+    from H(3) headcnp headcnp' nn' have ?case by auto}
+  moreover
+  {assume nn': "n > n'"  hence np: "n > 0" by simp 
+    hence headcnp:"head (CN c n p) = CN c n p"  by (cases n, simp_all)
+    from prems have degcnpeq: "degree (CN c' n' p') = degree (CN c n p)" by simp
+    from np have degcnp: "degree (CN c n p) = 0" by (cases n, simp_all)
+    with degcnpeq have "n' > 0" by (cases n', simp_all)
+    hence headcnp': "head (CN c' n' p') = CN c' n' p'" by (cases n', auto)
+    from H(3) headcnp headcnp' nn' have ?case by auto}
+  ultimately show ?case  by blast
+qed auto 
+ 
+lemma shift1_head : "isnpolyh p n0 \<Longrightarrow> head (shift1 p) = head p"
+by (induct p arbitrary: n0 rule: head.induct, simp_all add: shift1_def)
+
+lemma funpow_shift1_head: "isnpolyh p n0 \<Longrightarrow> p\<noteq> 0\<^sub>p \<Longrightarrow> head (funpow k shift1 p) = head p"
+proof(induct k arbitrary: n0 p)
+  case (Suc k n0 p) hence "isnpolyh (shift1 p) 0" by (simp add: shift1_isnpolyh)
+  with prems have "head (funpow k shift1 (shift1 p)) = head (shift1 p)"
+    and "head (shift1 p) = head p" by (simp_all add: shift1_head) 
+  thus ?case by simp
+qed auto  
+
+lemma shift1_degree: "degree (shift1 p) = 1 + degree p"
+  by (simp add: shift1_def)
+lemma funpow_shift1_degree: "degree (funpow k shift1 p) = k + degree p "
+  by (induct k arbitrary: p, auto simp add: shift1_degree)
+
+lemma funpow_shift1_nz: "p \<noteq> 0\<^sub>p \<Longrightarrow> funpow n shift1 p \<noteq> 0\<^sub>p"
+  by (induct n arbitrary: p, simp_all add: funpow_def)
+
+lemma head_isnpolyh_Suc[simp]: "isnpolyh p (Suc n) \<Longrightarrow> head p = p"
+  by (induct p arbitrary: n rule: degree.induct, auto)
+lemma headn_0[simp]: "isnpolyh p n \<Longrightarrow> m < n \<Longrightarrow> headn p m = p"
+  by (induct p arbitrary: n rule: degreen.induct, auto)
+lemma head_isnpolyh_Suc': "n > 0 \<Longrightarrow> isnpolyh p n \<Longrightarrow> head p = p"
+  by (induct p arbitrary: n rule: degree.induct, auto)
+lemma head_head[simp]: "isnpolyh p n0 \<Longrightarrow> head (head p) = head p"
+  by (induct p rule: head.induct, auto)
+
+lemma polyadd_eq_const_degree: 
+  "\<lbrakk> isnpolyh p n0 ; isnpolyh q n1 ; polyadd (p,q) = C c\<rbrakk> \<Longrightarrow> degree p = degree q" 
+  using polyadd_eq_const_degreen degree_eq_degreen0 by simp
+
+lemma polyadd_head: assumes np: "isnpolyh p n0" and nq: "isnpolyh q n1"
+  and deg: "degree p \<noteq> degree q"
+  shows "head (p +\<^sub>p q) = (if degree p < degree q then head q else head p)"
+using np nq deg
+apply(induct p q arbitrary: n0 n1 rule: polyadd.induct,simp_all)
+apply (case_tac n', simp, simp)
+apply (case_tac n, simp, simp)
+apply (case_tac n, case_tac n', simp add: Let_def)
+apply (case_tac "pa +\<^sub>p p' = 0\<^sub>p")
+apply (clarsimp simp add: polyadd_eq_const_degree)
+apply clarsimp
+apply (erule_tac impE,blast)
+apply (erule_tac impE,blast)
+apply clarsimp
+apply simp
+apply (case_tac n', simp_all)
+done
+
+lemma polymul_head_polyeq: 
+   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "\<lbrakk>isnpolyh p n0; isnpolyh q n1 ; p \<noteq> 0\<^sub>p ; q \<noteq> 0\<^sub>p \<rbrakk> \<Longrightarrow> head (p *\<^sub>p q) = head p *\<^sub>p head q"
+proof (induct p q arbitrary: n0 n1 rule: polymul.induct)
+  case (2 a b c' n' p' n0 n1)
+  hence "isnpolyh (head (CN c' n' p')) n1" "isnormNum (a,b)"  by (simp_all add: head_isnpolyh)
+  thus ?case using prems by (cases n', auto) 
+next 
+  case (3 c n p a' b' n0 n1) 
+  hence "isnpolyh (head (CN c n p)) n0" "isnormNum (a',b')"  by (simp_all add: head_isnpolyh)
+  thus ?case using prems by (cases n, auto)
+next
+  case (4 c n p c' n' p' n0 n1)
+  hence norm: "isnpolyh p n" "isnpolyh c (Suc n)" "isnpolyh p' n'" "isnpolyh c' (Suc n')"
+    "isnpolyh (CN c n p) n" "isnpolyh (CN c' n' p') n'"
+    by simp_all
+  have "n < n' \<or> n' < n \<or> n = n'" by arith
+  moreover 
+  {assume nn': "n < n'" hence ?case 
+      thm prems
+      using norm 
+    prems(6)[rule_format, OF nn' norm(1,6)]
+    prems(7)[rule_format, OF nn' norm(2,6)] by (simp, cases n, simp,cases n', simp_all)}
+  moreover {assume nn': "n'< n"
+    hence stupid: "n' < n \<and> \<not> n < n'" by simp
+    hence ?case using norm prems(4) [rule_format, OF stupid norm(5,3)]
+      prems(5)[rule_format, OF stupid norm(5,4)] 
+      by (simp,cases n',simp,cases n,auto)}
+  moreover {assume nn': "n' = n"
+    hence stupid: "\<not> n' < n \<and> \<not> n < n'" by simp
+    from nn' polymul_normh[OF norm(5,4)] 
+    have ncnpc':"isnpolyh (CN c n p *\<^sub>p c') n" by (simp add: min_def)
+    from nn' polymul_normh[OF norm(5,3)] norm 
+    have ncnpp':"isnpolyh (CN c n p *\<^sub>p p') n" by simp
+    from nn' ncnpp' polymul_eq0_iff[OF norm(5,3)] norm(6)
+    have ncnpp0':"isnpolyh (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n" by simp 
+    from polyadd_normh[OF ncnpc' ncnpp0'] 
+    have nth: "isnpolyh ((CN c n p *\<^sub>p c') +\<^sub>p (CN 0\<^sub>p n (CN c n p *\<^sub>p p'))) n" 
+      by (simp add: min_def)
+    {assume np: "n > 0"
+      with nn' head_isnpolyh_Suc'[OF np nth]
+	head_isnpolyh_Suc'[OF np norm(5)] head_isnpolyh_Suc'[OF np norm(6)[simplified nn']]
+      have ?case by simp}
+    moreover
+    {moreover assume nz: "n = 0"
+      from polymul_degreen[OF norm(5,4), where m="0"]
+	polymul_degreen[OF norm(5,3), where m="0"] nn' nz degree_eq_degreen0
+      norm(5,6) degree_npolyhCN[OF norm(6)]
+    have dth:"degree (CN c 0 p *\<^sub>p c') < degree (CN 0\<^sub>p 0 (CN c 0 p *\<^sub>p p'))" by simp
+    hence dth':"degree (CN c 0 p *\<^sub>p c') \<noteq> degree (CN 0\<^sub>p 0 (CN c 0 p *\<^sub>p p'))" by simp
+    from polyadd_head[OF ncnpc'[simplified nz] ncnpp0'[simplified nz] dth'] dth
+    have ?case   using norm prems(2)[rule_format, OF stupid norm(5,3)]
+	prems(3)[rule_format, OF stupid norm(5,4)] nn' nz by simp }
+    ultimately have ?case by (cases n) auto} 
+  ultimately show ?case by blast
+qed simp_all
+
+lemma degree_coefficients: "degree p = length (coefficients p) - 1"
+  by(induct p rule: degree.induct, auto)
+
+lemma degree_head[simp]: "degree (head p) = 0"
+  by (induct p rule: head.induct, auto)
+
+lemma degree_CN: "isnpolyh p n \<Longrightarrow> degree (CN c n p) \<le> 1+ degree p"
+  by (cases n, simp_all)
+lemma degree_CN': "isnpolyh p n \<Longrightarrow> degree (CN c n p) \<ge>  degree p"
+  by (cases n, simp_all)
+
+lemma polyadd_different_degree: "\<lbrakk>isnpolyh p n0 ; isnpolyh q n1; degree p \<noteq> degree q\<rbrakk> \<Longrightarrow> degree (polyadd(p,q)) = max (degree p) (degree q)"
+  using polyadd_different_degreen degree_eq_degreen0 by simp
+
+lemma degreen_polyneg: "isnpolyh p n0 \<Longrightarrow> degreen (~\<^sub>p p) m = degreen p m"
+  by (induct p arbitrary: n0 rule: polyneg.induct, auto)
+
+lemma degree_polymul:
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  and np: "isnpolyh p n0" and nq: "isnpolyh q n1"
+  shows "degree (p *\<^sub>p q) \<le> degree p + degree q"
+  using polymul_degreen[OF np nq, where m="0"]  degree_eq_degreen0 by simp
+
+lemma polyneg_degree: "isnpolyh p n \<Longrightarrow> degree (polyneg p) = degree p"
+  by (induct p arbitrary: n rule: degree.induct, auto)
+
+lemma polyneg_head: "isnpolyh p n \<Longrightarrow> head(polyneg p) = polyneg (head p)"
+  by (induct p arbitrary: n rule: degree.induct, auto)
+
+subsection {* Correctness of polynomial pseudo division *}
+
+lemma polydivide_aux_real_domintros:
+  assumes call1: "\<lbrakk>s \<noteq> 0\<^sub>p; \<not> degree s < n; a = head s\<rbrakk> 
+  \<Longrightarrow> polydivide_aux_dom (a, n, p, k, s -\<^sub>p funpow (degree s - n) shift1 p)"
+  and call2 : "\<lbrakk>s \<noteq> 0\<^sub>p; \<not> degree s < n; a \<noteq> head s\<rbrakk> 
+  \<Longrightarrow> polydivide_aux_dom(a, n, p,Suc k, a *\<^sub>p s -\<^sub>p (head s *\<^sub>p funpow (degree s - n) shift1 p))"
+  shows "polydivide_aux_dom (a, n, p, k, s)"
+proof (rule accpI, erule polydivide_aux_rel.cases)
+  fix y aa ka na pa sa x xa xb
+  assume eqs: "y = (aa, na, pa, ka, sa -\<^sub>p xb)" "(a, n, p, k, s) = (aa, na, pa, ka, sa)"
+     and \<Gamma>1': "sa \<noteq> 0\<^sub>p" "x = head sa" "xa = degree sa" "\<not> xa < na" 
+    "xb = funpow (xa - na) shift1 pa" "aa = x"
+
+  hence \<Gamma>1: "s \<noteq> 0\<^sub>p" "a = head s" "xa = degree s" "\<not> degree s < n" "\<not> xa < na" 
+    "xb = funpow (xa - na) shift1 pa" "aa = x" by auto
+
+  with call1 have "polydivide_aux_dom (a, n, p, k, s -\<^sub>p funpow (degree s - n) shift1 p)"
+    by auto
+  with eqs \<Gamma>1 show "polydivide_aux_dom y" by auto
+next
+  fix y aa ka na pa sa x xa xb
+  assume eqs: "y = (aa, na, pa, Suc ka, aa *\<^sub>p sa -\<^sub>p (x *\<^sub>p xb))" 
+    "(a, n, p, k, s) =(aa, na, pa, ka, sa)"
+    and \<Gamma>2': "sa \<noteq> 0\<^sub>p" "x = head sa" "xa = degree sa" "\<not> xa < na"
+    "xb = funpow (xa - na) shift1 pa" "aa \<noteq> x"
+  hence \<Gamma>2: "s \<noteq> 0\<^sub>p" "a \<noteq> head s" "xa = degree s" "\<not> degree s < n"
+    "xb = funpow (xa - na) shift1 pa" "aa \<noteq> x" by auto
+  with call2 have "polydivide_aux_dom (a, n, p, Suc k, a *\<^sub>p s -\<^sub>p (head s *\<^sub>p funpow (degree s - n) shift1 p))" by auto
+  with eqs \<Gamma>2'  show "polydivide_aux_dom y" by auto
+qed
+
+lemma polydivide_aux_properties:
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  and np: "isnpolyh p n0" and ns: "isnpolyh s n1"
+  and ap: "head p = a" and ndp: "degree p = n" and pnz: "p \<noteq> 0\<^sub>p"
+  shows "polydivide_aux_dom (a,n,p,k,s) \<and> 
+  (polydivide_aux (a,n,p,k,s) = (k',r) \<longrightarrow> (k' \<ge> k) \<and> (degree r = 0 \<or> degree r < degree p) 
+          \<and> (\<exists>nr. isnpolyh r nr) \<and> (\<exists>q n1. isnpolyh q n1 \<and> ((polypow (k' - k) a) *\<^sub>p s = p *\<^sub>p q +\<^sub>p r)))"
+  using ns
+proof(induct d\<equiv>"degree s" arbitrary: s k k' r n1 rule: nat_less_induct)
+  fix d s k k' r n1
+  let ?D = "polydivide_aux_dom"
+  let ?dths = "?D (a, n, p, k, s)"
+  let ?qths = "\<exists>q n1. isnpolyh q n1 \<and> (a ^\<^sub>p (k' - k) *\<^sub>p s = p *\<^sub>p q +\<^sub>p r)"
+  let ?qrths = "polydivide_aux (a, n, p, k, s) = (k', r) \<longrightarrow>  k \<le> k' \<and> (degree r = 0 \<or> degree r < degree p) 
+    \<and> (\<exists>nr. isnpolyh r nr) \<and> ?qths"
+  let ?ths = "?dths \<and> ?qrths"
+  let ?b = "head s"
+  let ?p' = "funpow (d - n) shift1 p"
+  let ?xdn = "funpow (d - n) shift1 1\<^sub>p"
+  let ?akk' = "a ^\<^sub>p (k' - k)"
+  assume H: "\<forall>m<d. \<forall>x xa xb xc xd.
+    isnpolyh x xd \<longrightarrow>
+    m = degree x \<longrightarrow> ?D (a, n, p, xa, x) \<and>
+    (polydivide_aux (a, n, p, xa, x) = (xb, xc) \<longrightarrow>
+    xa \<le> xb \<and> (degree xc = 0 \<or> degree xc < degree p) \<and> 
+    (\<exists> nr. isnpolyh xc nr) \<and>
+    (\<exists>q n1. isnpolyh q n1 \<and> a ^\<^sub>p xb - xa *\<^sub>p x = p *\<^sub>p q +\<^sub>p xc))"
+    and ns: "isnpolyh s n1" and ds: "d = degree s"
+  from np have np0: "isnpolyh p 0" 
+    using isnpolyh_mono[where n="n0" and n'="0" and p="p"]  by simp
+  have np': "isnpolyh ?p' 0" using funpow_shift1_isnpoly[OF np0[simplified isnpoly_def[symmetric]] pnz, where n="d -n"] isnpoly_def by simp
+  have headp': "head ?p' = head p" using funpow_shift1_head[OF np pnz] by simp
+  from funpow_shift1_isnpoly[where p="1\<^sub>p"] have nxdn: "isnpolyh ?xdn 0" by (simp add: isnpoly_def)
+  from polypow_normh [OF head_isnpolyh[OF np0], where k="k' - k"] ap 
+  have nakk':"isnpolyh ?akk' 0" by blast
+  {assume sz: "s = 0\<^sub>p"
+    hence dom: ?dths apply - apply (rule polydivide_aux_real_domintros) apply simp_all done
+    from polydivide_aux.psimps[OF dom] sz
+    have ?qrths using np apply clarsimp by (rule exI[where x="0\<^sub>p"], simp)
+    hence ?ths using dom by blast}
+  moreover
+  {assume sz: "s \<noteq> 0\<^sub>p"
+    {assume dn: "d < n"
+      with sz ds  have dom:"?dths" by - (rule polydivide_aux_real_domintros,simp_all) 
+      from polydivide_aux.psimps[OF dom] sz dn ds
+      have "?qrths" using ns ndp np by auto (rule exI[where x="0\<^sub>p"],simp)
+      with dom have ?ths by blast}
+    moreover 
+    {assume dn': "\<not> d < n" hence dn: "d \<ge> n" by arith
+      have degsp': "degree s = degree ?p'" 
+	using ds dn ndp funpow_shift1_degree[where k = "d - n" and p="p"] by simp
+      {assume ba: "?b = a"
+	hence headsp': "head s = head ?p'" using ap headp' by simp
+	have nr: "isnpolyh (s -\<^sub>p ?p') 0" using polysub_normh[OF ns np'] by simp
+	from ds degree_polysub_samehead[OF ns np' headsp' degsp']
+	have "degree (s -\<^sub>p ?p') < d \<or> s -\<^sub>p ?p' = 0\<^sub>p" by simp
+	moreover 
+	{assume deglt:"degree (s -\<^sub>p ?p') < d"
+	  from  H[rule_format, OF deglt nr,simplified] 
+	  have domsp: "?D (a, n, p, k, s -\<^sub>p ?p')" by blast 
+	  have dom: ?dths apply (rule polydivide_aux_real_domintros) 
+	    using ba ds dn' domsp by simp_all
+	  from polydivide_aux.psimps[OF dom] sz dn' ba ds
+	  have eq: "polydivide_aux (a,n,p,k,s) = polydivide_aux (a,n,p,k, s -\<^sub>p ?p')"
+	    by (simp add: Let_def)
+	  {assume h1: "polydivide_aux (a, n, p, k, s) = (k', r)"
+	    from H[rule_format, OF deglt nr, where xa = "k" and xb="k'" and xc="r", simplified]
+	      trans[OF eq[symmetric] h1]
+	    have kk': "k \<le> k'" and nr:"\<exists>nr. isnpolyh r nr" and dr: "degree r = 0 \<or> degree r < degree p"
+	      and q1:"\<exists>q nq. isnpolyh q nq \<and> (a ^\<^sub>p k' - k *\<^sub>p (s -\<^sub>p ?p') = p *\<^sub>p q +\<^sub>p r)" by auto
+	    from q1 obtain q n1 where nq: "isnpolyh q n1" 
+	      and asp:"a^\<^sub>p (k' - k) *\<^sub>p (s -\<^sub>p ?p') = p *\<^sub>p q +\<^sub>p r"  by blast
+	    from nr obtain nr where nr': "isnpolyh r nr" by blast
+	    from polymul_normh[OF nakk' ns] have nakks': "isnpolyh (a ^\<^sub>p (k' - k) *\<^sub>p s) 0" by simp
+	    from polyadd_normh[OF polymul_normh[OF nakk' nxdn] nq]
+	    have nq': "isnpolyh (?akk' *\<^sub>p ?xdn +\<^sub>p q) 0" by simp
+	    from polyadd_normh[OF polymul_normh[OF np 
+	      polyadd_normh[OF polymul_normh[OF nakk' nxdn] nq]] nr']
+	    have nqr': "isnpolyh (p *\<^sub>p (?akk' *\<^sub>p ?xdn +\<^sub>p q) +\<^sub>p r) 0" by simp 
+	    from asp have "\<forall> (bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a^\<^sub>p (k' - k) *\<^sub>p (s -\<^sub>p ?p')) = 
+	      Ipoly bs (p *\<^sub>p q +\<^sub>p r)" by simp
+	    hence " \<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a^\<^sub>p (k' - k)*\<^sub>p s) = 
+	      Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs ?p' + Ipoly bs p * Ipoly bs q + Ipoly bs r" 
+	      by (simp add: ring_simps)
+	    hence " \<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = 
+	      Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs (funpow (d - n) shift1 1\<^sub>p *\<^sub>p p) 
+	      + Ipoly bs p * Ipoly bs q + Ipoly bs r"
+	      by (auto simp only: funpow_shift1_1) 
+	    hence "\<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = 
+	      Ipoly bs p * (Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs (funpow (d - n) shift1 1\<^sub>p) 
+	      + Ipoly bs q) + Ipoly bs r" by (simp add: ring_simps)
+	    hence "\<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = 
+	      Ipoly bs (p *\<^sub>p ((a^\<^sub>p (k' - k)) *\<^sub>p (funpow (d - n) shift1 1\<^sub>p) +\<^sub>p q) +\<^sub>p r)" by simp
+	    with isnpolyh_unique[OF nakks' nqr']
+	    have "a ^\<^sub>p (k' - k) *\<^sub>p s = 
+	      p *\<^sub>p ((a^\<^sub>p (k' - k)) *\<^sub>p (funpow (d - n) shift1 1\<^sub>p) +\<^sub>p q) +\<^sub>p r" by blast
+	    hence ?qths using nq'
+	      apply (rule_tac x="(a^\<^sub>p (k' - k)) *\<^sub>p (funpow (d - n) shift1 1\<^sub>p) +\<^sub>p q" in exI)
+	      apply (rule_tac x="0" in exI) by simp
+	    with kk' nr dr have "k \<le> k' \<and> (degree r = 0 \<or> degree r < degree p) \<and> (\<exists>nr. isnpolyh r nr) \<and> ?qths"
+	      by blast } hence ?qrths by blast
+	  with dom have ?ths by blast} 
+	moreover 
+	{assume spz:"s -\<^sub>p ?p' = 0\<^sub>p"
+	  hence domsp: "?D (a, n, p, k, s -\<^sub>p ?p')" 
+	    apply (simp) by (rule polydivide_aux_real_domintros, simp_all)
+	  have dom: ?dths apply (rule polydivide_aux_real_domintros) 
+	    using ba ds dn' domsp by simp_all
+	  from spz isnpolyh_unique[OF polysub_normh[OF ns np'], where q="0\<^sub>p", symmetric, where ?'a = "'a::{ring_char_0,division_by_zero,field}"]
+	  have " \<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs s = Ipoly bs ?p'" by simp
+	  hence "\<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs s = Ipoly bs (?xdn *\<^sub>p p)" using np nxdn apply simp
+	    by (simp only: funpow_shift1_1) simp
+	  hence sp': "s = ?xdn *\<^sub>p p" using isnpolyh_unique[OF ns polymul_normh[OF nxdn np]] by blast
+	  {assume h1: "polydivide_aux (a,n,p,k,s) = (k',r)"
+	    from polydivide_aux.psimps[OF dom] sz dn' ba ds
+	    have eq: "polydivide_aux (a,n,p,k,s) = polydivide_aux (a,n,p,k, s -\<^sub>p ?p')"
+	      by (simp add: Let_def)
+	    also have "\<dots> = (k,0\<^sub>p)" using polydivide_aux.psimps[OF domsp] spz by simp
+	    finally have "(k',r) = (k,0\<^sub>p)" using h1 by simp
+	    with sp' ns np nxdn polyadd_0(1)[OF polymul_normh[OF np nxdn]]
+	      polyadd_0(2)[OF polymul_normh[OF np nxdn]] have ?qrths
+	      apply auto
+	      apply (rule exI[where x="?xdn"]) 	      
+	      apply auto
+	      apply (rule polymul_commute)
+	      apply simp_all
+	      done}
+	  with dom have ?ths by blast}
+	ultimately have ?ths by blast }
+      moreover
+      {assume ba: "?b \<noteq> a"
+	from polysub_normh[OF polymul_normh[OF head_isnpolyh[OF np0, simplified ap] ns] 
+	  polymul_normh[OF head_isnpolyh[OF ns] np']]
+	have nth: "isnpolyh ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) 0" by(simp add: min_def)
+	have nzths: "a *\<^sub>p s \<noteq> 0\<^sub>p" "?b *\<^sub>p ?p' \<noteq> 0\<^sub>p"
+	  using polymul_eq0_iff[OF head_isnpolyh[OF np0, simplified ap] ns] 
+	    polymul_eq0_iff[OF head_isnpolyh[OF ns] np']head_nz[OF np0] ap pnz sz head_nz[OF ns]
+	    funpow_shift1_nz[OF pnz] by simp_all
+	from polymul_head_polyeq[OF head_isnpolyh[OF np] ns] head_nz[OF np] sz ap head_head[OF np] pnz
+	  polymul_head_polyeq[OF head_isnpolyh[OF ns] np'] head_nz [OF ns] sz funpow_shift1_nz[OF pnz, where n="d - n"]
+	have hdth: "head (a *\<^sub>p s) = head (?b *\<^sub>p ?p')" 
+	  using head_head[OF ns] funpow_shift1_head[OF np pnz]
+	    polymul_commute[OF head_isnpolyh[OF np] head_isnpolyh[OF ns]]
+	  by (simp add: ap)
+	from polymul_degreen[OF head_isnpolyh[OF np] ns, where m="0"]
+	  head_nz[OF np] pnz sz ap[symmetric]
+	  funpow_shift1_nz[OF pnz, where n="d - n"]
+	  polymul_degreen[OF head_isnpolyh[OF ns] np', where m="0"]  head_nz[OF ns]
+	  ndp ds[symmetric] dn
+	have degth: "degree (a *\<^sub>p s) = degree (?b *\<^sub>p ?p') "
+	  by (simp add: degree_eq_degreen0[symmetric] funpow_shift1_degree)
+	{assume dth: "degree ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) < d"
+	  have th: "?D (a, n, p, Suc k, (a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p'))"
+	    using H[rule_format, OF dth nth, simplified] by blast 
+	  have dom: ?dths
+	    using ba ds dn' th apply simp apply (rule polydivide_aux_real_domintros)  
+	    using ba ds dn'  by simp_all
+	  from polysub_normh[OF polymul_normh[OF head_isnpolyh[OF np] ns] polymul_normh[OF head_isnpolyh[OF ns]np']]
+	  ap have nasbp': "isnpolyh ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) 0" by simp
+	  {assume h1:"polydivide_aux (a,n,p,k,s) = (k', r)"
+	    from h1  polydivide_aux.psimps[OF dom] sz dn' ba ds
+	    have eq:"polydivide_aux (a,n,p,Suc k,(a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) = (k',r)"
+	      by (simp add: Let_def)
+	    with H[rule_format, OF dth nasbp', simplified, where xa="Suc k" and xb="k'" and xc="r"]
+	    obtain q nq nr where kk': "Suc k \<le> k'" and nr: "isnpolyh r nr" and nq: "isnpolyh q nq" 
+	      and dr: "degree r = 0 \<or> degree r < degree p"
+	      and qr: "a ^\<^sub>p (k' - Suc k) *\<^sub>p ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) = p *\<^sub>p q +\<^sub>p r" by auto
+	    from kk' have kk'':"Suc (k' - Suc k) = k' - k" by arith
+	    {fix bs:: "'a::{ring_char_0,division_by_zero,field} list"
+	      
+	    from qr isnpolyh_unique[OF polypow_normh[OF head_isnpolyh[OF np], where k="k' - Suc k", simplified ap] nasbp', symmetric]
+	    have "Ipoly bs (a ^\<^sub>p (k' - Suc k) *\<^sub>p ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p'))) = Ipoly bs (p *\<^sub>p q +\<^sub>p r)" by simp
+	    hence "Ipoly bs a ^ (Suc (k' - Suc k)) * Ipoly bs s = Ipoly bs p * Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?p' + Ipoly bs r"
+	      by (simp add: ring_simps power_Suc)
+	    hence "Ipoly bs a ^ (k' - k)  * Ipoly bs s = Ipoly bs p * Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?xdn * Ipoly bs p + Ipoly bs r"
+	      by (simp add:kk'' funpow_shift1_1[where n="d - n" and p="p"])
+	    hence "Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = Ipoly bs p * (Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?xdn) + Ipoly bs r"
+	      by (simp add: ring_simps)}
+	    hence ieq:"\<forall>(bs :: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = 
+	      Ipoly bs (p *\<^sub>p (q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn)) +\<^sub>p r)" by auto 
+	    let ?q = "q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn)"
+	    from polyadd_normh[OF nq polymul_normh[OF polymul_normh[OF polypow_normh[OF head_isnpolyh[OF np], where k="k' - Suc k"] head_isnpolyh[OF ns], simplified ap ] nxdn]]
+	    have nqw: "isnpolyh ?q 0" by simp
+	    from ieq isnpolyh_unique[OF polymul_normh[OF polypow_normh[OF head_isnpolyh[OF np], where k="k' - k"] ns, simplified ap] polyadd_normh[OF polymul_normh[OF np nqw] nr]]
+	    have asth: "(a ^\<^sub>p (k' - k) *\<^sub>p s) = p *\<^sub>p (q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn)) +\<^sub>p r" by blast
+	    from dr kk' nr h1 asth nqw have ?qrths apply simp
+	      apply (rule conjI)
+	      apply (rule exI[where x="nr"], simp)
+	      apply (rule exI[where x="(q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn))"], simp)
+	      apply (rule exI[where x="0"], simp)
+	      done}
+	  hence ?qrths by blast
+	  with dom have ?ths by blast}
+	moreover 
+	{assume spz: "a *\<^sub>p s -\<^sub>p (?b *\<^sub>p ?p') = 0\<^sub>p"
+	  hence domsp: "?D (a, n, p, Suc k, a *\<^sub>p s -\<^sub>p (?b *\<^sub>p ?p'))" 
+	    apply (simp) by (rule polydivide_aux_real_domintros, simp_all)
+	  have dom: ?dths using sz ba dn' ds domsp 
+	    by - (rule polydivide_aux_real_domintros, simp_all)
+	  {fix bs :: "'a::{ring_char_0,division_by_zero,field} list"
+	    from isnpolyh_unique[OF nth, where ?'a="'a" and q="0\<^sub>p",simplified,symmetric] spz
+	  have "Ipoly bs (a*\<^sub>p s) = Ipoly bs ?b * Ipoly bs ?p'" by simp
+	  hence "Ipoly bs (a*\<^sub>p s) = Ipoly bs (?b *\<^sub>p ?xdn) * Ipoly bs p" 
+	    by (simp add: funpow_shift1_1[where n="d - n" and p="p"])
+	  hence "Ipoly bs (a*\<^sub>p s) = Ipoly bs (p *\<^sub>p (?b *\<^sub>p ?xdn))" by simp
+	}
+	hence hth: "\<forall> (bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a*\<^sub>p s) = Ipoly bs (p *\<^sub>p (?b *\<^sub>p ?xdn))" ..
+	  from hth
+	  have asq: "a *\<^sub>p s = p *\<^sub>p (?b *\<^sub>p ?xdn)" 
+	    using isnpolyh_unique[where ?'a = "'a::{ring_char_0,division_by_zero,field}", OF polymul_normh[OF head_isnpolyh[OF np] ns] 
+                    polymul_normh[OF np polymul_normh[OF head_isnpolyh[OF ns] nxdn]],
+	      simplified ap] by simp
+	  {assume h1: "polydivide_aux (a,n,p,k,s) = (k', r)"
+	  from h1 sz ds ba dn' spz polydivide_aux.psimps[OF dom] polydivide_aux.psimps[OF domsp] 
+	  have "(k',r) = (Suc k, 0\<^sub>p)" by (simp add: Let_def)
+	  with h1 np head_isnpolyh[OF np, simplified ap] ns polymul_normh[OF head_isnpolyh[OF ns] nxdn]
+	    polymul_normh[OF np polymul_normh[OF head_isnpolyh[OF ns] nxdn]] asq
+	  have ?qrths apply (clarsimp simp add: Let_def)
+	    apply (rule exI[where x="?b *\<^sub>p ?xdn"]) apply simp
+	    apply (rule exI[where x="0"], simp)
+	    done}
+	hence ?qrths by blast
+	with dom have ?ths by blast}
+	ultimately have ?ths using  degree_polysub_samehead[OF polymul_normh[OF head_isnpolyh[OF np0, simplified ap] ns] polymul_normh[OF head_isnpolyh[OF ns] np'] hdth degth] polymul_degreen[OF head_isnpolyh[OF np] ns, where m="0"]
+	  head_nz[OF np] pnz sz ap[symmetric] ds[symmetric] 
+	  by (simp add: degree_eq_degreen0[symmetric]) blast }
+      ultimately have ?ths by blast
+    }
+    ultimately have ?ths by blast}
+  ultimately show ?ths by blast
+qed
+
+lemma polydivide_properties: 
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  and np: "isnpolyh p n0" and ns: "isnpolyh s n1" and pnz: "p \<noteq> 0\<^sub>p"
+  shows "(\<exists> k r. polydivide s p = (k,r) \<and> (\<exists>nr. isnpolyh r nr) \<and> (degree r = 0 \<or> degree r < degree p) 
+  \<and> (\<exists>q n1. isnpolyh q n1 \<and> ((polypow k (head p)) *\<^sub>p s = p *\<^sub>p q +\<^sub>p r)))"
+proof-
+  have trv: "head p = head p" "degree p = degree p" by simp_all
+  from polydivide_aux_properties[OF np ns trv pnz, where k="0"] 
+  have d: "polydivide_aux_dom (head p, degree p, p, 0, s)" by blast
+  from polydivide_def[where s="s" and p="p"] polydivide_aux.psimps[OF d]
+  have ex: "\<exists> k r. polydivide s p = (k,r)" by auto
+  then obtain k r where kr: "polydivide s p = (k,r)" by blast
+  from trans[OF meta_eq_to_obj_eq[OF polydivide_def[where s="s" and p="p"], symmetric] kr]
+    polydivide_aux_properties[OF np ns trv pnz, where k="0" and k'="k" and r="r"]
+  have "(degree r = 0 \<or> degree r < degree p) \<and>
+   (\<exists>nr. isnpolyh r nr) \<and> (\<exists>q n1. isnpolyh q n1 \<and> head p ^\<^sub>p k - 0 *\<^sub>p s = p *\<^sub>p q +\<^sub>p r)" by blast
+  with kr show ?thesis 
+    apply -
+    apply (rule exI[where x="k"])
+    apply (rule exI[where x="r"])
+    apply simp
+    done
+qed
+
+subsection{* More about polypoly and pnormal etc *}
+
+definition "isnonconstant p = (\<not> isconstant p)"
+
+lemma last_map: "xs \<noteq> [] ==> last (map f xs) = f (last xs)" by (induct xs, auto)
+
+lemma isnonconstant_pnormal_iff: assumes nc: "isnonconstant p" 
+  shows "pnormal (polypoly bs p) \<longleftrightarrow> Ipoly bs (head p) \<noteq> 0" 
+proof
+  let ?p = "polypoly bs p"  
+  assume H: "pnormal ?p"
+  have csz: "coefficients p \<noteq> []" using nc by (cases p, auto)
+  
+  from coefficients_head[of p] last_map[OF csz, of "Ipoly bs"]  
+    pnormal_last_nonzero[OF H]
+  show "Ipoly bs (head p) \<noteq> 0" by (simp add: polypoly_def)
+next
+  assume h: "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0"
+  let ?p = "polypoly bs p"
+  have csz: "coefficients p \<noteq> []" using nc by (cases p, auto)
+  hence pz: "?p \<noteq> []" by (simp add: polypoly_def) 
+  hence lg: "length ?p > 0" by simp
+  from h coefficients_head[of p] last_map[OF csz, of "Ipoly bs"] 
+  have lz: "last ?p \<noteq> 0" by (simp add: polypoly_def)
+  from pnormal_last_length[OF lg lz] show "pnormal ?p" .
+qed
+
+lemma isnonconstant_coefficients_length: "isnonconstant p \<Longrightarrow> length (coefficients p) > 1"
+  unfolding isnonconstant_def
+  apply (cases p, simp_all)
+  apply (case_tac nat, auto)
+  done
+lemma isnonconstant_nonconstant: assumes inc: "isnonconstant p"
+  shows "nonconstant (polypoly bs p) \<longleftrightarrow> Ipoly bs (head p) \<noteq> 0"
+proof
+  let ?p = "polypoly bs p"
+  assume nc: "nonconstant ?p"
+  from isnonconstant_pnormal_iff[OF inc, of bs] nc
+  show "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0" unfolding nonconstant_def by blast
+next
+  let ?p = "polypoly bs p"
+  assume h: "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0"
+  from isnonconstant_pnormal_iff[OF inc, of bs] h
+  have pn: "pnormal ?p" by blast
+  {fix x assume H: "?p = [x]" 
+    from H have "length (coefficients p) = 1" unfolding polypoly_def by auto
+    with isnonconstant_coefficients_length[OF inc] have False by arith}
+  thus "nonconstant ?p" using pn unfolding nonconstant_def by blast  
+qed
+
+lemma pnormal_length: "p\<noteq>[] \<Longrightarrow> pnormal p \<longleftrightarrow> length (pnormalize p) = length p"
+  unfolding pnormal_def
+ apply (induct p rule: pnormalize.induct, simp_all)
+ apply (case_tac "p=[]", simp_all)
+ done
+
+lemma degree_degree: assumes inc: "isnonconstant p"
+  shows "degree p = Polynomial_List.degree (polypoly bs p) \<longleftrightarrow> \<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0"
+proof
+  let  ?p = "polypoly bs p"
+  assume H: "degree p = Polynomial_List.degree ?p"
+  from isnonconstant_coefficients_length[OF inc] have pz: "?p \<noteq> []"
+    unfolding polypoly_def by auto
+  from H degree_coefficients[of p] isnonconstant_coefficients_length[OF inc]
+  have lg:"length (pnormalize ?p) = length ?p"
+    unfolding Polynomial_List.degree_def polypoly_def by simp
+  hence "pnormal ?p" using pnormal_length[OF pz] by blast 
+  with isnonconstant_pnormal_iff[OF inc]  
+  show "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0" by blast
+next
+  let  ?p = "polypoly bs p"  
+  assume H: "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0"
+  with isnonconstant_pnormal_iff[OF inc] have "pnormal ?p" by blast
+  with degree_coefficients[of p] isnonconstant_coefficients_length[OF inc]
+  show "degree p = Polynomial_List.degree ?p" 
+    unfolding polypoly_def pnormal_def Polynomial_List.degree_def by auto
+qed
+
+section{* Swaps ; Division by a certain variable *}
+consts swap:: "nat \<Rightarrow> nat \<Rightarrow> poly \<Rightarrow> poly"
+primrec
+  "swap n m (C x) = C x"
+  "swap n m (Bound k) = Bound (if k = n then m else if k=m then n else k)"
+  "swap n m (Neg t) = Neg (swap n m t)"
+  "swap n m (Add s t) = Add (swap n m s) (swap n m t)"
+  "swap n m (Sub s t) = Sub (swap n m s) (swap n m t)"
+  "swap n m (Mul s t) = Mul (swap n m s) (swap n m t)"
+  "swap n m (Pw t k) = Pw (swap n m t) k"
+  "swap n m (CN c k p) = CN (swap n m c) (if k = n then m else if k=m then n else k)
+  (swap n m p)"
+
+lemma swap: assumes nbs: "n < length bs" and mbs: "m < length bs"
+  shows "Ipoly bs (swap n m t) = Ipoly ((bs[n:= bs!m])[m:= bs!n]) t"
+proof (induct t)
+  case (Bound k) thus ?case using nbs mbs by simp 
+next
+  case (CN c k p) thus ?case using nbs mbs by simp 
+qed simp_all
+lemma swap_swap_id[simp]: "swap n m (swap m n t) = t"
+  by (induct t,simp_all)
+
+lemma swap_commute: "swap n m p = swap m n p" by (induct p, simp_all)
+
+lemma swap_same_id[simp]: "swap n n t = t"
+  by (induct t, simp_all)
+
+definition "swapnorm n m t = polynate (swap n m t)"
+
+lemma swapnorm: assumes nbs: "n < length bs" and mbs: "m < length bs"
+  shows "((Ipoly bs (swapnorm n m t) :: 'a\<Colon>{ring_char_0,division_by_zero,field})) = Ipoly ((bs[n:= bs!m])[m:= bs!n]) t"
+  using swap[OF prems] swapnorm_def by simp
+
+lemma swapnorm_isnpoly[simp]: 
+    assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  shows "isnpoly (swapnorm n m p)"
+  unfolding swapnorm_def by simp
+
+definition "polydivideby n s p = 
+    (let ss = swapnorm 0 n s ; sp = swapnorm 0 n p ; h = head sp; (k,r) = polydivide ss sp
+     in (k,swapnorm 0 n h,swapnorm 0 n r))"
+
+lemma swap_nz [simp]: " (swap n m p = 0\<^sub>p) = (p = 0\<^sub>p)" by (induct p, simp_all)
+
+consts isweaknpoly :: "poly \<Rightarrow> bool"
+recdef isweaknpoly "measure size"
+  "isweaknpoly (C c) = True"
+  "isweaknpoly (CN c n p) \<longleftrightarrow> isweaknpoly c \<and> isweaknpoly p"
+  "isweaknpoly p = False"	
+
+lemma isnpolyh_isweaknpoly: "isnpolyh p n0 \<Longrightarrow> isweaknpoly p" 
+  by (induct p arbitrary: n0, auto)
+
+lemma swap_isweanpoly: "isweaknpoly p \<Longrightarrow> isweaknpoly (swap n m p)" 
+  by (induct p, auto)
+
+end
\ No newline at end of file
--- a/src/HOL/FunDef.thy	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/FunDef.thy	Mon Oct 26 09:03:57 2009 +0100
@@ -9,20 +9,20 @@
 uses
   "Tools/prop_logic.ML"
   "Tools/sat_solver.ML"
-  ("Tools/Function/fundef_lib.ML")
-  ("Tools/Function/fundef_common.ML")
+  ("Tools/Function/function_lib.ML")
+  ("Tools/Function/function_common.ML")
   ("Tools/Function/inductive_wrap.ML")
   ("Tools/Function/context_tree.ML")
-  ("Tools/Function/fundef_core.ML")
+  ("Tools/Function/function_core.ML")
   ("Tools/Function/sum_tree.ML")
   ("Tools/Function/mutual.ML")
   ("Tools/Function/pattern_split.ML")
-  ("Tools/Function/fundef.ML")
-  ("Tools/Function/auto_term.ML")
+  ("Tools/Function/function.ML")
+  ("Tools/Function/relation.ML")
   ("Tools/Function/measure_functions.ML")
   ("Tools/Function/lexicographic_order.ML")
   ("Tools/Function/pat_completeness.ML")
-  ("Tools/Function/fundef_datatype.ML")
+  ("Tools/Function/fun.ML")
   ("Tools/Function/induction_scheme.ML")
   ("Tools/Function/termination.ML")
   ("Tools/Function/decompose.ML")
@@ -104,25 +104,25 @@
   "wf R \<Longrightarrow> wfP (in_rel R)"
   by (simp add: wfP_def)
 
-use "Tools/Function/fundef_lib.ML"
-use "Tools/Function/fundef_common.ML"
+use "Tools/Function/function_lib.ML"
+use "Tools/Function/function_common.ML"
 use "Tools/Function/inductive_wrap.ML"
 use "Tools/Function/context_tree.ML"
-use "Tools/Function/fundef_core.ML"
+use "Tools/Function/function_core.ML"
 use "Tools/Function/sum_tree.ML"
 use "Tools/Function/mutual.ML"
 use "Tools/Function/pattern_split.ML"
-use "Tools/Function/auto_term.ML"
-use "Tools/Function/fundef.ML"
+use "Tools/Function/relation.ML"
+use "Tools/Function/function.ML"
 use "Tools/Function/pat_completeness.ML"
-use "Tools/Function/fundef_datatype.ML"
+use "Tools/Function/fun.ML"
 use "Tools/Function/induction_scheme.ML"
 
 setup {* 
-  Fundef.setup
+  Function.setup
   #> Pat_Completeness.setup
-  #> FundefDatatype.setup
-  #> InductionScheme.setup
+  #> Function_Fun.setup
+  #> Induction_Scheme.setup
 *}
 
 subsection {* Measure Functions *}
@@ -142,7 +142,7 @@
 by (rule is_measure_trivial)
 
 use "Tools/Function/lexicographic_order.ML"
-setup LexicographicOrder.setup 
+setup Lexicographic_Order.setup 
 
 
 subsection {* Congruence Rules *}
@@ -320,7 +320,7 @@
 
 ML_val -- "setup inactive"
 {*
-  Context.theory_map (FundefCommon.set_termination_prover (ScnpReconstruct.decomp_scnp 
+  Context.theory_map (Function_Common.set_termination_prover (ScnpReconstruct.decomp_scnp 
   [ScnpSolve.MAX, ScnpSolve.MIN, ScnpSolve.MS])) 
 *}
 
--- a/src/HOL/HOL.thy	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/HOL.thy	Mon Oct 26 09:03:57 2009 +0100
@@ -2002,8 +2002,12 @@
 *} "solve goal by normalization"
 
 
+subsection {* Counterexample Search Units *}
+
 subsubsection {* Quickcheck *}
 
+quickcheck_params [size = 5, iterations = 50]
+
 ML {*
 structure Quickcheck_RecFun_Simps = Named_Thms
 (
@@ -2014,37 +2018,8 @@
 
 setup Quickcheck_RecFun_Simps.setup
 
-setup {*
-  Quickcheck.add_generator ("SML", Codegen.test_term)
-*}
 
-quickcheck_params [size = 5, iterations = 50]
-
-subsection {* Preprocessing for the predicate compiler *}
-
-ML {*
-structure Predicate_Compile_Alternative_Defs = Named_Thms
-(
-  val name = "code_pred_def"
-  val description = "alternative definitions of constants for the Predicate Compiler"
-)
-*}
-
-ML {*
-structure Predicate_Compile_Inline_Defs = Named_Thms
-(
-  val name = "code_pred_inline"
-  val description = "inlining definitions for the Predicate Compiler"
-)
-*}
-
-setup {*
-  Predicate_Compile_Alternative_Defs.setup
-  #> Predicate_Compile_Inline_Defs.setup
-  #> Predicate_Compile_Preproc_Const_Defs.setup
-*}
-
-subsection {* Nitpick setup *}
+subsubsection {* Nitpick setup *}
 
 text {* This will be relocated once Nitpick is moved to HOL. *}
 
@@ -2079,6 +2054,31 @@
 *}
 
 
+subsection {* Preprocessing for the predicate compiler *}
+
+ML {*
+structure Predicate_Compile_Alternative_Defs = Named_Thms
+(
+  val name = "code_pred_def"
+  val description = "alternative definitions of constants for the Predicate Compiler"
+)
+*}
+
+ML {*
+structure Predicate_Compile_Inline_Defs = Named_Thms
+(
+  val name = "code_pred_inline"
+  val description = "inlining definitions for the Predicate Compiler"
+)
+*}
+
+setup {*
+  Predicate_Compile_Alternative_Defs.setup
+  #> Predicate_Compile_Inline_Defs.setup
+  #> Predicate_Compile_Preproc_Const_Defs.setup
+*}
+
+
 subsection {* Legacy tactics and ML bindings *}
 
 ML {*
--- a/src/HOL/IsaMakefile	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/IsaMakefile	Mon Oct 26 09:03:57 2009 +0100
@@ -156,15 +156,14 @@
   Tools/Datatype/datatype_realizer.ML \
   Tools/Datatype/datatype_rep_proofs.ML \
   Tools/dseq.ML \
-  Tools/Function/auto_term.ML \
   Tools/Function/context_tree.ML \
   Tools/Function/decompose.ML \
   Tools/Function/descent.ML \
-  Tools/Function/fundef_common.ML \
-  Tools/Function/fundef_core.ML \
-  Tools/Function/fundef_datatype.ML \
-  Tools/Function/fundef_lib.ML \
-  Tools/Function/fundef.ML \
+  Tools/Function/function_common.ML \
+  Tools/Function/function_core.ML \
+  Tools/Function/function_lib.ML \
+  Tools/Function/function.ML \
+  Tools/Function/fun.ML \
   Tools/Function/induction_scheme.ML \
   Tools/Function/inductive_wrap.ML \
   Tools/Function/lexicographic_order.ML \
@@ -172,6 +171,7 @@
   Tools/Function/mutual.ML \
   Tools/Function/pat_completeness.ML \
   Tools/Function/pattern_split.ML \
+  Tools/Function/relation.ML \
   Tools/Function/scnp_reconstruct.ML \
   Tools/Function/scnp_solve.ML \
   Tools/Function/size.ML \
@@ -358,7 +358,7 @@
   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
+  Library/OptionalSugar.thy Library/SML_Quickcheck.thy
 	@cd Library; $(ISABELLE_TOOL) usedir $(OUT)/HOL Library
 
 
--- a/src/HOL/Library/Library.thy	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Library/Library.thy	Mon Oct 26 09:03:57 2009 +0100
@@ -50,6 +50,7 @@
   Ramsey
   Reflection
   RBT
+  SML_Quickcheck
   State_Monad
   Sum_Of_Squares
   Univ_Poly
--- a/src/HOL/Library/Multiset.thy	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Library/Multiset.thy	Mon Oct 26 09:03:57 2009 +0100
@@ -1607,7 +1607,7 @@
       rtac @{thm nonempty_plus} ORELSE' rtac @{thm nonempty_single}
 
   val regroup_munion_conv =
-      FundefLib.regroup_conv @{const_name Multiset.Mempty} @{const_name plus}
+      Function_Lib.regroup_conv @{const_name Multiset.Mempty} @{const_name plus}
         (map (fn t => t RS eq_reflection) (@{thms union_ac} @ @{thms empty_idemp}))
 
   fun unfold_pwleq_tac i =
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/SML_Quickcheck.thy	Mon Oct 26 09:03:57 2009 +0100
@@ -0,0 +1,12 @@
+
+header {* Install quickcheck of SML code generator *}
+
+theory SML_Quickcheck
+imports Main
+begin
+
+setup {*
+  Quickcheck.add_generator ("SML", Codegen.test_term)
+*}
+
+end
--- a/src/HOL/Nominal/nominal_datatype.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Nominal/nominal_datatype.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -567,7 +567,7 @@
     val rep_set_names'' = map (Sign.full_bname thy3) rep_set_names';
 
     val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy4) =
-        Inductive.add_inductive_global (serial_string ())
+        Inductive.add_inductive_global (serial ())
           {quiet_mode = false, verbose = false, kind = Thm.internalK,
            alt_name = Binding.name big_rep_name, coind = false, no_elim = true, no_ind = false,
            skip_mono = true, fork_mono = false}
@@ -1506,7 +1506,7 @@
 
     val ({intrs = rec_intrs, elims = rec_elims, raw_induct = rec_induct, ...}, thy11) =
       thy10 |>
-        Inductive.add_inductive_global (serial_string ())
+        Inductive.add_inductive_global (serial ())
           {quiet_mode = #quiet config, verbose = false, kind = Thm.internalK,
            alt_name = Binding.name big_rec_name, coind = false, no_elim = false, no_ind = false,
            skip_mono = true, fork_mono = false}
--- a/src/HOL/Nominal/nominal_primrec.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Nominal/nominal_primrec.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -280,7 +280,7 @@
       else primrec_err ("functions " ^ commas_quote names2 ^
         "\nare not mutually recursive");
     val (defs_thms, lthy') = lthy |>
-      set_group ? LocalTheory.set_group (serial_string ()) |>
+      set_group ? LocalTheory.set_group (serial ()) |>
       fold_map (apfst (snd o snd) oo
         LocalTheory.define Thm.definitionK o fst) defs';
     val qualify = Binding.qualify false
--- a/src/HOL/Predicate.thy	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Predicate.thy	Mon Oct 26 09:03:57 2009 +0100
@@ -471,49 +471,49 @@
   "is_empty (A \<squnion> B) \<longleftrightarrow> is_empty A \<and> is_empty B"
   by (auto simp add: is_empty_def intro: sup_eq_bot_eq1 sup_eq_bot_eq2)
 
-definition singleton :: "'a pred \<Rightarrow> 'a" where
-  "singleton A = (if \<exists>!x. eval A x then THE x. eval A x else undefined)"
+definition singleton :: "(unit => 'a) \<Rightarrow> 'a pred \<Rightarrow> 'a" where
+  "singleton dfault A = (if \<exists>!x. eval A x then THE x. eval A x else dfault ())"
 
 lemma singleton_eqI:
-  "\<exists>!x. eval A x \<Longrightarrow> eval A x \<Longrightarrow> singleton A = x"
+  "\<exists>!x. eval A x \<Longrightarrow> eval A x \<Longrightarrow> singleton dfault A = x"
   by (auto simp add: singleton_def)
 
 lemma eval_singletonI:
-  "\<exists>!x. eval A x \<Longrightarrow> eval A (singleton A)"
+  "\<exists>!x. eval A x \<Longrightarrow> eval A (singleton dfault A)"
 proof -
   assume assm: "\<exists>!x. eval A x"
   then obtain x where "eval A x" ..
-  moreover with assm have "singleton A = x" by (rule singleton_eqI)
+  moreover with assm have "singleton dfault A = x" by (rule singleton_eqI)
   ultimately show ?thesis by simp 
 qed
 
 lemma single_singleton:
-  "\<exists>!x. eval A x \<Longrightarrow> single (singleton A) = A"
+  "\<exists>!x. eval A x \<Longrightarrow> single (singleton dfault A) = A"
 proof -
   assume assm: "\<exists>!x. eval A x"
-  then have "eval A (singleton A)"
+  then have "eval A (singleton dfault A)"
     by (rule eval_singletonI)
-  moreover from assm have "\<And>x. eval A x \<Longrightarrow> singleton A = x"
+  moreover from assm have "\<And>x. eval A x \<Longrightarrow> singleton dfault A = x"
     by (rule singleton_eqI)
-  ultimately have "eval (single (singleton A)) = eval A"
+  ultimately have "eval (single (singleton dfault A)) = eval A"
     by (simp (no_asm_use) add: single_def expand_fun_eq) blast
   then show ?thesis by (simp add: eval_inject)
 qed
 
 lemma singleton_undefinedI:
-  "\<not> (\<exists>!x. eval A x) \<Longrightarrow> singleton A = undefined"
+  "\<not> (\<exists>!x. eval A x) \<Longrightarrow> singleton dfault A = dfault ()"
   by (simp add: singleton_def)
 
 lemma singleton_bot:
-  "singleton \<bottom> = undefined"
+  "singleton dfault \<bottom> = dfault ()"
   by (auto simp add: bot_pred_def intro: singleton_undefinedI)
 
 lemma singleton_single:
-  "singleton (single x) = x"
+  "singleton dfault (single x) = x"
   by (auto simp add: intro: singleton_eqI singleI elim: singleE)
 
 lemma singleton_sup_single_single:
-  "singleton (single x \<squnion> single y) = (if x = y then x else undefined)"
+  "singleton dfault (single x \<squnion> single y) = (if x = y then x else dfault ())"
 proof (cases "x = y")
   case True then show ?thesis by (simp add: singleton_single)
 next
@@ -523,25 +523,25 @@
   by (auto intro: supI1 supI2 singleI)
   with False have "\<not> (\<exists>!z. eval (single x \<squnion> single y) z)"
     by blast
-  then have "singleton (single x \<squnion> single y) = undefined"
+  then have "singleton dfault (single x \<squnion> single y) = dfault ()"
     by (rule singleton_undefinedI)
   with False show ?thesis by simp
 qed
 
 lemma singleton_sup_aux:
-  "singleton (A \<squnion> B) = (if A = \<bottom> then singleton B
-    else if B = \<bottom> then singleton A
-    else singleton
-      (single (singleton A) \<squnion> single (singleton B)))"
+  "singleton dfault (A \<squnion> B) = (if A = \<bottom> then singleton dfault B
+    else if B = \<bottom> then singleton dfault A
+    else singleton dfault
+      (single (singleton dfault A) \<squnion> single (singleton dfault B)))"
 proof (cases "(\<exists>!x. eval A x) \<and> (\<exists>!y. eval B y)")
   case True then show ?thesis by (simp add: single_singleton)
 next
   case False
   from False have A_or_B:
-    "singleton A = undefined \<or> singleton B = undefined"
+    "singleton dfault A = dfault () \<or> singleton dfault B = dfault ()"
     by (auto intro!: singleton_undefinedI)
-  then have rhs: "singleton
-    (single (singleton A) \<squnion> single (singleton B)) = undefined"
+  then have rhs: "singleton dfault
+    (single (singleton dfault A) \<squnion> single (singleton dfault B)) = dfault ()"
     by (auto simp add: singleton_sup_single_single singleton_single)
   from False have not_unique:
     "\<not> (\<exists>!x. eval A x) \<or> \<not> (\<exists>!y. eval B y)" by simp
@@ -551,7 +551,7 @@
       by (blast elim: not_bot)
     with True not_unique have "\<not> (\<exists>!x. eval (A \<squnion> B) x)"
       by (auto simp add: sup_pred_def bot_pred_def)
-    then have "singleton (A \<squnion> B) = undefined" by (rule singleton_undefinedI)
+    then have "singleton dfault (A \<squnion> B) = dfault ()" by (rule singleton_undefinedI)
     with True rhs show ?thesis by simp
   next
     case False then show ?thesis by auto
@@ -559,10 +559,10 @@
 qed
 
 lemma singleton_sup:
-  "singleton (A \<squnion> B) = (if A = \<bottom> then singleton B
-    else if B = \<bottom> then singleton A
-    else if singleton A = singleton B then singleton A else undefined)"
-using singleton_sup_aux [of A B] by (simp only: singleton_sup_single_single)
+  "singleton dfault (A \<squnion> B) = (if A = \<bottom> then singleton dfault B
+    else if B = \<bottom> then singleton dfault A
+    else if singleton dfault A = singleton dfault B then singleton dfault A else dfault ())"
+using singleton_sup_aux [of dfault A B] by (simp only: singleton_sup_single_single)
 
 
 subsubsection {* Derived operations *}
@@ -743,36 +743,43 @@
   "is_empty (Seq f) \<longleftrightarrow> null (f ())"
   by (simp add: null_is_empty Seq_def)
 
-primrec the_only :: "'a seq \<Rightarrow> 'a" where
-  [code del]: "the_only Empty = undefined"
-  | "the_only (Insert x P) = (if is_empty P then x else let y = singleton P in if x = y then x else undefined)"
-  | "the_only (Join P xq) = (if is_empty P then the_only xq else if null xq then singleton P
-       else let x = singleton P; y = the_only xq in
-       if x = y then x else undefined)"
+primrec the_only :: "(unit \<Rightarrow> 'a) \<Rightarrow> 'a seq \<Rightarrow> 'a" where
+  [code del]: "the_only dfault Empty = dfault ()"
+  | "the_only dfault (Insert x P) = (if is_empty P then x else let y = singleton dfault P in if x = y then x else dfault ())"
+  | "the_only dfault (Join P xq) = (if is_empty P then the_only dfault xq else if null xq then singleton dfault P
+       else let x = singleton dfault P; y = the_only dfault xq in
+       if x = y then x else dfault ())"
 
 lemma the_only_singleton:
-  "the_only xq = singleton (pred_of_seq xq)"
+  "the_only dfault xq = singleton dfault (pred_of_seq xq)"
   by (induct xq)
     (auto simp add: singleton_bot singleton_single is_empty_def
     null_is_empty Let_def singleton_sup)
 
 lemma singleton_code [code]:
-  "singleton (Seq f) = (case f ()
-   of Empty \<Rightarrow> undefined
+  "singleton dfault (Seq f) = (case f ()
+   of Empty \<Rightarrow> dfault ()
     | Insert x P \<Rightarrow> if is_empty P then x
-        else let y = singleton P in
-          if x = y then x else undefined
-    | Join P xq \<Rightarrow> if is_empty P then the_only xq
-        else if null xq then singleton P
-        else let x = singleton P; y = the_only xq in
-          if x = y then x else undefined)"
+        else let y = singleton dfault P in
+          if x = y then x else dfault ()
+    | Join P xq \<Rightarrow> if is_empty P then the_only dfault xq
+        else if null xq then singleton dfault P
+        else let x = singleton dfault P; y = the_only dfault xq in
+          if x = y then x else dfault ())"
   by (cases "f ()")
    (auto simp add: Seq_def the_only_singleton is_empty_def
       null_is_empty singleton_bot singleton_single singleton_sup Let_def)
 
-lemma meta_fun_cong:
-"f == g ==> f x == g x"
-by simp
+definition not_unique :: "'a pred => 'a"
+where
+  [code del]: "not_unique A = (THE x. eval A x)"
+
+definition the :: "'a pred => 'a"
+where
+  [code del]: "the A = (THE x. eval A x)"
+
+lemma the_eq[code]: "the A = singleton (\<lambda>x. not_unique A) A"
+by (auto simp add: the_def singleton_def not_unique_def)
 
 ML {*
 signature PREDICATE =
@@ -819,6 +826,8 @@
 code_const Seq and Empty and Insert and Join
   (Eval "Predicate.Seq" and "Predicate.Empty" and "Predicate.Insert/ (_,/ _)" and "Predicate.Join/ (_,/ _)")
 
+code_abort not_unique
+
 text {* dummy setup for @{text code_pred} and @{text values} keywords *}
 
 ML {*
@@ -852,6 +861,6 @@
 
 hide (open) type pred seq
 hide (open) const Pred eval single bind is_empty singleton if_pred not_pred
-  Empty Insert Join Seq member pred_of_seq "apply" adjunct null the_only eq map
+  Empty Insert Join Seq member pred_of_seq "apply" adjunct null the_only eq map not_unique the
 
 end
--- a/src/HOL/Product_Type.thy	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Product_Type.thy	Mon Oct 26 09:03:57 2009 +0100
@@ -6,7 +6,7 @@
 header {* Cartesian products *}
 
 theory Product_Type
-imports Inductive
+imports Inductive Nat
 uses
   ("Tools/split_rule.ML")
   ("Tools/inductive_set.ML")
--- a/src/HOL/SizeChange/sct.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/SizeChange/sct.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -112,7 +112,7 @@
     end
 
 fun bind_many [] = I
-  | bind_many vs = FundefLib.tupled_lambda (foldr1 HOLogic.mk_prod vs)
+  | bind_many vs = Function_Lib.tupled_lambda (foldr1 HOLogic.mk_prod vs)
 
 (* Builds relation descriptions from a relation definition *)
 fun mk_reldescs (Abs a) =
--- a/src/HOL/Tools/Datatype/datatype_abs_proofs.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/Datatype/datatype_abs_proofs.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -153,7 +153,7 @@
         (descr' ~~ recTs ~~ rec_sets') ([], 0);
 
     val ({intrs = rec_intrs, elims = rec_elims, ...}, thy1) =
-        Inductive.add_inductive_global (serial_string ())
+        Inductive.add_inductive_global (serial ())
           {quiet_mode = #quiet config, verbose = false, kind = Thm.internalK,
             alt_name = Binding.name big_rec_name', coind = false, no_elim = false, no_ind = true,
             skip_mono = true, fork_mono = false}
@@ -321,7 +321,7 @@
                 fns2 @ (flat (Library.drop (i + 1, case_dummy_fns))) )));
           val ([def_thm], thy') =
             thy
-            |> Sign.declare_const [] decl |> snd
+            |> Sign.declare_const decl |> snd
             |> (PureThy.add_defs false o map Thm.no_attributes) [def];
 
         in (defs @ [def_thm], thy')
--- a/src/HOL/Tools/Datatype/datatype_rep_proofs.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/Datatype/datatype_rep_proofs.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -170,7 +170,7 @@
         ((1 upto (length constrs)) ~~ constrs)) (descr' ~~ rep_set_names');
 
     val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy2) =
-        Inductive.add_inductive_global (serial_string ())
+        Inductive.add_inductive_global (serial ())
           {quiet_mode = #quiet config, verbose = false, kind = Thm.internalK,
            alt_name = Binding.name big_rec_name, coind = false, no_elim = true, no_ind = false,
            skip_mono = true, fork_mono = false}
--- a/src/HOL/Tools/Function/auto_term.ML	Fri Oct 23 13:23:18 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,36 +0,0 @@
-(*  Title:      HOL/Tools/Function/auto_term.ML
-    Author:     Alexander Krauss, TU Muenchen
-
-A package for general recursive function definitions.
-Method "relation" to commence a termination proof using a user-specified relation.
-*)
-
-signature FUNDEF_RELATION =
-sig
-  val relation_tac: Proof.context -> term -> int -> tactic
-  val setup: theory -> theory
-end
-
-structure FundefRelation : FUNDEF_RELATION =
-struct
-
-fun inst_thm ctxt rel st =
-    let
-      val cert = Thm.cterm_of (ProofContext.theory_of ctxt)
-      val rel' = cert (singleton (Variable.polymorphic ctxt) rel)
-      val st' = Thm.incr_indexes (#maxidx (Thm.rep_cterm rel') + 1) st
-      val Rvar = cert (Var (the_single (Term.add_vars (prop_of st') [])))
-    in 
-      Drule.cterm_instantiate [(Rvar, rel')] st' 
-    end
-
-fun relation_tac ctxt rel i = 
-    TRY (FundefCommon.apply_termination_rule ctxt i)
-    THEN PRIMITIVE (inst_thm ctxt rel)
-
-val setup =
-  Method.setup @{binding relation}
-    (Args.term >> (fn rel => fn ctxt => SIMPLE_METHOD' (relation_tac ctxt rel)))
-    "proves termination using a user-specified wellfounded relation"
-
-end
--- a/src/HOL/Tools/Function/context_tree.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/Function/context_tree.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -5,15 +5,15 @@
 Builds and traverses trees of nested contexts along a term.
 *)
 
-signature FUNDEF_CTXTREE =
+signature FUNCTION_CTXTREE =
 sig
     type ctxt = (string * typ) list * thm list (* poor man's contexts: fixes + assumes *)
     type ctx_tree
 
     (* FIXME: This interface is a mess and needs to be cleaned up! *)
-    val get_fundef_congs : Proof.context -> thm list
-    val add_fundef_cong : thm -> Context.generic -> Context.generic
-    val map_fundef_congs : (thm list -> thm list) -> Context.generic -> Context.generic
+    val get_function_congs : Proof.context -> thm list
+    val add_function_cong : thm -> Context.generic -> Context.generic
+    val map_function_congs : (thm list -> thm list) -> Context.generic -> Context.generic
 
     val cong_add: attribute
     val cong_del: attribute
@@ -36,15 +36,15 @@
     val rewrite_by_tree : theory -> term -> thm -> (thm * thm) list -> ctx_tree -> thm * (thm * thm) list
 end
 
-structure FundefCtxTree : FUNDEF_CTXTREE =
+structure Function_Ctx_Tree : FUNCTION_CTXTREE =
 struct
 
 type ctxt = (string * typ) list * thm list
 
-open FundefCommon
-open FundefLib
+open Function_Common
+open Function_Lib
 
-structure FundefCongs = GenericDataFun
+structure FunctionCongs = GenericDataFun
 (
   type T = thm list
   val empty = []
@@ -52,14 +52,14 @@
   fun merge _ = Thm.merge_thms
 );
 
-val get_fundef_congs = FundefCongs.get o Context.Proof
-val map_fundef_congs = FundefCongs.map
-val add_fundef_cong = FundefCongs.map o Thm.add_thm
+val get_function_congs = FunctionCongs.get o Context.Proof
+val map_function_congs = FunctionCongs.map
+val add_function_cong = FunctionCongs.map o Thm.add_thm
 
 (* congruence rules *)
 
-val cong_add = Thm.declaration_attribute (map_fundef_congs o Thm.add_thm o safe_mk_meta_eq);
-val cong_del = Thm.declaration_attribute (map_fundef_congs o Thm.del_thm o safe_mk_meta_eq);
+val cong_add = Thm.declaration_attribute (map_function_congs o Thm.add_thm o safe_mk_meta_eq);
+val cong_del = Thm.declaration_attribute (map_function_congs o Thm.del_thm o safe_mk_meta_eq);
 
 
 type depgraph = int IntGraph.T
@@ -128,7 +128,7 @@
 
 fun mk_tree fvar h ctxt t =
     let 
-      val congs = get_fundef_congs ctxt
+      val congs = get_function_congs ctxt
       val congs_deps = map (fn c => (c, cong_deps c)) (congs @ default_congs) (* FIXME: Save in theory *)
 
       fun matchcall (a $ b) = if a = Free fvar then SOME b else NONE
--- a/src/HOL/Tools/Function/decompose.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/Function/decompose.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -33,8 +33,8 @@
                                       Const (@{const_name Set.empty}, fastype_of c1))
                        |> HOLogic.mk_Trueprop (* "C1 O C2 = {}" *)
 
-            val chain = case FundefLib.try_proof (cterm_of thy goal) chain_tac of
-                          FundefLib.Solved thm => SOME thm
+            val chain = case Function_Lib.try_proof (cterm_of thy goal) chain_tac of
+                          Function_Lib.Solved thm => SOME thm
                         | _ => NONE
           in
             Termination.note_chain c1 c2 chain D
@@ -62,7 +62,7 @@
    let
      val is = map (fn c => find_index (curry op aconv c) cs') cs
    in
-     CONVERSION (Conv.arg_conv (Conv.arg_conv (FundefLib.regroup_union_conv is))) i
+     CONVERSION (Conv.arg_conv (Conv.arg_conv (Function_Lib.regroup_union_conv is))) i
    end)
 
 
--- a/src/HOL/Tools/Function/descent.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/Function/descent.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -35,7 +35,7 @@
                (measures_of p) (measures_of q) D
       end
   in
-    cont (FundefCommon.PROFILE "deriving descents" (fold derive cs) D) i
+    cont (Function_Common.PROFILE "deriving descents" (fold derive cs) D) i
   end)
 
 val derive_diag = gen_descent true
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/fun.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -0,0 +1,178 @@
+(*  Title:      HOL/Tools/Function/fun.ML
+    Author:     Alexander Krauss, TU Muenchen
+
+Sequential mode for function definitions
+Command "fun" for fully automated function definitions
+*)
+
+signature FUNCTION_FUN =
+sig
+    val add_fun : Function_Common.function_config ->
+      (binding * typ option * mixfix) list -> (Attrib.binding * term) list ->
+      bool -> local_theory -> Proof.context
+    val add_fun_cmd : Function_Common.function_config ->
+      (binding * string option * mixfix) list -> (Attrib.binding * string) list ->
+      bool -> local_theory -> Proof.context
+
+    val setup : theory -> theory
+end
+
+structure Function_Fun : FUNCTION_FUN =
+struct
+
+open Function_Lib
+open Function_Common
+
+
+fun check_pats ctxt geq =
+    let 
+      fun err str = error (cat_lines ["Malformed definition:",
+                                      str ^ " not allowed in sequential mode.",
+                                      Syntax.string_of_term ctxt geq])
+      val thy = ProofContext.theory_of ctxt
+                
+      fun check_constr_pattern (Bound _) = ()
+        | check_constr_pattern t =
+          let
+            val (hd, args) = strip_comb t
+          in
+            (((case Datatype.info_of_constr thy (dest_Const hd) of
+                 SOME _ => ()
+               | NONE => err "Non-constructor pattern")
+              handle TERM ("dest_Const", _) => err "Non-constructor patterns");
+             map check_constr_pattern args; 
+             ())
+          end
+          
+      val (fname, qs, gs, args, rhs) = split_def ctxt geq 
+                                       
+      val _ = if not (null gs) then err "Conditional equations" else ()
+      val _ = map check_constr_pattern args
+                  
+                  (* just count occurrences to check linearity *)
+      val _ = if fold (fold_aterms (fn Bound _ => Integer.add 1 | _ => I)) args 0 > length qs
+              then err "Nonlinear patterns" else ()
+    in
+      ()
+    end
+    
+val by_pat_completeness_auto =
+    Proof.global_future_terminal_proof
+      (Method.Basic Pat_Completeness.pat_completeness,
+       SOME (Method.Source_i (Args.src (("HOL.auto", []), Position.none))))
+
+fun termination_by method int =
+    Function.termination_proof NONE
+    #> Proof.global_future_terminal_proof (Method.Basic method, NONE) int
+
+fun mk_catchall fixes arity_of =
+    let
+      fun mk_eqn ((fname, fT), _) =
+          let 
+            val n = arity_of fname
+            val (argTs, rT) = chop n (binder_types fT)
+                                   |> apsnd (fn Ts => Ts ---> body_type fT) 
+                              
+            val qs = map Free (Name.invent_list [] "a" n ~~ argTs)
+          in
+            HOLogic.mk_eq(list_comb (Free (fname, fT), qs),
+                          Const ("HOL.undefined", rT))
+              |> HOLogic.mk_Trueprop
+              |> fold_rev Logic.all qs
+          end
+    in
+      map mk_eqn fixes
+    end
+
+fun add_catchall ctxt fixes spec =
+  let val fqgars = map (split_def ctxt) spec
+      val arity_of = map (fn (fname,_,_,args,_) => (fname, length args)) fqgars
+                     |> AList.lookup (op =) #> the
+  in
+    spec @ mk_catchall fixes arity_of
+  end
+
+fun warn_if_redundant ctxt origs tss =
+    let
+        fun msg t = "Ignoring redundant equation: " ^ quote (Syntax.string_of_term ctxt t)
+                    
+        val (tss', _) = chop (length origs) tss
+        fun check (t, []) = (warning (msg t); [])
+          | check (t, s) = s
+    in
+        (map check (origs ~~ tss'); tss)
+    end
+
+
+fun sequential_preproc (config as FunctionConfig {sequential, ...}) ctxt fixes spec =
+      if sequential then
+        let
+          val (bnds, eqss) = split_list spec
+                            
+          val eqs = map the_single eqss
+                    
+          val feqs = eqs
+                      |> tap (check_defs ctxt fixes) (* Standard checks *)
+                      |> tap (map (check_pats ctxt))    (* More checks for sequential mode *)
+
+          val compleqs = add_catchall ctxt fixes feqs   (* Completion *)
+
+          val spliteqs = warn_if_redundant ctxt feqs
+                           (Function_Split.split_all_equations ctxt compleqs)
+
+          fun restore_spec thms =
+              bnds ~~ Library.take (length bnds, Library.unflat spliteqs thms)
+              
+          val spliteqs' = flat (Library.take (length bnds, spliteqs))
+          val fnames = map (fst o fst) fixes
+          val indices = map (fn eq => find_index (curry op = (fname_of eq)) fnames) spliteqs'
+
+          fun sort xs = partition_list (fn i => fn (j,_) => i = j) 0 (length fnames - 1) (indices ~~ xs)
+                                       |> map (map snd)
+
+
+          val bnds' = bnds @ replicate (length spliteqs - length bnds) Attrib.empty_binding
+
+          (* using theorem names for case name currently disabled *)
+          val case_names = map_index (fn (i, (_, es)) => mk_case_names i "" (length es)) 
+                                     (bnds' ~~ spliteqs)
+                           |> flat
+        in
+          (flat spliteqs, restore_spec, sort, case_names)
+        end
+      else
+        Function_Common.empty_preproc check_defs config ctxt fixes spec
+
+val setup =
+  Context.theory_map (Function_Common.set_preproc sequential_preproc)
+
+
+val fun_config = FunctionConfig { sequential=true, default="%x. undefined" (*FIXME dynamic scoping*), 
+  domintros=false, partials=false, tailrec=false }
+
+fun gen_fun add config fixes statements int lthy =
+  let val group = serial () in
+    lthy
+      |> LocalTheory.set_group group
+      |> add fixes statements config
+      |> by_pat_completeness_auto int
+      |> LocalTheory.restore
+      |> LocalTheory.set_group group
+      |> termination_by (Function_Common.get_termination_prover lthy) int
+  end;
+
+val add_fun = gen_fun Function.add_function
+val add_fun_cmd = gen_fun Function.add_function_cmd
+
+
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val _ =
+  OuterSyntax.local_theory' "fun" "define general recursive functions (short version)" K.thy_decl
+  (function_parser fun_config
+     >> (fn ((config, fixes), statements) => add_fun_cmd config fixes statements));
+
+end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/function.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -0,0 +1,227 @@
+(*  Title:      HOL/Tools/Function/fundef.ML
+    Author:     Alexander Krauss, TU Muenchen
+
+A package for general recursive function definitions.
+Isar commands.
+*)
+
+signature FUNCTION =
+sig
+    val add_function :  (binding * typ option * mixfix) list
+                       -> (Attrib.binding * term) list
+                       -> Function_Common.function_config
+                       -> local_theory
+                       -> Proof.state
+    val add_function_cmd :  (binding * string option * mixfix) list
+                      -> (Attrib.binding * string) list
+                      -> Function_Common.function_config
+                      -> local_theory
+                      -> Proof.state
+
+    val termination_proof : term option -> local_theory -> Proof.state
+    val termination_proof_cmd : string option -> local_theory -> Proof.state
+    val termination : term option -> local_theory -> Proof.state
+    val termination_cmd : string option -> local_theory -> Proof.state
+
+    val setup : theory -> theory
+    val get_congs : Proof.context -> thm list
+end
+
+
+structure Function : FUNCTION =
+struct
+
+open Function_Lib
+open Function_Common
+
+val simp_attribs = map (Attrib.internal o K)
+    [Simplifier.simp_add,
+     Code.add_default_eqn_attribute,
+     Nitpick_Simps.add,
+     Quickcheck_RecFun_Simps.add]
+
+val psimp_attribs = map (Attrib.internal o K)
+    [Simplifier.simp_add,
+     Nitpick_Psimps.add]
+
+fun note_theorem ((name, atts), ths) =
+  LocalTheory.note Thm.generatedK ((Binding.qualified_name name, atts), ths)
+
+fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_"
+
+fun add_simps fnames post sort extra_qualify label moreatts simps lthy =
+    let
+      val spec = post simps
+                   |> map (apfst (apsnd (fn ats => moreatts @ ats)))
+                   |> map (apfst (apfst extra_qualify))
+
+      val (saved_spec_simps, lthy) =
+        fold_map (LocalTheory.note Thm.generatedK) spec lthy
+
+      val saved_simps = maps snd saved_spec_simps
+      val simps_by_f = sort saved_simps
+
+      fun add_for_f fname simps =
+        note_theorem ((Long_Name.qualify fname label, []), simps) #> snd
+    in
+      (saved_simps,
+       fold2 add_for_f fnames simps_by_f lthy)
+    end
+
+fun gen_add_function is_external prep default_constraint fixspec eqns config lthy =
+    let
+      val constrn_fxs = map (fn (b, T, mx) => (b, SOME (the_default default_constraint T), mx))
+      val ((fixes0, spec0), ctxt') = prep (constrn_fxs fixspec) eqns lthy
+      val fixes = map (apfst (apfst Binding.name_of)) fixes0;
+      val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0;
+      val (eqs, post, sort_cont, cnames) = get_preproc lthy config ctxt' fixes spec
+
+      val defname = mk_defname fixes
+      val FunctionConfig {partials, ...} = config
+
+      val ((goalstate, cont), lthy) =
+          Function_Mutual.prepare_function_mutual config defname fixes eqs lthy
+
+      fun afterqed [[proof]] lthy =
+        let
+          val FunctionResult {fs, R, psimps, trsimps,  simple_pinducts, termination,
+                            domintros, cases, ...} =
+          cont (Thm.close_derivation proof)
+
+          val fnames = map (fst o fst) fixes
+          val qualify = Long_Name.qualify defname
+          val addsmps = add_simps fnames post sort_cont
+
+          val (((psimps', pinducts'), (_, [termination'])), lthy) =
+            lthy
+            |> addsmps (Binding.qualify false "partial") "psimps"
+                 psimp_attribs psimps
+            ||> fold_option (snd oo addsmps I "simps" simp_attribs) trsimps
+            ||>> note_theorem ((qualify "pinduct",
+                   [Attrib.internal (K (RuleCases.case_names cnames)),
+                    Attrib.internal (K (RuleCases.consumes 1)),
+                    Attrib.internal (K (Induct.induct_pred ""))]), simple_pinducts)
+            ||>> note_theorem ((qualify "termination", []), [termination])
+            ||> (snd o note_theorem ((qualify "cases",
+                   [Attrib.internal (K (RuleCases.case_names cnames))]), [cases]))
+            ||> fold_option (snd oo curry note_theorem (qualify "domintros", [])) domintros
+
+          val cdata = FunctionCtxData { add_simps=addsmps, case_names=cnames, psimps=psimps',
+                                      pinducts=snd pinducts', termination=termination',
+                                      fs=fs, R=R, defname=defname }
+          val _ =
+            if not is_external then ()
+            else Specification.print_consts lthy (K false) (map fst fixes)
+        in
+          lthy
+          |> LocalTheory.declaration (add_function_data o morph_function_data cdata)
+        end
+    in
+      lthy
+        |> is_external ? LocalTheory.set_group (serial ())
+        |> Proof.theorem_i NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]]
+        |> Proof.refine (Method.primitive_text (fn _ => goalstate)) |> Seq.hd
+    end
+
+val add_function = gen_add_function false Specification.check_spec (TypeInfer.anyT HOLogic.typeS)
+val add_function_cmd = gen_add_function true Specification.read_spec "_::type"
+
+fun gen_termination_proof prep_term raw_term_opt lthy =
+    let
+      val term_opt = Option.map (prep_term lthy) raw_term_opt
+      val data = the (case term_opt of
+                        SOME t => (import_function_data t lthy
+                          handle Option.Option =>
+                            error ("Not a function: " ^ quote (Syntax.string_of_term lthy t)))
+                      | NONE => (import_last_function lthy handle Option.Option => error "Not a function"))
+
+        val FunctionCtxData { termination, R, add_simps, case_names, psimps,
+                            pinducts, defname, ...} = data
+        val domT = domain_type (fastype_of R)
+        val goal = HOLogic.mk_Trueprop
+                     (HOLogic.mk_all ("x", domT, mk_acc domT R $ Free ("x", domT)))
+        fun afterqed [[totality]] lthy =
+          let
+            val totality = Thm.close_derivation totality
+            val remove_domain_condition =
+              full_simplify (HOL_basic_ss addsimps [totality, True_implies_equals])
+            val tsimps = map remove_domain_condition psimps
+            val tinduct = map remove_domain_condition pinducts
+            val qualify = Long_Name.qualify defname;
+          in
+            lthy
+            |> add_simps I "simps" simp_attribs tsimps |> snd
+            |> note_theorem
+               ((qualify "induct",
+                 [Attrib.internal (K (RuleCases.case_names case_names))]),
+                tinduct) |> snd
+          end
+    in
+      lthy
+      |> ProofContext.note_thmss ""
+         [((Binding.empty, [ContextRules.rule_del]), [([allI], [])])] |> snd
+      |> ProofContext.note_thmss ""
+         [((Binding.empty, [ContextRules.intro_bang (SOME 1)]), [([allI], [])])] |> snd
+      |> ProofContext.note_thmss ""
+         [((Binding.name "termination", [ContextRules.intro_bang (SOME 0)]),
+           [([Goal.norm_result termination], [])])] |> snd
+      |> Proof.theorem_i NONE afterqed [[(goal, [])]]
+    end
+
+val termination_proof = gen_termination_proof Syntax.check_term;
+val termination_proof_cmd = gen_termination_proof Syntax.read_term;
+
+fun termination term_opt lthy =
+  lthy
+  |> LocalTheory.set_group (serial ())
+  |> termination_proof term_opt;
+
+fun termination_cmd term_opt lthy =
+  lthy
+  |> LocalTheory.set_group (serial ())
+  |> termination_proof_cmd term_opt;
+
+
+(* Datatype hook to declare datatype congs as "function_congs" *)
+
+
+fun add_case_cong n thy =
+    Context.theory_map (Function_Ctx_Tree.map_function_congs (Thm.add_thm
+                          (Datatype.the_info thy n
+                           |> #case_cong
+                           |> safe_mk_meta_eq)))
+                       thy
+
+val setup_case_cong = Datatype.interpretation (K (fold add_case_cong))
+
+
+(* setup *)
+
+val setup =
+  Attrib.setup @{binding fundef_cong}
+    (Attrib.add_del Function_Ctx_Tree.cong_add Function_Ctx_Tree.cong_del)
+    "declaration of congruence rule for function definitions"
+  #> setup_case_cong
+  #> Function_Relation.setup
+  #> Function_Common.Termination_Simps.setup
+
+val get_congs = Function_Ctx_Tree.get_function_congs
+
+
+(* outer syntax *)
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val _ =
+  OuterSyntax.local_theory_to_proof "function" "define general recursive functions" K.thy_goal
+  (function_parser default_config
+     >> (fn ((config, fixes), statements) => add_function_cmd fixes statements config));
+
+val _ =
+  OuterSyntax.local_theory_to_proof "termination" "prove termination of a recursive function" K.thy_goal
+  (Scan.option P.term >> termination_cmd);
+
+end;
+
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/function_common.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -0,0 +1,348 @@
+(*  Title:      HOL/Tools/Function/fundef_common.ML
+    Author:     Alexander Krauss, TU Muenchen
+
+A package for general recursive function definitions. 
+Common definitions and other infrastructure.
+*)
+
+structure Function_Common =
+struct
+
+local open Function_Lib in
+
+(* Profiling *)
+val profile = Unsynchronized.ref false;
+
+fun PROFILE msg = if !profile then timeap_msg msg else I
+
+
+val acc_const_name = @{const_name accp}
+fun mk_acc domT R =
+    Const (acc_const_name, (domT --> domT --> HOLogic.boolT) --> domT --> HOLogic.boolT) $ R 
+
+val function_name = suffix "C"
+val graph_name = suffix "_graph"
+val rel_name = suffix "_rel"
+val dom_name = suffix "_dom"
+
+(* Termination rules *)
+
+structure TerminationRule = GenericDataFun
+(
+  type T = thm list
+  val empty = []
+  val extend = I
+  fun merge _ = Thm.merge_thms
+);
+
+val get_termination_rules = TerminationRule.get
+val store_termination_rule = TerminationRule.map o cons
+val apply_termination_rule = resolve_tac o get_termination_rules o Context.Proof
+
+
+(* Function definition result data *)
+
+datatype function_result =
+  FunctionResult of
+     {
+      fs: term list,
+      G: term,
+      R: term,
+
+      psimps : thm list, 
+      trsimps : thm list option, 
+
+      simple_pinducts : thm list, 
+      cases : thm,
+      termination : thm,
+      domintros : thm list option
+     }
+
+
+datatype function_context_data =
+  FunctionCtxData of
+     {
+      defname : string,
+
+      (* contains no logical entities: invariant under morphisms *)
+      add_simps : (binding -> binding) -> string -> Attrib.src list -> thm list 
+                  -> local_theory -> thm list * local_theory,
+      case_names : string list,
+
+      fs : term list,
+      R : term,
+      
+      psimps: thm list,
+      pinducts: thm list,
+      termination: thm
+     }
+
+fun morph_function_data (FunctionCtxData {add_simps, case_names, fs, R, 
+                                      psimps, pinducts, termination, defname}) phi =
+    let
+      val term = Morphism.term phi val thm = Morphism.thm phi val fact = Morphism.fact phi
+      val name = Binding.name_of o Morphism.binding phi o Binding.name
+    in
+      FunctionCtxData { add_simps = add_simps, case_names = case_names,
+                      fs = map term fs, R = term R, psimps = fact psimps, 
+                      pinducts = fact pinducts, termination = thm termination,
+                      defname = name defname }
+    end
+
+structure FunctionData = GenericDataFun
+(
+  type T = (term * function_context_data) Item_Net.T;
+  val empty = Item_Net.init
+    (op aconv o pairself fst : (term * function_context_data) * (term * function_context_data) -> bool)
+    fst;
+  val copy = I;
+  val extend = I;
+  fun merge _ (tab1, tab2) = Item_Net.merge (tab1, tab2)
+);
+
+val get_function = FunctionData.get o Context.Proof;
+
+
+(* Generally useful?? *)
+fun lift_morphism thy f = 
+    let 
+      val term = Drule.term_rule thy f
+    in
+      Morphism.thm_morphism f $> Morphism.term_morphism term 
+       $> Morphism.typ_morphism (Logic.type_map term)
+    end
+
+fun import_function_data t ctxt =
+    let
+      val thy = ProofContext.theory_of ctxt
+      val ct = cterm_of thy t
+      val inst_morph = lift_morphism thy o Thm.instantiate 
+
+      fun match (trm, data) = 
+          SOME (morph_function_data data (inst_morph (Thm.match (cterm_of thy trm, ct))))
+          handle Pattern.MATCH => NONE
+    in 
+      get_first match (Item_Net.retrieve (get_function ctxt) t)
+    end
+
+fun import_last_function ctxt =
+    case Item_Net.content (get_function ctxt) of
+      [] => NONE
+    | (t, data) :: _ =>
+      let 
+        val ([t'], ctxt') = Variable.import_terms true [t] ctxt
+      in
+        import_function_data t' ctxt'
+      end
+
+val all_function_data = Item_Net.content o get_function
+
+fun add_function_data (data as FunctionCtxData {fs, termination, ...}) =
+    FunctionData.map (fold (fn f => Item_Net.insert (f, data)) fs)
+    #> store_termination_rule termination
+
+
+(* Simp rules for termination proofs *)
+
+structure Termination_Simps = Named_Thms
+(
+  val name = "termination_simp" 
+  val description = "Simplification rule for termination proofs"
+);
+
+
+(* Default Termination Prover *)
+
+structure TerminationProver = GenericDataFun
+(
+  type T = Proof.context -> Proof.method
+  val empty = (fn _ => error "Termination prover not configured")
+  val extend = I
+  fun merge _ (a,b) = b (* FIXME *)
+);
+
+val set_termination_prover = TerminationProver.put
+val get_termination_prover = TerminationProver.get o Context.Proof
+
+
+(* Configuration management *)
+datatype function_opt 
+  = Sequential
+  | Default of string
+  | DomIntros
+  | No_Partials
+  | Tailrec
+
+datatype function_config
+  = FunctionConfig of
+   {
+    sequential: bool,
+    default: string,
+    domintros: bool,
+    partials: bool,
+    tailrec: bool
+   }
+
+fun apply_opt Sequential (FunctionConfig {sequential, default, domintros, partials, tailrec}) =
+    FunctionConfig {sequential=true, default=default, domintros=domintros, partials=partials, tailrec=tailrec}
+  | apply_opt (Default d) (FunctionConfig {sequential, default, domintros, partials, tailrec}) =
+    FunctionConfig {sequential=sequential, default=d, domintros=domintros, partials=partials, tailrec=tailrec}
+  | apply_opt DomIntros (FunctionConfig {sequential, default, domintros, partials, tailrec}) =
+    FunctionConfig {sequential=sequential, default=default, domintros=true, partials=partials, tailrec=tailrec}
+  | apply_opt Tailrec (FunctionConfig {sequential, default, domintros, partials, tailrec}) =
+    FunctionConfig {sequential=sequential, default=default, domintros=domintros, partials=partials, tailrec=true}
+  | apply_opt No_Partials (FunctionConfig {sequential, default, domintros, partials, tailrec}) =
+    FunctionConfig {sequential=sequential, default=default, domintros=domintros, partials=false, tailrec=true}
+
+val default_config =
+  FunctionConfig { sequential=false, default="%x. undefined" (*FIXME dynamic scoping*), 
+    domintros=false, partials=true, tailrec=false }
+
+
+(* Analyzing function equations *)
+
+fun split_def ctxt geq =
+    let
+      fun input_error msg = cat_lines [msg, Syntax.string_of_term ctxt geq]
+      val qs = Term.strip_qnt_vars "all" geq
+      val imp = Term.strip_qnt_body "all" geq
+      val (gs, eq) = Logic.strip_horn imp
+
+      val (f_args, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop eq)
+          handle TERM _ => error (input_error "Not an equation")
+
+      val (head, args) = strip_comb f_args
+
+      val fname = fst (dest_Free head)
+          handle TERM _ => error (input_error "Head symbol must not be a bound variable")
+    in
+      (fname, qs, gs, args, rhs)
+    end
+
+(* Check for all sorts of errors in the input *)
+fun check_defs ctxt fixes eqs =
+    let
+      val fnames = map (fst o fst) fixes
+                                
+      fun check geq = 
+          let
+            fun input_error msg = error (cat_lines [msg, Syntax.string_of_term ctxt geq])
+                                  
+            val fqgar as (fname, qs, gs, args, rhs) = split_def ctxt geq
+                                 
+            val _ = fname mem fnames 
+                    orelse input_error 
+                             ("Head symbol of left hand side must be " 
+                              ^ plural "" "one out of " fnames ^ commas_quote fnames)
+                                            
+            val _ = length args > 0 orelse input_error "Function has no arguments:"
+
+            fun add_bvs t is = add_loose_bnos (t, 0, is)
+            val rvs = (subtract (op =) (fold add_bvs args []) (add_bvs rhs []))
+                        |> map (fst o nth (rev qs))
+                      
+            val _ = null rvs orelse input_error 
+                        ("Variable" ^ plural " " "s " rvs ^ commas_quote rvs
+                         ^ " occur" ^ plural "s" "" rvs ^ " on right hand side only:")
+                                    
+            val _ = forall (not o Term.exists_subterm 
+                             (fn Free (n, _) => n mem fnames | _ => false)) (gs @ args)
+                    orelse input_error "Defined function may not occur in premises or arguments"
+
+            val freeargs = map (fn t => subst_bounds (rev (map Free qs), t)) args
+            val funvars = filter (fn q => exists (exists_subterm (fn (Free q') $ _ => q = q' | _ => false)) freeargs) qs
+            val _ = null funvars
+                    orelse (warning (cat_lines 
+                    ["Bound variable" ^ plural " " "s " funvars 
+                     ^ commas_quote (map fst funvars) ^  
+                     " occur" ^ plural "s" "" funvars ^ " in function position.",  
+                     "Misspelled constructor???"]); true)
+          in
+            (fname, length args)
+          end
+
+      val _ = AList.group (op =) (map check eqs)
+        |> map (fn (fname, ars) =>
+             length (distinct (op =) ars) = 1
+             orelse error ("Function " ^ quote fname ^
+                           " has different numbers of arguments in different equations"))
+
+      fun check_sorts ((fname, fT), _) =
+          Sorts.of_sort (Sign.classes_of (ProofContext.theory_of ctxt)) (fT, HOLogic.typeS)
+          orelse error (cat_lines 
+          ["Type of " ^ quote fname ^ " is not of sort " ^ quote "type" ^ ":",
+           setmp_CRITICAL show_sorts true (Syntax.string_of_typ ctxt) fT])
+
+      val _ = map check_sorts fixes
+    in
+      ()
+    end
+
+(* Preprocessors *)
+
+type fixes = ((string * typ) * mixfix) list
+type 'a spec = (Attrib.binding * 'a list) list
+type preproc = function_config -> Proof.context -> fixes -> term spec 
+               -> (term list * (thm list -> thm spec) * (thm list -> thm list list) * string list)
+
+val fname_of = fst o dest_Free o fst o strip_comb o fst 
+ o HOLogic.dest_eq o HOLogic.dest_Trueprop o Logic.strip_imp_concl o snd o dest_all_all
+
+fun mk_case_names i "" k = mk_case_names i (string_of_int (i + 1)) k
+  | mk_case_names _ n 0 = []
+  | mk_case_names _ n 1 = [n]
+  | mk_case_names _ n k = map (fn i => n ^ "_" ^ string_of_int i) (1 upto k)
+
+fun empty_preproc check _ ctxt fixes spec =
+    let 
+      val (bnds, tss) = split_list spec
+      val ts = flat tss
+      val _ = check ctxt fixes ts
+      val fnames = map (fst o fst) fixes
+      val indices = map (fn eq => find_index (curry op = (fname_of eq)) fnames) ts
+
+      fun sort xs = partition_list (fn i => fn (j,_) => i = j) 0 (length fnames - 1) 
+                                   (indices ~~ xs)
+                        |> map (map snd)
+
+      (* using theorem names for case name currently disabled *)
+      val cnames = map_index (fn (i, _) => mk_case_names i "" 1) bnds |> flat
+    in
+      (ts, curry op ~~ bnds o Library.unflat tss, sort, cnames)
+    end
+
+structure Preprocessor = GenericDataFun
+(
+  type T = preproc
+  val empty : T = empty_preproc check_defs
+  val extend = I
+  fun merge _ (a, _) = a
+);
+
+val get_preproc = Preprocessor.get o Context.Proof
+val set_preproc = Preprocessor.map o K
+
+
+
+local 
+  structure P = OuterParse and K = OuterKeyword
+
+  val option_parser = 
+      P.group "option" ((P.reserved "sequential" >> K Sequential)
+                    || ((P.reserved "default" |-- P.term) >> Default)
+                    || (P.reserved "domintros" >> K DomIntros)
+                    || (P.reserved "no_partials" >> K No_Partials)
+                    || (P.reserved "tailrec" >> K Tailrec))
+
+  fun config_parser default = 
+      (Scan.optional (P.$$$ "(" |-- P.!!! (P.list1 option_parser) --| P.$$$ ")") [])
+        >> (fn opts => fold apply_opt opts default)
+in
+  fun function_parser default_cfg = 
+      config_parser default_cfg -- P.fixes -- SpecParse.where_alt_specs
+end
+
+
+end
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/function_core.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -0,0 +1,956 @@
+(*  Title:      HOL/Tools/Function/function_core.ML
+    Author:     Alexander Krauss, TU Muenchen
+
+A package for general recursive function definitions:
+Main functionality.
+*)
+
+signature FUNCTION_CORE =
+sig
+    val trace: bool Unsynchronized.ref
+
+    val prepare_function : Function_Common.function_config
+                         -> string (* defname *)
+                         -> ((bstring * typ) * mixfix) list (* defined symbol *)
+                         -> ((bstring * typ) list * term list * term * term) list (* specification *)
+                         -> local_theory
+
+                         -> (term   (* f *)
+                             * thm  (* goalstate *)
+                             * (thm -> Function_Common.function_result) (* continuation *)
+                            ) * local_theory
+
+end
+
+structure Function_Core : FUNCTION_CORE =
+struct
+
+val trace = Unsynchronized.ref false;
+fun trace_msg msg = if ! trace then tracing (msg ()) else ();
+
+val boolT = HOLogic.boolT
+val mk_eq = HOLogic.mk_eq
+
+open Function_Lib
+open Function_Common
+
+datatype globals =
+   Globals of {
+         fvar: term,
+         domT: typ,
+         ranT: typ,
+         h: term,
+         y: term,
+         x: term,
+         z: term,
+         a: term,
+         P: term,
+         D: term,
+         Pbool:term
+}
+
+
+datatype rec_call_info =
+  RCInfo of
+  {
+   RIvs: (string * typ) list,  (* Call context: fixes and assumes *)
+   CCas: thm list,
+   rcarg: term,                 (* The recursive argument *)
+
+   llRI: thm,
+   h_assum: term
+  }
+
+
+datatype clause_context =
+  ClauseContext of
+  {
+    ctxt : Proof.context,
+
+    qs : term list,
+    gs : term list,
+    lhs: term,
+    rhs: term,
+
+    cqs: cterm list,
+    ags: thm list,
+    case_hyp : thm
+  }
+
+
+fun transfer_clause_ctx thy (ClauseContext { ctxt, qs, gs, lhs, rhs, cqs, ags, case_hyp }) =
+    ClauseContext { ctxt = ProofContext.transfer thy ctxt,
+                    qs = qs, gs = gs, lhs = lhs, rhs = rhs, cqs = cqs, ags = ags, case_hyp = case_hyp }
+
+
+datatype clause_info =
+  ClauseInfo of
+     {
+      no: int,
+      qglr : ((string * typ) list * term list * term * term),
+      cdata : clause_context,
+
+      tree: Function_Ctx_Tree.ctx_tree,
+      lGI: thm,
+      RCs: rec_call_info list
+     }
+
+
+(* Theory dependencies. *)
+val Pair_inject = @{thm Product_Type.Pair_inject};
+
+val acc_induct_rule = @{thm accp_induct_rule};
+
+val ex1_implies_ex = @{thm FunDef.fundef_ex1_existence};
+val ex1_implies_un = @{thm FunDef.fundef_ex1_uniqueness};
+val ex1_implies_iff = @{thm FunDef.fundef_ex1_iff};
+
+val acc_downward = @{thm accp_downward};
+val accI = @{thm accp.accI};
+val case_split = @{thm HOL.case_split};
+val fundef_default_value = @{thm FunDef.fundef_default_value};
+val not_acc_down = @{thm not_accp_down};
+
+
+
+fun find_calls tree =
+    let
+      fun add_Ri (fixes,assumes) (_ $ arg) _ (_, xs) = ([], (fixes, assumes, arg) :: xs)
+        | add_Ri _ _ _ _ = raise Match
+    in
+      rev (Function_Ctx_Tree.traverse_tree add_Ri tree [])
+    end
+
+
+(** building proof obligations *)
+
+fun mk_compat_proof_obligations domT ranT fvar f glrs =
+    let
+      fun mk_impl ((qs, gs, lhs, rhs),(qs', gs', lhs', rhs')) =
+          let
+            val shift = incr_boundvars (length qs')
+          in
+            Logic.mk_implies
+              (HOLogic.mk_Trueprop (HOLogic.eq_const domT $ shift lhs $ lhs'),
+                HOLogic.mk_Trueprop (HOLogic.eq_const ranT $ shift rhs $ rhs'))
+              |> fold_rev (curry Logic.mk_implies) (map shift gs @ gs')
+              |> fold_rev (fn (n,T) => fn b => Term.all T $ Abs(n,T,b)) (qs @ qs')
+              |> curry abstract_over fvar
+              |> curry subst_bound f
+          end
+    in
+      map mk_impl (unordered_pairs glrs)
+    end
+
+
+fun mk_completeness (Globals {x, Pbool, ...}) clauses qglrs =
+    let
+        fun mk_case (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) =
+            HOLogic.mk_Trueprop Pbool
+                     |> curry Logic.mk_implies (HOLogic.mk_Trueprop (mk_eq (x, lhs)))
+                     |> fold_rev (curry Logic.mk_implies) gs
+                     |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
+    in
+        HOLogic.mk_Trueprop Pbool
+                 |> fold_rev (curry Logic.mk_implies o mk_case) (clauses ~~ qglrs)
+                 |> mk_forall_rename ("x", x)
+                 |> mk_forall_rename ("P", Pbool)
+    end
+
+(** making a context with it's own local bindings **)
+
+fun mk_clause_context x ctxt (pre_qs,pre_gs,pre_lhs,pre_rhs) =
+    let
+      val (qs, ctxt') = Variable.variant_fixes (map fst pre_qs) ctxt
+                                           |>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs
+
+      val thy = ProofContext.theory_of ctxt'
+
+      fun inst t = subst_bounds (rev qs, t)
+      val gs = map inst pre_gs
+      val lhs = inst pre_lhs
+      val rhs = inst pre_rhs
+
+      val cqs = map (cterm_of thy) qs
+      val ags = map (assume o cterm_of thy) gs
+
+      val case_hyp = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (x, lhs))))
+    in
+      ClauseContext { ctxt = ctxt', qs = qs, gs = gs, lhs = lhs, rhs = rhs,
+                      cqs = cqs, ags = ags, case_hyp = case_hyp }
+    end
+
+
+(* lowlevel term function. FIXME: remove *)
+fun abstract_over_list vs body =
+  let
+    fun abs lev v tm =
+      if v aconv tm then Bound lev
+      else
+        (case tm of
+          Abs (a, T, t) => Abs (a, T, abs (lev + 1) v t)
+        | t $ u => abs lev v t $ abs lev v u
+        | t => t);
+  in
+    fold_index (fn (i, v) => fn t => abs i v t) vs body
+  end
+
+
+
+fun mk_clause_info globals G f no cdata qglr tree RCs GIntro_thm RIntro_thms =
+    let
+        val Globals {h, fvar, x, ...} = globals
+
+        val ClauseContext { ctxt, qs, cqs, ags, ... } = cdata
+        val cert = Thm.cterm_of (ProofContext.theory_of ctxt)
+
+        (* Instantiate the GIntro thm with "f" and import into the clause context. *)
+        val lGI = GIntro_thm
+                    |> forall_elim (cert f)
+                    |> fold forall_elim cqs
+                    |> fold Thm.elim_implies ags
+
+        fun mk_call_info (rcfix, rcassm, rcarg) RI =
+            let
+                val llRI = RI
+                             |> fold forall_elim cqs
+                             |> fold (forall_elim o cert o Free) rcfix
+                             |> fold Thm.elim_implies ags
+                             |> fold Thm.elim_implies rcassm
+
+                val h_assum =
+                    HOLogic.mk_Trueprop (G $ rcarg $ (h $ rcarg))
+                              |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
+                              |> fold_rev (Logic.all o Free) rcfix
+                              |> Pattern.rewrite_term (ProofContext.theory_of ctxt) [(f, h)] []
+                              |> abstract_over_list (rev qs)
+            in
+                RCInfo {RIvs=rcfix, rcarg=rcarg, CCas=rcassm, llRI=llRI, h_assum=h_assum}
+            end
+
+        val RC_infos = map2 mk_call_info RCs RIntro_thms
+    in
+        ClauseInfo
+            {
+             no=no,
+             cdata=cdata,
+             qglr=qglr,
+
+             lGI=lGI,
+             RCs=RC_infos,
+             tree=tree
+            }
+    end
+
+
+
+
+
+
+
+(* replace this by a table later*)
+fun store_compat_thms 0 thms = []
+  | store_compat_thms n thms =
+    let
+        val (thms1, thms2) = chop n thms
+    in
+        (thms1 :: store_compat_thms (n - 1) thms2)
+    end
+
+(* expects i <= j *)
+fun lookup_compat_thm i j cts =
+    nth (nth cts (i - 1)) (j - i)
+
+(* Returns "Gsi, Gsj, lhs_i = lhs_j |-- rhs_j_f = rhs_i_f" *)
+(* if j < i, then turn around *)
+fun get_compat_thm thy cts i j ctxi ctxj =
+    let
+      val ClauseContext {cqs=cqsi,ags=agsi,lhs=lhsi,...} = ctxi
+      val ClauseContext {cqs=cqsj,ags=agsj,lhs=lhsj,...} = ctxj
+
+      val lhsi_eq_lhsj = cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj)))
+    in if j < i then
+         let
+           val compat = lookup_compat_thm j i cts
+         in
+           compat         (* "!!qj qi. Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
+                |> fold forall_elim (cqsj @ cqsi) (* "Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
+                |> fold Thm.elim_implies agsj
+                |> fold Thm.elim_implies agsi
+                |> Thm.elim_implies ((assume lhsi_eq_lhsj) RS sym) (* "Gsj, Gsi, lhsi = lhsj |-- rhsj = rhsi" *)
+         end
+       else
+         let
+           val compat = lookup_compat_thm i j cts
+         in
+               compat        (* "!!qi qj. Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
+                 |> fold forall_elim (cqsi @ cqsj) (* "Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
+                 |> fold Thm.elim_implies agsi
+                 |> fold Thm.elim_implies agsj
+                 |> Thm.elim_implies (assume lhsi_eq_lhsj)
+                 |> (fn thm => thm RS sym) (* "Gsi, Gsj, lhsi = lhsj |-- rhsj = rhsi" *)
+         end
+    end
+
+
+
+
+(* Generates the replacement lemma in fully quantified form. *)
+fun mk_replacement_lemma thy h ih_elim clause =
+    let
+        val ClauseInfo {cdata=ClauseContext {qs, lhs, rhs, cqs, ags, case_hyp, ...}, RCs, tree, ...} = clause
+        local open Conv in
+        val ih_conv = arg1_conv o arg_conv o arg_conv
+        end
+
+        val ih_elim_case = Conv.fconv_rule (ih_conv (K (case_hyp RS eq_reflection))) ih_elim
+
+        val Ris = map (fn RCInfo {llRI, ...} => llRI) RCs
+        val h_assums = map (fn RCInfo {h_assum, ...} => assume (cterm_of thy (subst_bounds (rev qs, h_assum)))) RCs
+
+        val (eql, _) = Function_Ctx_Tree.rewrite_by_tree thy h ih_elim_case (Ris ~~ h_assums) tree
+
+        val replace_lemma = (eql RS meta_eq_to_obj_eq)
+                                |> implies_intr (cprop_of case_hyp)
+                                |> fold_rev (implies_intr o cprop_of) h_assums
+                                |> fold_rev (implies_intr o cprop_of) ags
+                                |> fold_rev forall_intr cqs
+                                |> Thm.close_derivation
+    in
+      replace_lemma
+    end
+
+
+fun mk_uniqueness_clause thy globals f compat_store clausei clausej RLj =
+    let
+        val Globals {h, y, x, fvar, ...} = globals
+        val ClauseInfo {no=i, cdata=cctxi as ClauseContext {ctxt=ctxti, lhs=lhsi, case_hyp, ...}, ...} = clausei
+        val ClauseInfo {no=j, qglr=cdescj, RCs=RCsj, ...} = clausej
+
+        val cctxj as ClauseContext {ags = agsj', lhs = lhsj', rhs = rhsj', qs = qsj', cqs = cqsj', ...}
+            = mk_clause_context x ctxti cdescj
+
+        val rhsj'h = Pattern.rewrite_term thy [(fvar,h)] [] rhsj'
+        val compat = get_compat_thm thy compat_store i j cctxi cctxj
+        val Ghsj' = map (fn RCInfo {h_assum, ...} => assume (cterm_of thy (subst_bounds (rev qsj', h_assum)))) RCsj
+
+        val RLj_import =
+            RLj |> fold forall_elim cqsj'
+                |> fold Thm.elim_implies agsj'
+                |> fold Thm.elim_implies Ghsj'
+
+        val y_eq_rhsj'h = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (y, rhsj'h))))
+        val lhsi_eq_lhsj' = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj')))) (* lhs_i = lhs_j' |-- lhs_i = lhs_j' *)
+    in
+        (trans OF [case_hyp, lhsi_eq_lhsj']) (* lhs_i = lhs_j' |-- x = lhs_j' *)
+        |> implies_elim RLj_import (* Rj1' ... Rjk', lhs_i = lhs_j' |-- rhs_j'_h = rhs_j'_f *)
+        |> (fn it => trans OF [it, compat]) (* lhs_i = lhs_j', Gj', Rj1' ... Rjk' |-- rhs_j'_h = rhs_i_f *)
+        |> (fn it => trans OF [y_eq_rhsj'h, it]) (* lhs_i = lhs_j', Gj', Rj1' ... Rjk', y = rhs_j_h' |-- y = rhs_i_f *)
+        |> fold_rev (implies_intr o cprop_of) Ghsj'
+        |> fold_rev (implies_intr o cprop_of) agsj' (* lhs_i = lhs_j' , y = rhs_j_h' |-- Gj', Rj1'...Rjk' ==> y = rhs_i_f *)
+        |> implies_intr (cprop_of y_eq_rhsj'h)
+        |> implies_intr (cprop_of lhsi_eq_lhsj')
+        |> fold_rev forall_intr (cterm_of thy h :: cqsj')
+    end
+
+
+
+fun mk_uniqueness_case ctxt thy globals G f ihyp ih_intro G_cases compat_store clauses rep_lemmas clausei =
+    let
+        val Globals {x, y, ranT, fvar, ...} = globals
+        val ClauseInfo {cdata = ClauseContext {lhs, rhs, qs, cqs, ags, case_hyp, ...}, lGI, RCs, ...} = clausei
+        val rhsC = Pattern.rewrite_term thy [(fvar, f)] [] rhs
+
+        val ih_intro_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ih_intro
+
+        fun prep_RC (RCInfo {llRI, RIvs, CCas, ...}) = (llRI RS ih_intro_case)
+                                                            |> fold_rev (implies_intr o cprop_of) CCas
+                                                            |> fold_rev (forall_intr o cterm_of thy o Free) RIvs
+
+        val existence = fold (curry op COMP o prep_RC) RCs lGI
+
+        val P = cterm_of thy (mk_eq (y, rhsC))
+        val G_lhs_y = assume (cterm_of thy (HOLogic.mk_Trueprop (G $ lhs $ y)))
+
+        val unique_clauses = map2 (mk_uniqueness_clause thy globals f compat_store clausei) clauses rep_lemmas
+
+        val uniqueness = G_cases
+                           |> forall_elim (cterm_of thy lhs)
+                           |> forall_elim (cterm_of thy y)
+                           |> forall_elim P
+                           |> Thm.elim_implies G_lhs_y
+                           |> fold Thm.elim_implies unique_clauses
+                           |> implies_intr (cprop_of G_lhs_y)
+                           |> forall_intr (cterm_of thy y)
+
+        val P2 = cterm_of thy (lambda y (G $ lhs $ y)) (* P2 y := (lhs, y): G *)
+
+        val exactly_one =
+            ex1I |> instantiate' [SOME (ctyp_of thy ranT)] [SOME P2, SOME (cterm_of thy rhsC)]
+                 |> curry (op COMP) existence
+                 |> curry (op COMP) uniqueness
+                 |> simplify (HOL_basic_ss addsimps [case_hyp RS sym])
+                 |> implies_intr (cprop_of case_hyp)
+                 |> fold_rev (implies_intr o cprop_of) ags
+                 |> fold_rev forall_intr cqs
+
+        val function_value =
+            existence
+              |> implies_intr ihyp
+              |> implies_intr (cprop_of case_hyp)
+              |> forall_intr (cterm_of thy x)
+              |> forall_elim (cterm_of thy lhs)
+              |> curry (op RS) refl
+    in
+        (exactly_one, function_value)
+    end
+
+
+
+
+fun prove_stuff ctxt globals G f R clauses complete compat compat_store G_elim f_def =
+    let
+        val Globals {h, domT, ranT, x, ...} = globals
+        val thy = ProofContext.theory_of ctxt
+
+        (* Inductive Hypothesis: !!z. (z,x):R ==> EX!y. (z,y):G *)
+        val ihyp = Term.all domT $ Abs ("z", domT,
+                                   Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
+                                     HOLogic.mk_Trueprop (Const ("Ex1", (ranT --> boolT) --> boolT) $
+                                                             Abs ("y", ranT, G $ Bound 1 $ Bound 0))))
+                       |> cterm_of thy
+
+        val ihyp_thm = assume ihyp |> Thm.forall_elim_vars 0
+        val ih_intro = ihyp_thm RS (f_def RS ex1_implies_ex)
+        val ih_elim = ihyp_thm RS (f_def RS ex1_implies_un)
+                        |> instantiate' [] [NONE, SOME (cterm_of thy h)]
+
+        val _ = trace_msg (K "Proving Replacement lemmas...")
+        val repLemmas = map (mk_replacement_lemma thy h ih_elim) clauses
+
+        val _ = trace_msg (K "Proving cases for unique existence...")
+        val (ex1s, values) =
+            split_list (map (mk_uniqueness_case ctxt thy globals G f ihyp ih_intro G_elim compat_store clauses repLemmas) clauses)
+
+        val _ = trace_msg (K "Proving: Graph is a function")
+        val graph_is_function = complete
+                                  |> Thm.forall_elim_vars 0
+                                  |> fold (curry op COMP) ex1s
+                                  |> implies_intr (ihyp)
+                                  |> implies_intr (cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ x)))
+                                  |> forall_intr (cterm_of thy x)
+                                  |> (fn it => Drule.compose_single (it, 2, acc_induct_rule)) (* "EX! y. (?x,y):G" *)
+                                  |> (fn it => fold (forall_intr o cterm_of thy o Var) (Term.add_vars (prop_of it) []) it)
+
+        val goalstate =  Conjunction.intr graph_is_function complete
+                          |> Thm.close_derivation
+                          |> Goal.protect
+                          |> fold_rev (implies_intr o cprop_of) compat
+                          |> implies_intr (cprop_of complete)
+    in
+      (goalstate, values)
+    end
+
+
+fun define_graph Gname fvar domT ranT clauses RCss lthy =
+    let
+      val GT = domT --> ranT --> boolT
+      val Gvar = Free (the_single (Variable.variant_frees lthy [] [(Gname, GT)]))
+
+      fun mk_GIntro (ClauseContext {qs, gs, lhs, rhs, ...}) RCs =
+          let
+            fun mk_h_assm (rcfix, rcassm, rcarg) =
+                HOLogic.mk_Trueprop (Gvar $ rcarg $ (fvar $ rcarg))
+                          |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
+                          |> fold_rev (Logic.all o Free) rcfix
+          in
+            HOLogic.mk_Trueprop (Gvar $ lhs $ rhs)
+                      |> fold_rev (curry Logic.mk_implies o mk_h_assm) RCs
+                      |> fold_rev (curry Logic.mk_implies) gs
+                      |> fold_rev Logic.all (fvar :: qs)
+          end
+
+      val G_intros = map2 mk_GIntro clauses RCss
+
+      val (GIntro_thms, (G, G_elim, G_induct, lthy)) =
+          Function_Inductive_Wrap.inductive_def G_intros ((dest_Free Gvar, NoSyn), lthy)
+    in
+      ((G, GIntro_thms, G_elim, G_induct), lthy)
+    end
+
+
+
+fun define_function fdefname (fname, mixfix) domT ranT G default lthy =
+    let
+      val f_def =
+          Abs ("x", domT, Const (@{const_name FunDef.THE_default}, ranT --> (ranT --> boolT) --> ranT) $ (default $ Bound 0) $
+                                Abs ("y", ranT, G $ Bound 1 $ Bound 0))
+              |> Syntax.check_term lthy
+
+      val ((f, (_, f_defthm)), lthy) =
+        LocalTheory.define Thm.internalK ((Binding.name (function_name fname), mixfix), ((Binding.name fdefname, []), f_def)) lthy
+    in
+      ((f, f_defthm), lthy)
+    end
+
+
+fun define_recursion_relation Rname domT ranT fvar f qglrs clauses RCss lthy =
+    let
+
+      val RT = domT --> domT --> boolT
+      val Rvar = Free (the_single (Variable.variant_frees lthy [] [(Rname, RT)]))
+
+      fun mk_RIntro (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) (rcfix, rcassm, rcarg) =
+          HOLogic.mk_Trueprop (Rvar $ rcarg $ lhs)
+                    |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
+                    |> fold_rev (curry Logic.mk_implies) gs
+                    |> fold_rev (Logic.all o Free) rcfix
+                    |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
+                    (* "!!qs xs. CS ==> G => (r, lhs) : R" *)
+
+      val R_intross = map2 (map o mk_RIntro) (clauses ~~ qglrs) RCss
+
+      val (RIntro_thmss, (R, R_elim, _, lthy)) =
+          fold_burrow Function_Inductive_Wrap.inductive_def R_intross ((dest_Free Rvar, NoSyn), lthy)
+    in
+      ((R, RIntro_thmss, R_elim), lthy)
+    end
+
+
+fun fix_globals domT ranT fvar ctxt =
+    let
+      val ([h, y, x, z, a, D, P, Pbool],ctxt') =
+          Variable.variant_fixes ["h_fd", "y_fd", "x_fd", "z_fd", "a_fd", "D_fd", "P_fd", "Pb_fd"] ctxt
+    in
+      (Globals {h = Free (h, domT --> ranT),
+                y = Free (y, ranT),
+                x = Free (x, domT),
+                z = Free (z, domT),
+                a = Free (a, domT),
+                D = Free (D, domT --> boolT),
+                P = Free (P, domT --> boolT),
+                Pbool = Free (Pbool, boolT),
+                fvar = fvar,
+                domT = domT,
+                ranT = ranT
+               },
+       ctxt')
+    end
+
+
+
+fun inst_RC thy fvar f (rcfix, rcassm, rcarg) =
+    let
+      fun inst_term t = subst_bound(f, abstract_over (fvar, t))
+    in
+      (rcfix, map (assume o cterm_of thy o inst_term o prop_of) rcassm, inst_term rcarg)
+    end
+
+
+
+(**********************************************************
+ *                   PROVING THE RULES
+ **********************************************************)
+
+fun mk_psimps thy globals R clauses valthms f_iff graph_is_function =
+    let
+      val Globals {domT, z, ...} = globals
+
+      fun mk_psimp (ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {cqs, lhs, ags, ...}, ...}) valthm =
+          let
+            val lhs_acc = cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ lhs)) (* "acc R lhs" *)
+            val z_smaller = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ lhs)) (* "R z lhs" *)
+          in
+            ((assume z_smaller) RS ((assume lhs_acc) RS acc_downward))
+              |> (fn it => it COMP graph_is_function)
+              |> implies_intr z_smaller
+              |> forall_intr (cterm_of thy z)
+              |> (fn it => it COMP valthm)
+              |> implies_intr lhs_acc
+              |> asm_simplify (HOL_basic_ss addsimps [f_iff])
+              |> fold_rev (implies_intr o cprop_of) ags
+              |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
+          end
+    in
+      map2 mk_psimp clauses valthms
+    end
+
+
+(** Induction rule **)
+
+
+val acc_subset_induct = @{thm Orderings.predicate1I} RS @{thm accp_subset_induct}
+
+
+fun mk_partial_induct_rule thy globals R complete_thm clauses =
+    let
+      val Globals {domT, x, z, a, P, D, ...} = globals
+      val acc_R = mk_acc domT R
+
+      val x_D = assume (cterm_of thy (HOLogic.mk_Trueprop (D $ x)))
+      val a_D = cterm_of thy (HOLogic.mk_Trueprop (D $ a))
+
+      val D_subset = cterm_of thy (Logic.all x
+        (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x), HOLogic.mk_Trueprop (acc_R $ x))))
+
+      val D_dcl = (* "!!x z. [| x: D; (z,x):R |] ==> z:D" *)
+                    Logic.all x
+                    (Logic.all z (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x),
+                                                    Logic.mk_implies (HOLogic.mk_Trueprop (R $ z $ x),
+                                                                      HOLogic.mk_Trueprop (D $ z)))))
+                    |> cterm_of thy
+
+
+  (* Inductive Hypothesis: !!z. (z,x):R ==> P z *)
+      val ihyp = Term.all domT $ Abs ("z", domT,
+               Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
+                 HOLogic.mk_Trueprop (P $ Bound 0)))
+           |> cterm_of thy
+
+      val aihyp = assume ihyp
+
+  fun prove_case clause =
+      let
+    val ClauseInfo {cdata = ClauseContext {ctxt, qs, cqs, ags, gs, lhs, case_hyp, ...}, RCs,
+                    qglr = (oqs, _, _, _), ...} = clause
+
+    val case_hyp_conv = K (case_hyp RS eq_reflection)
+    local open Conv in
+    val lhs_D = fconv_rule (arg_conv (arg_conv (case_hyp_conv))) x_D
+    val sih = fconv_rule (More_Conv.binder_conv (K (arg1_conv (arg_conv (arg_conv case_hyp_conv)))) ctxt) aihyp
+    end
+
+    fun mk_Prec (RCInfo {llRI, RIvs, CCas, rcarg, ...}) =
+        sih |> forall_elim (cterm_of thy rcarg)
+            |> Thm.elim_implies llRI
+            |> fold_rev (implies_intr o cprop_of) CCas
+            |> fold_rev (forall_intr o cterm_of thy o Free) RIvs
+
+    val P_recs = map mk_Prec RCs   (*  [P rec1, P rec2, ... ]  *)
+
+    val step = HOLogic.mk_Trueprop (P $ lhs)
+            |> fold_rev (curry Logic.mk_implies o prop_of) P_recs
+            |> fold_rev (curry Logic.mk_implies) gs
+            |> curry Logic.mk_implies (HOLogic.mk_Trueprop (D $ lhs))
+            |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
+            |> cterm_of thy
+
+    val P_lhs = assume step
+           |> fold forall_elim cqs
+           |> Thm.elim_implies lhs_D
+           |> fold Thm.elim_implies ags
+           |> fold Thm.elim_implies P_recs
+
+    val res = cterm_of thy (HOLogic.mk_Trueprop (P $ x))
+           |> Conv.arg_conv (Conv.arg_conv case_hyp_conv)
+           |> symmetric (* P lhs == P x *)
+           |> (fn eql => equal_elim eql P_lhs) (* "P x" *)
+           |> implies_intr (cprop_of case_hyp)
+           |> fold_rev (implies_intr o cprop_of) ags
+           |> fold_rev forall_intr cqs
+      in
+        (res, step)
+      end
+
+  val (cases, steps) = split_list (map prove_case clauses)
+
+  val istep = complete_thm
+                |> Thm.forall_elim_vars 0
+                |> fold (curry op COMP) cases (*  P x  *)
+                |> implies_intr ihyp
+                |> implies_intr (cprop_of x_D)
+                |> forall_intr (cterm_of thy x)
+
+  val subset_induct_rule =
+      acc_subset_induct
+        |> (curry op COMP) (assume D_subset)
+        |> (curry op COMP) (assume D_dcl)
+        |> (curry op COMP) (assume a_D)
+        |> (curry op COMP) istep
+        |> fold_rev implies_intr steps
+        |> implies_intr a_D
+        |> implies_intr D_dcl
+        |> implies_intr D_subset
+
+  val subset_induct_all = fold_rev (forall_intr o cterm_of thy) [P, a, D] subset_induct_rule
+
+  val simple_induct_rule =
+      subset_induct_rule
+        |> forall_intr (cterm_of thy D)
+        |> forall_elim (cterm_of thy acc_R)
+        |> assume_tac 1 |> Seq.hd
+        |> (curry op COMP) (acc_downward
+                              |> (instantiate' [SOME (ctyp_of thy domT)]
+                                               (map (SOME o cterm_of thy) [R, x, z]))
+                              |> forall_intr (cterm_of thy z)
+                              |> forall_intr (cterm_of thy x))
+        |> forall_intr (cterm_of thy a)
+        |> forall_intr (cterm_of thy P)
+    in
+      simple_induct_rule
+    end
+
+
+
+(* FIXME: This should probably use fixed goals, to be more reliable and faster *)
+fun mk_domain_intro ctxt (Globals {domT, ...}) R R_cases clause =
+    let
+      val thy = ProofContext.theory_of ctxt
+      val ClauseInfo {cdata = ClauseContext {qs, gs, lhs, rhs, cqs, ...},
+                      qglr = (oqs, _, _, _), ...} = clause
+      val goal = HOLogic.mk_Trueprop (mk_acc domT R $ lhs)
+                          |> fold_rev (curry Logic.mk_implies) gs
+                          |> cterm_of thy
+    in
+      Goal.init goal
+      |> (SINGLE (resolve_tac [accI] 1)) |> the
+      |> (SINGLE (eresolve_tac [Thm.forall_elim_vars 0 R_cases] 1))  |> the
+      |> (SINGLE (auto_tac (clasimpset_of ctxt))) |> the
+      |> Goal.conclude
+      |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
+    end
+
+
+
+(** Termination rule **)
+
+val wf_induct_rule = @{thm Wellfounded.wfP_induct_rule};
+val wf_in_rel = @{thm FunDef.wf_in_rel};
+val in_rel_def = @{thm FunDef.in_rel_def};
+
+fun mk_nest_term_case thy globals R' ihyp clause =
+    let
+      val Globals {x, z, ...} = globals
+      val ClauseInfo {cdata = ClauseContext {qs,cqs,ags,lhs,rhs,case_hyp,...},tree,
+                      qglr=(oqs, _, _, _), ...} = clause
+
+      val ih_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ihyp
+
+      fun step (fixes, assumes) (_ $ arg) u (sub,(hyps,thms)) =
+          let
+            val used = map (fn (ctx,thm) => Function_Ctx_Tree.export_thm thy ctx thm) (u @ sub)
+
+            val hyp = HOLogic.mk_Trueprop (R' $ arg $ lhs)
+                      |> fold_rev (curry Logic.mk_implies o prop_of) used (* additional hyps *)
+                      |> Function_Ctx_Tree.export_term (fixes, assumes)
+                      |> fold_rev (curry Logic.mk_implies o prop_of) ags
+                      |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
+                      |> cterm_of thy
+
+            val thm = assume hyp
+                      |> fold forall_elim cqs
+                      |> fold Thm.elim_implies ags
+                      |> Function_Ctx_Tree.import_thm thy (fixes, assumes)
+                      |> fold Thm.elim_implies used (*  "(arg, lhs) : R'"  *)
+
+            val z_eq_arg = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (z, arg))))
+
+            val acc = thm COMP ih_case
+            val z_acc_local = acc
+            |> Conv.fconv_rule (Conv.arg_conv (Conv.arg_conv (K (symmetric (z_eq_arg RS eq_reflection)))))
+
+            val ethm = z_acc_local
+                         |> Function_Ctx_Tree.export_thm thy (fixes,
+                                                          z_eq_arg :: case_hyp :: ags @ assumes)
+                         |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
+
+            val sub' = sub @ [(([],[]), acc)]
+          in
+            (sub', (hyp :: hyps, ethm :: thms))
+          end
+        | step _ _ _ _ = raise Match
+    in
+      Function_Ctx_Tree.traverse_tree step tree
+    end
+
+
+fun mk_nest_term_rule thy globals R R_cases clauses =
+    let
+      val Globals { domT, x, z, ... } = globals
+      val acc_R = mk_acc domT R
+
+      val R' = Free ("R", fastype_of R)
+
+      val Rrel = Free ("R", HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)))
+      val inrel_R = Const (@{const_name FunDef.in_rel}, HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)) --> fastype_of R) $ Rrel
+
+      val wfR' = cterm_of thy (HOLogic.mk_Trueprop (Const (@{const_name Wellfounded.wfP}, (domT --> domT --> boolT) --> boolT) $ R')) (* "wf R'" *)
+
+      (* Inductive Hypothesis: !!z. (z,x):R' ==> z : acc R *)
+      val ihyp = Term.all domT $ Abs ("z", domT,
+                                 Logic.mk_implies (HOLogic.mk_Trueprop (R' $ Bound 0 $ x),
+                                   HOLogic.mk_Trueprop (acc_R $ Bound 0)))
+                     |> cterm_of thy
+
+      val ihyp_a = assume ihyp |> Thm.forall_elim_vars 0
+
+      val R_z_x = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ x))
+
+      val (hyps,cases) = fold (mk_nest_term_case thy globals R' ihyp_a) clauses ([],[])
+    in
+      R_cases
+        |> forall_elim (cterm_of thy z)
+        |> forall_elim (cterm_of thy x)
+        |> forall_elim (cterm_of thy (acc_R $ z))
+        |> curry op COMP (assume R_z_x)
+        |> fold_rev (curry op COMP) cases
+        |> implies_intr R_z_x
+        |> forall_intr (cterm_of thy z)
+        |> (fn it => it COMP accI)
+        |> implies_intr ihyp
+        |> forall_intr (cterm_of thy x)
+        |> (fn it => Drule.compose_single(it,2,wf_induct_rule))
+        |> curry op RS (assume wfR')
+        |> forall_intr_vars
+        |> (fn it => it COMP allI)
+        |> fold implies_intr hyps
+        |> implies_intr wfR'
+        |> forall_intr (cterm_of thy R')
+        |> forall_elim (cterm_of thy (inrel_R))
+        |> curry op RS wf_in_rel
+        |> full_simplify (HOL_basic_ss addsimps [in_rel_def])
+        |> forall_intr (cterm_of thy Rrel)
+    end
+
+
+
+(* Tail recursion (probably very fragile)
+ *
+ * FIXME:
+ * - Need to do forall_elim_vars on psimps: Unneccesary, if psimps would be taken from the same context.
+ * - Must we really replace the fvar by f here?
+ * - Splitting is not configured automatically: Problems with case?
+ *)
+fun mk_trsimps octxt globals f G R f_def R_cases G_induct clauses psimps =
+    let
+      val Globals {domT, ranT, fvar, ...} = globals
+
+      val R_cases = Thm.forall_elim_vars 0 R_cases (* FIXME: Should be already in standard form. *)
+
+      val graph_implies_dom = (* "G ?x ?y ==> dom ?x"  *)
+          Goal.prove octxt ["x", "y"] [HOLogic.mk_Trueprop (G $ Free ("x", domT) $ Free ("y", ranT))]
+                     (HOLogic.mk_Trueprop (mk_acc domT R $ Free ("x", domT)))
+                     (fn {prems=[a], ...} =>
+                         ((rtac (G_induct OF [a]))
+                            THEN_ALL_NEW (rtac accI)
+                            THEN_ALL_NEW (etac R_cases)
+                            THEN_ALL_NEW (asm_full_simp_tac (simpset_of octxt))) 1)
+
+      val default_thm = (forall_intr_vars graph_implies_dom) COMP (f_def COMP fundef_default_value)
+
+      fun mk_trsimp clause psimp =
+          let
+            val ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {ctxt, cqs, qs, gs, lhs, rhs, ...}, ...} = clause
+            val thy = ProofContext.theory_of ctxt
+            val rhs_f = Pattern.rewrite_term thy [(fvar, f)] [] rhs
+
+            val trsimp = Logic.list_implies(gs, HOLogic.mk_Trueprop (HOLogic.mk_eq(f $ lhs, rhs_f))) (* "f lhs = rhs" *)
+            val lhs_acc = (mk_acc domT R $ lhs) (* "acc R lhs" *)
+            fun simp_default_tac ss = asm_full_simp_tac (ss addsimps [default_thm, Let_def])
+          in
+            Goal.prove ctxt [] [] trsimp
+                       (fn _ =>
+                           rtac (instantiate' [] [SOME (cterm_of thy lhs_acc)] case_split) 1
+                                THEN (rtac (Thm.forall_elim_vars 0 psimp) THEN_ALL_NEW assume_tac) 1
+                                THEN (simp_default_tac (simpset_of ctxt) 1)
+                                THEN (etac not_acc_down 1)
+                                THEN ((etac R_cases) THEN_ALL_NEW (simp_default_tac (simpset_of ctxt))) 1)
+              |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
+          end
+    in
+      map2 mk_trsimp clauses psimps
+    end
+
+
+fun prepare_function config defname [((fname, fT), mixfix)] abstract_qglrs lthy =
+    let
+      val FunctionConfig {domintros, tailrec, default=default_str, ...} = config
+
+      val fvar = Free (fname, fT)
+      val domT = domain_type fT
+      val ranT = range_type fT
+
+      val default = Syntax.parse_term lthy default_str
+        |> TypeInfer.constrain fT |> Syntax.check_term lthy
+
+      val (globals, ctxt') = fix_globals domT ranT fvar lthy
+
+      val Globals { x, h, ... } = globals
+
+      val clauses = map (mk_clause_context x ctxt') abstract_qglrs
+
+      val n = length abstract_qglrs
+
+      fun build_tree (ClauseContext { ctxt, rhs, ...}) =
+            Function_Ctx_Tree.mk_tree (fname, fT) h ctxt rhs
+
+      val trees = map build_tree clauses
+      val RCss = map find_calls trees
+
+      val ((G, GIntro_thms, G_elim, G_induct), lthy) =
+          PROFILE "def_graph" (define_graph (graph_name defname) fvar domT ranT clauses RCss) lthy
+
+      val ((f, f_defthm), lthy) =
+          PROFILE "def_fun" (define_function (defname ^ "_sumC_def") (fname, mixfix) domT ranT G default) lthy
+
+      val RCss = map (map (inst_RC (ProofContext.theory_of lthy) fvar f)) RCss
+      val trees = map (Function_Ctx_Tree.inst_tree (ProofContext.theory_of lthy) fvar f) trees
+
+      val ((R, RIntro_thmss, R_elim), lthy) =
+          PROFILE "def_rel" (define_recursion_relation (rel_name defname) domT ranT fvar f abstract_qglrs clauses RCss) lthy
+
+      val (_, lthy) =
+          LocalTheory.abbrev Syntax.mode_default ((Binding.name (dom_name defname), NoSyn), mk_acc domT R) lthy
+
+      val newthy = ProofContext.theory_of lthy
+      val clauses = map (transfer_clause_ctx newthy) clauses
+
+      val cert = cterm_of (ProofContext.theory_of lthy)
+
+      val xclauses = PROFILE "xclauses" (map7 (mk_clause_info globals G f) (1 upto n) clauses abstract_qglrs trees RCss GIntro_thms) RIntro_thmss
+
+      val complete = mk_completeness globals clauses abstract_qglrs |> cert |> assume
+      val compat = mk_compat_proof_obligations domT ranT fvar f abstract_qglrs |> map (cert #> assume)
+
+      val compat_store = store_compat_thms n compat
+
+      val (goalstate, values) = PROFILE "prove_stuff" (prove_stuff lthy globals G f R xclauses complete compat compat_store G_elim) f_defthm
+
+      val mk_trsimps = mk_trsimps lthy globals f G R f_defthm R_elim G_induct xclauses
+
+      fun mk_partial_rules provedgoal =
+          let
+            val newthy = theory_of_thm provedgoal (*FIXME*)
+
+            val (graph_is_function, complete_thm) =
+                provedgoal
+                  |> Conjunction.elim
+                  |> apfst (Thm.forall_elim_vars 0)
+
+            val f_iff = graph_is_function RS (f_defthm RS ex1_implies_iff)
+
+            val psimps = PROFILE "Proving simplification rules" (mk_psimps newthy globals R xclauses values f_iff) graph_is_function
+
+            val simple_pinduct = PROFILE "Proving partial induction rule"
+                                                           (mk_partial_induct_rule newthy globals R complete_thm) xclauses
+
+
+            val total_intro = PROFILE "Proving nested termination rule" (mk_nest_term_rule newthy globals R R_elim) xclauses
+
+            val dom_intros = if domintros
+                             then SOME (PROFILE "Proving domain introduction rules" (map (mk_domain_intro lthy globals R R_elim)) xclauses)
+                             else NONE
+            val trsimps = if tailrec then SOME (mk_trsimps psimps) else NONE
+
+          in
+            FunctionResult {fs=[f], G=G, R=R, cases=complete_thm,
+                          psimps=psimps, simple_pinducts=[simple_pinduct],
+                          termination=total_intro, trsimps=trsimps,
+                          domintros=dom_intros}
+          end
+    in
+      ((f, goalstate, mk_partial_rules), lthy)
+    end
+
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/function_lib.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -0,0 +1,180 @@
+(*  Title:      HOL/Tools/Function/fundef_lib.ML
+    Author:     Alexander Krauss, TU Muenchen
+
+A package for general recursive function definitions. 
+Some fairly general functions that should probably go somewhere else... 
+*)
+
+structure Function_Lib =
+struct
+
+fun map_option f NONE = NONE 
+  | map_option f (SOME x) = SOME (f x);
+
+fun fold_option f NONE y = y
+  | fold_option f (SOME x) y = f x y;
+
+fun fold_map_option f NONE y = (NONE, y)
+  | fold_map_option f (SOME x) y = apfst SOME (f x y);
+
+(* Ex: "The variable" ^ plural " is" "s are" vs *)
+fun plural sg pl [x] = sg
+  | plural sg pl _ = pl
+
+(* lambda-abstracts over an arbitrarily nested tuple
+  ==> hologic.ML? *)
+fun tupled_lambda vars t =
+    case vars of
+      (Free v) => lambda (Free v) t
+    | (Var v) => lambda (Var v) t
+    | (Const ("Pair", Type ("fun", [Ta, Type ("fun", [Tb, _])]))) $ us $ vs =>  
+      (HOLogic.split_const (Ta,Tb, fastype_of t)) $ (tupled_lambda us (tupled_lambda vs t))
+    | _ => raise Match
+                 
+                 
+fun dest_all (Const ("all", _) $ Abs (a as (_,T,_))) =
+    let
+      val (n, body) = Term.dest_abs a
+    in
+      (Free (n, T), body)
+    end
+  | dest_all _ = raise Match
+                         
+
+(* Removes all quantifiers from a term, replacing bound variables by frees. *)
+fun dest_all_all (t as (Const ("all",_) $ _)) = 
+    let
+      val (v,b) = dest_all t
+      val (vs, b') = dest_all_all b
+    in
+      (v :: vs, b')
+    end
+  | dest_all_all t = ([],t)
+                     
+
+(* FIXME: similar to Variable.focus *)
+fun dest_all_all_ctx ctx (Const ("all", _) $ Abs (a as (n,T,b))) =
+    let
+      val [(n', _)] = Variable.variant_frees ctx [] [(n,T)]
+      val (_, ctx') = ProofContext.add_fixes [(Binding.name n', SOME T, NoSyn)] ctx
+
+      val (n'', body) = Term.dest_abs (n', T, b) 
+      val _ = (n' = n'') orelse error "dest_all_ctx"
+      (* Note: We assume that n' does not occur in the body. Otherwise it would be fixed. *)
+
+      val (ctx'', vs, bd) = dest_all_all_ctx ctx' body
+    in
+      (ctx'', (n', T) :: vs, bd)
+    end
+  | dest_all_all_ctx ctx t = 
+    (ctx, [], t)
+
+
+fun map3 _ [] [] [] = []
+  | map3 f (x :: xs) (y :: ys) (z :: zs) = f x y z :: map3 f xs ys zs
+  | map3 _ _ _ _ = raise Library.UnequalLengths;
+
+fun map4 _ [] [] [] [] = []
+  | map4 f (x :: xs) (y :: ys) (z :: zs) (u :: us) = f x y z u :: map4 f xs ys zs us
+  | map4 _ _ _ _ _ = raise Library.UnequalLengths;
+
+fun map6 _ [] [] [] [] [] [] = []
+  | map6 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (w :: ws) = f x y z u v w :: map6 f xs ys zs us vs ws
+  | map6 _ _ _ _ _ _ _ = raise Library.UnequalLengths;
+
+fun map7 _ [] [] [] [] [] [] [] = []
+  | map7 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (w :: ws) (b :: bs) = f x y z u v w b :: map7 f xs ys zs us vs ws bs
+  | map7 _ _ _ _ _ _ _ _ = raise Library.UnequalLengths;
+
+
+
+(* forms all "unordered pairs": [1, 2, 3] ==> [(1, 1), (1, 2), (1, 3), (2, 2), (2, 3), (3, 3)] *)
+(* ==> library *)
+fun unordered_pairs [] = []
+  | unordered_pairs (x::xs) = map (pair x) (x::xs) @ unordered_pairs xs
+
+
+(* Replaces Frees by name. Works with loose Bounds. *)
+fun replace_frees assoc =
+    map_aterms (fn c as Free (n, _) => the_default c (AList.lookup (op =) assoc n)
+                 | t => t)
+
+
+fun rename_bound n (Q $ Abs(_, T, b)) = (Q $ Abs(n, T, b))
+  | rename_bound n _ = raise Match
+
+fun mk_forall_rename (n, v) =
+    rename_bound n o Logic.all v 
+
+fun forall_intr_rename (n, cv) thm =
+    let
+      val allthm = forall_intr cv thm
+      val (_ $ abs) = prop_of allthm
+    in
+      Thm.rename_boundvars abs (Abs (n, dummyT, Term.dummy_pattern dummyT)) allthm
+    end
+
+
+(* Returns the frees in a term in canonical order, excluding the fixes from the context *)
+fun frees_in_term ctxt t =
+    Term.add_frees t []
+    |> filter_out (Variable.is_fixed ctxt o fst)
+    |> rev
+
+
+datatype proof_attempt = Solved of thm | Stuck of thm | Fail
+
+fun try_proof cgoal tac = 
+    case SINGLE tac (Goal.init cgoal) of
+      NONE => Fail
+    | SOME st =>
+        if Thm.no_prems st
+        then Solved (Goal.finish (Syntax.init_pretty_global (Thm.theory_of_cterm cgoal)) st)
+        else Stuck st
+
+
+fun dest_binop_list cn (t as (Const (n, _) $ a $ b)) = 
+    if cn = n then dest_binop_list cn a @ dest_binop_list cn b else [ t ]
+  | dest_binop_list _ t = [ t ]
+
+
+(* separate two parts in a +-expression:
+   "a + b + c + d + e" --> "(a + b + d) + (c + e)"
+
+   Here, + can be any binary operation that is AC.
+
+   cn - The name of the binop-constructor (e.g. @{const_name Un})
+   ac - the AC rewrite rules for cn
+   is - the list of indices of the expressions that should become the first part
+        (e.g. [0,1,3] in the above example)
+*)
+
+fun regroup_conv neu cn ac is ct =
+ let
+   val mk = HOLogic.mk_binop cn
+   val t = term_of ct
+   val xs = dest_binop_list cn t
+   val js = subtract (op =) is (0 upto (length xs) - 1)
+   val ty = fastype_of t
+   val thy = theory_of_cterm ct
+ in
+   Goal.prove_internal []
+     (cterm_of thy
+       (Logic.mk_equals (t,
+          if is = []
+          then mk (Const (neu, ty), foldr1 mk (map (nth xs) js))
+          else if js = []
+            then mk (foldr1 mk (map (nth xs) is), Const (neu, ty))
+            else mk (foldr1 mk (map (nth xs) is), foldr1 mk (map (nth xs) js)))))
+     (K (rewrite_goals_tac ac
+         THEN rtac Drule.reflexive_thm 1))
+ end
+
+(* instance for unions *)
+fun regroup_union_conv t = regroup_conv @{const_name Set.empty} @{const_name Lattices.sup}
+  (map (fn t => t RS eq_reflection) (@{thms Un_ac} @
+                                     @{thms Un_empty_right} @
+                                     @{thms Un_empty_left})) t
+
+
+end
--- a/src/HOL/Tools/Function/fundef.ML	Fri Oct 23 13:23:18 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,226 +0,0 @@
-(*  Title:      HOL/Tools/Function/fundef.ML
-    Author:     Alexander Krauss, TU Muenchen
-
-A package for general recursive function definitions.
-Isar commands.
-*)
-
-signature FUNDEF =
-sig
-    val add_fundef :  (binding * typ option * mixfix) list
-                       -> (Attrib.binding * term) list
-                       -> FundefCommon.fundef_config
-                       -> local_theory
-                       -> Proof.state
-    val add_fundef_cmd :  (binding * string option * mixfix) list
-                      -> (Attrib.binding * string) list
-                      -> FundefCommon.fundef_config
-                      -> local_theory
-                      -> Proof.state
-
-    val termination_proof : term option -> local_theory -> Proof.state
-    val termination_proof_cmd : string option -> local_theory -> Proof.state
-    val termination : term option -> local_theory -> Proof.state
-    val termination_cmd : string option -> local_theory -> Proof.state
-
-    val setup : theory -> theory
-    val get_congs : Proof.context -> thm list
-end
-
-
-structure Fundef : FUNDEF =
-struct
-
-open FundefLib
-open FundefCommon
-
-val simp_attribs = map (Attrib.internal o K)
-    [Simplifier.simp_add,
-     Code.add_default_eqn_attribute,
-     Nitpick_Simps.add,
-     Quickcheck_RecFun_Simps.add]
-
-val psimp_attribs = map (Attrib.internal o K)
-    [Simplifier.simp_add,
-     Nitpick_Psimps.add]
-
-fun note_theorem ((name, atts), ths) =
-  LocalTheory.note Thm.generatedK ((Binding.qualified_name name, atts), ths)
-
-fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_"
-
-fun add_simps fnames post sort extra_qualify label moreatts simps lthy =
-    let
-      val spec = post simps
-                   |> map (apfst (apsnd (fn ats => moreatts @ ats)))
-                   |> map (apfst (apfst extra_qualify))
-
-      val (saved_spec_simps, lthy) =
-        fold_map (LocalTheory.note Thm.generatedK) spec lthy
-
-      val saved_simps = maps snd saved_spec_simps
-      val simps_by_f = sort saved_simps
-
-      fun add_for_f fname simps =
-        note_theorem ((Long_Name.qualify fname label, []), simps) #> snd
-    in
-      (saved_simps,
-       fold2 add_for_f fnames simps_by_f lthy)
-    end
-
-fun gen_add_fundef is_external prep default_constraint fixspec eqns config lthy =
-    let
-      val constrn_fxs = map (fn (b, T, mx) => (b, SOME (the_default default_constraint T), mx))
-      val ((fixes0, spec0), ctxt') = prep (constrn_fxs fixspec) eqns lthy
-      val fixes = map (apfst (apfst Binding.name_of)) fixes0;
-      val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0;
-      val (eqs, post, sort_cont, cnames) = FundefCommon.get_preproc lthy config ctxt' fixes spec
-
-      val defname = mk_defname fixes
-
-      val ((goalstate, cont), lthy) =
-          FundefMutual.prepare_fundef_mutual config defname fixes eqs lthy
-
-      fun afterqed [[proof]] lthy =
-        let
-          val FundefResult {fs, R, psimps, trsimps,  simple_pinducts, termination,
-                            domintros, cases, ...} =
-          cont (Thm.close_derivation proof)
-
-          val fnames = map (fst o fst) fixes
-          val qualify = Long_Name.qualify defname
-          val addsmps = add_simps fnames post sort_cont
-
-          val (((psimps', pinducts'), (_, [termination'])), lthy) =
-            lthy
-            |> addsmps (Binding.qualify false "partial") "psimps"
-                 psimp_attribs psimps
-            ||> fold_option (snd oo addsmps I "simps" simp_attribs) trsimps
-            ||>> note_theorem ((qualify "pinduct",
-                   [Attrib.internal (K (RuleCases.case_names cnames)),
-                    Attrib.internal (K (RuleCases.consumes 1)),
-                    Attrib.internal (K (Induct.induct_pred ""))]), simple_pinducts)
-            ||>> note_theorem ((qualify "termination", []), [termination])
-            ||> (snd o note_theorem ((qualify "cases",
-                   [Attrib.internal (K (RuleCases.case_names cnames))]), [cases]))
-            ||> fold_option (snd oo curry note_theorem (qualify "domintros", [])) domintros
-
-          val cdata = FundefCtxData { add_simps=addsmps, case_names=cnames, psimps=psimps',
-                                      pinducts=snd pinducts', termination=termination',
-                                      fs=fs, R=R, defname=defname }
-          val _ =
-            if not is_external then ()
-            else Specification.print_consts lthy (K false) (map fst fixes)
-        in
-          lthy
-          |> LocalTheory.declaration (add_fundef_data o morph_fundef_data cdata)
-        end
-    in
-      lthy
-        |> is_external ? LocalTheory.set_group (serial_string ())
-        |> Proof.theorem_i NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]]
-        |> Proof.refine (Method.primitive_text (fn _ => goalstate)) |> Seq.hd
-    end
-
-val add_fundef = gen_add_fundef false Specification.check_spec (TypeInfer.anyT HOLogic.typeS)
-val add_fundef_cmd = gen_add_fundef true Specification.read_spec "_::type"
-
-fun gen_termination_proof prep_term raw_term_opt lthy =
-    let
-      val term_opt = Option.map (prep_term lthy) raw_term_opt
-      val data = the (case term_opt of
-                        SOME t => (import_fundef_data t lthy
-                          handle Option.Option =>
-                            error ("Not a function: " ^ quote (Syntax.string_of_term lthy t)))
-                      | NONE => (import_last_fundef lthy handle Option.Option => error "Not a function"))
-
-        val FundefCtxData { termination, R, add_simps, case_names, psimps,
-                            pinducts, defname, ...} = data
-        val domT = domain_type (fastype_of R)
-        val goal = HOLogic.mk_Trueprop
-                     (HOLogic.mk_all ("x", domT, mk_acc domT R $ Free ("x", domT)))
-        fun afterqed [[totality]] lthy =
-          let
-            val totality = Thm.close_derivation totality
-            val remove_domain_condition =
-              full_simplify (HOL_basic_ss addsimps [totality, True_implies_equals])
-            val tsimps = map remove_domain_condition psimps
-            val tinduct = map remove_domain_condition pinducts
-            val qualify = Long_Name.qualify defname;
-          in
-            lthy
-            |> add_simps I "simps" simp_attribs tsimps |> snd
-            |> note_theorem
-               ((qualify "induct",
-                 [Attrib.internal (K (RuleCases.case_names case_names))]),
-                tinduct) |> snd
-          end
-    in
-      lthy
-      |> ProofContext.note_thmss ""
-         [((Binding.empty, [ContextRules.rule_del]), [([allI], [])])] |> snd
-      |> ProofContext.note_thmss ""
-         [((Binding.empty, [ContextRules.intro_bang (SOME 1)]), [([allI], [])])] |> snd
-      |> ProofContext.note_thmss ""
-         [((Binding.name "termination", [ContextRules.intro_bang (SOME 0)]),
-           [([Goal.norm_result termination], [])])] |> snd
-      |> Proof.theorem_i NONE afterqed [[(goal, [])]]
-    end
-
-val termination_proof = gen_termination_proof Syntax.check_term;
-val termination_proof_cmd = gen_termination_proof Syntax.read_term;
-
-fun termination term_opt lthy =
-  lthy
-  |> LocalTheory.set_group (serial_string ())
-  |> termination_proof term_opt;
-
-fun termination_cmd term_opt lthy =
-  lthy
-  |> LocalTheory.set_group (serial_string ())
-  |> termination_proof_cmd term_opt;
-
-
-(* Datatype hook to declare datatype congs as "fundef_congs" *)
-
-
-fun add_case_cong n thy =
-    Context.theory_map (FundefCtxTree.map_fundef_congs (Thm.add_thm
-                          (Datatype.the_info thy n
-                           |> #case_cong
-                           |> safe_mk_meta_eq)))
-                       thy
-
-val setup_case_cong = Datatype.interpretation (K (fold add_case_cong))
-
-
-(* setup *)
-
-val setup =
-  Attrib.setup @{binding fundef_cong}
-    (Attrib.add_del FundefCtxTree.cong_add FundefCtxTree.cong_del)
-    "declaration of congruence rule for function definitions"
-  #> setup_case_cong
-  #> FundefRelation.setup
-  #> FundefCommon.Termination_Simps.setup
-
-val get_congs = FundefCtxTree.get_fundef_congs
-
-
-(* outer syntax *)
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val _ =
-  OuterSyntax.local_theory_to_proof "function" "define general recursive functions" K.thy_goal
-  (fundef_parser default_config
-     >> (fn ((config, fixes), statements) => add_fundef_cmd fixes statements config));
-
-val _ =
-  OuterSyntax.local_theory_to_proof "termination" "prove termination of a recursive function" K.thy_goal
-  (Scan.option P.term >> termination_cmd);
-
-end;
-
-
-end
--- a/src/HOL/Tools/Function/fundef_common.ML	Fri Oct 23 13:23:18 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,343 +0,0 @@
-(*  Title:      HOL/Tools/Function/fundef_common.ML
-    Author:     Alexander Krauss, TU Muenchen
-
-A package for general recursive function definitions. 
-Common definitions and other infrastructure.
-*)
-
-structure FundefCommon =
-struct
-
-local open FundefLib in
-
-(* Profiling *)
-val profile = Unsynchronized.ref false;
-
-fun PROFILE msg = if !profile then timeap_msg msg else I
-
-
-val acc_const_name = @{const_name accp}
-fun mk_acc domT R =
-    Const (acc_const_name, (domT --> domT --> HOLogic.boolT) --> domT --> HOLogic.boolT) $ R 
-
-val function_name = suffix "C"
-val graph_name = suffix "_graph"
-val rel_name = suffix "_rel"
-val dom_name = suffix "_dom"
-
-(* Termination rules *)
-
-structure TerminationRule = GenericDataFun
-(
-  type T = thm list
-  val empty = []
-  val extend = I
-  fun merge _ = Thm.merge_thms
-);
-
-val get_termination_rules = TerminationRule.get
-val store_termination_rule = TerminationRule.map o cons
-val apply_termination_rule = resolve_tac o get_termination_rules o Context.Proof
-
-
-(* Function definition result data *)
-
-datatype fundef_result =
-  FundefResult of
-     {
-      fs: term list,
-      G: term,
-      R: term,
-
-      psimps : thm list, 
-      trsimps : thm list option, 
-
-      simple_pinducts : thm list, 
-      cases : thm,
-      termination : thm,
-      domintros : thm list option
-     }
-
-
-datatype fundef_context_data =
-  FundefCtxData of
-     {
-      defname : string,
-
-      (* contains no logical entities: invariant under morphisms *)
-      add_simps : (binding -> binding) -> string -> Attrib.src list -> thm list 
-                  -> local_theory -> thm list * local_theory,
-      case_names : string list,
-
-      fs : term list,
-      R : term,
-      
-      psimps: thm list,
-      pinducts: thm list,
-      termination: thm
-     }
-
-fun morph_fundef_data (FundefCtxData {add_simps, case_names, fs, R, 
-                                      psimps, pinducts, termination, defname}) phi =
-    let
-      val term = Morphism.term phi val thm = Morphism.thm phi val fact = Morphism.fact phi
-      val name = Binding.name_of o Morphism.binding phi o Binding.name
-    in
-      FundefCtxData { add_simps = add_simps, case_names = case_names,
-                      fs = map term fs, R = term R, psimps = fact psimps, 
-                      pinducts = fact pinducts, termination = thm termination,
-                      defname = name defname }
-    end
-
-structure FundefData = GenericDataFun
-(
-  type T = (term * fundef_context_data) Item_Net.T;
-  val empty = Item_Net.init
-    (op aconv o pairself fst : (term * fundef_context_data) * (term * fundef_context_data) -> bool)
-    fst;
-  val copy = I;
-  val extend = I;
-  fun merge _ (tab1, tab2) = Item_Net.merge (tab1, tab2)
-);
-
-val get_fundef = FundefData.get o Context.Proof;
-
-
-(* Generally useful?? *)
-fun lift_morphism thy f = 
-    let 
-      val term = Drule.term_rule thy f
-    in
-      Morphism.thm_morphism f $> Morphism.term_morphism term 
-       $> Morphism.typ_morphism (Logic.type_map term)
-    end
-
-fun import_fundef_data t ctxt =
-    let
-      val thy = ProofContext.theory_of ctxt
-      val ct = cterm_of thy t
-      val inst_morph = lift_morphism thy o Thm.instantiate 
-
-      fun match (trm, data) = 
-          SOME (morph_fundef_data data (inst_morph (Thm.match (cterm_of thy trm, ct))))
-          handle Pattern.MATCH => NONE
-    in 
-      get_first match (Item_Net.retrieve (get_fundef ctxt) t)
-    end
-
-fun import_last_fundef ctxt =
-    case Item_Net.content (get_fundef ctxt) of
-      [] => NONE
-    | (t, data) :: _ =>
-      let 
-        val ([t'], ctxt') = Variable.import_terms true [t] ctxt
-      in
-        import_fundef_data t' ctxt'
-      end
-
-val all_fundef_data = Item_Net.content o get_fundef
-
-fun add_fundef_data (data as FundefCtxData {fs, termination, ...}) =
-    FundefData.map (fold (fn f => Item_Net.insert (f, data)) fs)
-    #> store_termination_rule termination
-
-
-(* Simp rules for termination proofs *)
-
-structure Termination_Simps = Named_Thms
-(
-  val name = "termination_simp" 
-  val description = "Simplification rule for termination proofs"
-);
-
-
-(* Default Termination Prover *)
-
-structure TerminationProver = GenericDataFun
-(
-  type T = Proof.context -> Proof.method
-  val empty = (fn _ => error "Termination prover not configured")
-  val extend = I
-  fun merge _ (a,b) = b (* FIXME *)
-);
-
-val set_termination_prover = TerminationProver.put
-val get_termination_prover = TerminationProver.get o Context.Proof
-
-
-(* Configuration management *)
-datatype fundef_opt 
-  = Sequential
-  | Default of string
-  | DomIntros
-  | Tailrec
-
-datatype fundef_config
-  = FundefConfig of
-   {
-    sequential: bool,
-    default: string,
-    domintros: bool,
-    tailrec: bool
-   }
-
-fun apply_opt Sequential (FundefConfig {sequential, default, domintros,tailrec}) = 
-    FundefConfig {sequential=true, default=default, domintros=domintros, tailrec=tailrec}
-  | apply_opt (Default d) (FundefConfig {sequential, default, domintros,tailrec}) = 
-    FundefConfig {sequential=sequential, default=d, domintros=domintros, tailrec=tailrec}
-  | apply_opt DomIntros (FundefConfig {sequential, default, domintros,tailrec}) =
-    FundefConfig {sequential=sequential, default=default, domintros=true,tailrec=tailrec}
-  | apply_opt Tailrec (FundefConfig {sequential, default, domintros,tailrec}) =
-    FundefConfig {sequential=sequential, default=default, domintros=domintros,tailrec=true}
-
-val default_config =
-  FundefConfig { sequential=false, default="%x. undefined" (*FIXME dynamic scoping*), 
-                 domintros=false, tailrec=false }
-
-
-(* Analyzing function equations *)
-
-fun split_def ctxt geq =
-    let
-      fun input_error msg = cat_lines [msg, Syntax.string_of_term ctxt geq]
-      val qs = Term.strip_qnt_vars "all" geq
-      val imp = Term.strip_qnt_body "all" geq
-      val (gs, eq) = Logic.strip_horn imp
-
-      val (f_args, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop eq)
-          handle TERM _ => error (input_error "Not an equation")
-
-      val (head, args) = strip_comb f_args
-
-      val fname = fst (dest_Free head)
-          handle TERM _ => error (input_error "Head symbol must not be a bound variable")
-    in
-      (fname, qs, gs, args, rhs)
-    end
-
-(* Check for all sorts of errors in the input *)
-fun check_defs ctxt fixes eqs =
-    let
-      val fnames = map (fst o fst) fixes
-                                
-      fun check geq = 
-          let
-            fun input_error msg = error (cat_lines [msg, Syntax.string_of_term ctxt geq])
-                                  
-            val fqgar as (fname, qs, gs, args, rhs) = split_def ctxt geq
-                                 
-            val _ = fname mem fnames 
-                    orelse input_error 
-                             ("Head symbol of left hand side must be " 
-                              ^ plural "" "one out of " fnames ^ commas_quote fnames)
-                                            
-            val _ = length args > 0 orelse input_error "Function has no arguments:"
-
-            fun add_bvs t is = add_loose_bnos (t, 0, is)
-            val rvs = (subtract (op =) (fold add_bvs args []) (add_bvs rhs []))
-                        |> map (fst o nth (rev qs))
-                      
-            val _ = null rvs orelse input_error 
-                        ("Variable" ^ plural " " "s " rvs ^ commas_quote rvs
-                         ^ " occur" ^ plural "s" "" rvs ^ " on right hand side only:")
-                                    
-            val _ = forall (not o Term.exists_subterm 
-                             (fn Free (n, _) => n mem fnames | _ => false)) (gs @ args)
-                    orelse input_error "Defined function may not occur in premises or arguments"
-
-            val freeargs = map (fn t => subst_bounds (rev (map Free qs), t)) args
-            val funvars = filter (fn q => exists (exists_subterm (fn (Free q') $ _ => q = q' | _ => false)) freeargs) qs
-            val _ = null funvars
-                    orelse (warning (cat_lines 
-                    ["Bound variable" ^ plural " " "s " funvars 
-                     ^ commas_quote (map fst funvars) ^  
-                     " occur" ^ plural "s" "" funvars ^ " in function position.",  
-                     "Misspelled constructor???"]); true)
-          in
-            (fname, length args)
-          end
-
-      val _ = AList.group (op =) (map check eqs)
-        |> map (fn (fname, ars) =>
-             length (distinct (op =) ars) = 1
-             orelse error ("Function " ^ quote fname ^
-                           " has different numbers of arguments in different equations"))
-
-      fun check_sorts ((fname, fT), _) =
-          Sorts.of_sort (Sign.classes_of (ProofContext.theory_of ctxt)) (fT, HOLogic.typeS)
-          orelse error (cat_lines 
-          ["Type of " ^ quote fname ^ " is not of sort " ^ quote "type" ^ ":",
-           setmp_CRITICAL show_sorts true (Syntax.string_of_typ ctxt) fT])
-
-      val _ = map check_sorts fixes
-    in
-      ()
-    end
-
-(* Preprocessors *)
-
-type fixes = ((string * typ) * mixfix) list
-type 'a spec = (Attrib.binding * 'a list) list
-type preproc = fundef_config -> Proof.context -> fixes -> term spec 
-               -> (term list * (thm list -> thm spec) * (thm list -> thm list list) * string list)
-
-val fname_of = fst o dest_Free o fst o strip_comb o fst 
- o HOLogic.dest_eq o HOLogic.dest_Trueprop o Logic.strip_imp_concl o snd o dest_all_all
-
-fun mk_case_names i "" k = mk_case_names i (string_of_int (i + 1)) k
-  | mk_case_names _ n 0 = []
-  | mk_case_names _ n 1 = [n]
-  | mk_case_names _ n k = map (fn i => n ^ "_" ^ string_of_int i) (1 upto k)
-
-fun empty_preproc check _ ctxt fixes spec =
-    let 
-      val (bnds, tss) = split_list spec
-      val ts = flat tss
-      val _ = check ctxt fixes ts
-      val fnames = map (fst o fst) fixes
-      val indices = map (fn eq => find_index (curry op = (fname_of eq)) fnames) ts
-
-      fun sort xs = partition_list (fn i => fn (j,_) => i = j) 0 (length fnames - 1) 
-                                   (indices ~~ xs)
-                        |> map (map snd)
-
-      (* using theorem names for case name currently disabled *)
-      val cnames = map_index (fn (i, _) => mk_case_names i "" 1) bnds |> flat
-    in
-      (ts, curry op ~~ bnds o Library.unflat tss, sort, cnames)
-    end
-
-structure Preprocessor = GenericDataFun
-(
-  type T = preproc
-  val empty : T = empty_preproc check_defs
-  val extend = I
-  fun merge _ (a, _) = a
-);
-
-val get_preproc = Preprocessor.get o Context.Proof
-val set_preproc = Preprocessor.map o K
-
-
-
-local 
-  structure P = OuterParse and K = OuterKeyword
-
-  val option_parser = 
-      P.group "option" ((P.reserved "sequential" >> K Sequential)
-                    || ((P.reserved "default" |-- P.term) >> Default)
-                    || (P.reserved "domintros" >> K DomIntros)
-                    || (P.reserved "tailrec" >> K Tailrec))
-
-  fun config_parser default = 
-      (Scan.optional (P.$$$ "(" |-- P.!!! (P.list1 option_parser) --| P.$$$ ")") [])
-        >> (fn opts => fold apply_opt opts default)
-in
-  fun fundef_parser default_cfg = 
-      config_parser default_cfg -- P.fixes -- SpecParse.where_alt_specs
-end
-
-
-end
-end
-
--- a/src/HOL/Tools/Function/fundef_core.ML	Fri Oct 23 13:23:18 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,956 +0,0 @@
-(*  Title:      HOL/Tools/Function/fundef_core.ML
-    Author:     Alexander Krauss, TU Muenchen
-
-A package for general recursive function definitions:
-Main functionality.
-*)
-
-signature FUNDEF_CORE =
-sig
-    val trace: bool Unsynchronized.ref
-
-    val prepare_fundef : FundefCommon.fundef_config
-                         -> string (* defname *)
-                         -> ((bstring * typ) * mixfix) list (* defined symbol *)
-                         -> ((bstring * typ) list * term list * term * term) list (* specification *)
-                         -> local_theory
-
-                         -> (term   (* f *)
-                             * thm  (* goalstate *)
-                             * (thm -> FundefCommon.fundef_result) (* continuation *)
-                            ) * local_theory
-
-end
-
-structure FundefCore : FUNDEF_CORE =
-struct
-
-val trace = Unsynchronized.ref false;
-fun trace_msg msg = if ! trace then tracing (msg ()) else ();
-
-val boolT = HOLogic.boolT
-val mk_eq = HOLogic.mk_eq
-
-open FundefLib
-open FundefCommon
-
-datatype globals =
-   Globals of {
-         fvar: term,
-         domT: typ,
-         ranT: typ,
-         h: term,
-         y: term,
-         x: term,
-         z: term,
-         a: term,
-         P: term,
-         D: term,
-         Pbool:term
-}
-
-
-datatype rec_call_info =
-  RCInfo of
-  {
-   RIvs: (string * typ) list,  (* Call context: fixes and assumes *)
-   CCas: thm list,
-   rcarg: term,                 (* The recursive argument *)
-
-   llRI: thm,
-   h_assum: term
-  }
-
-
-datatype clause_context =
-  ClauseContext of
-  {
-    ctxt : Proof.context,
-
-    qs : term list,
-    gs : term list,
-    lhs: term,
-    rhs: term,
-
-    cqs: cterm list,
-    ags: thm list,
-    case_hyp : thm
-  }
-
-
-fun transfer_clause_ctx thy (ClauseContext { ctxt, qs, gs, lhs, rhs, cqs, ags, case_hyp }) =
-    ClauseContext { ctxt = ProofContext.transfer thy ctxt,
-                    qs = qs, gs = gs, lhs = lhs, rhs = rhs, cqs = cqs, ags = ags, case_hyp = case_hyp }
-
-
-datatype clause_info =
-  ClauseInfo of
-     {
-      no: int,
-      qglr : ((string * typ) list * term list * term * term),
-      cdata : clause_context,
-
-      tree: FundefCtxTree.ctx_tree,
-      lGI: thm,
-      RCs: rec_call_info list
-     }
-
-
-(* Theory dependencies. *)
-val Pair_inject = @{thm Product_Type.Pair_inject};
-
-val acc_induct_rule = @{thm accp_induct_rule};
-
-val ex1_implies_ex = @{thm FunDef.fundef_ex1_existence};
-val ex1_implies_un = @{thm FunDef.fundef_ex1_uniqueness};
-val ex1_implies_iff = @{thm FunDef.fundef_ex1_iff};
-
-val acc_downward = @{thm accp_downward};
-val accI = @{thm accp.accI};
-val case_split = @{thm HOL.case_split};
-val fundef_default_value = @{thm FunDef.fundef_default_value};
-val not_acc_down = @{thm not_accp_down};
-
-
-
-fun find_calls tree =
-    let
-      fun add_Ri (fixes,assumes) (_ $ arg) _ (_, xs) = ([], (fixes, assumes, arg) :: xs)
-        | add_Ri _ _ _ _ = raise Match
-    in
-      rev (FundefCtxTree.traverse_tree add_Ri tree [])
-    end
-
-
-(** building proof obligations *)
-
-fun mk_compat_proof_obligations domT ranT fvar f glrs =
-    let
-      fun mk_impl ((qs, gs, lhs, rhs),(qs', gs', lhs', rhs')) =
-          let
-            val shift = incr_boundvars (length qs')
-          in
-            Logic.mk_implies
-              (HOLogic.mk_Trueprop (HOLogic.eq_const domT $ shift lhs $ lhs'),
-                HOLogic.mk_Trueprop (HOLogic.eq_const ranT $ shift rhs $ rhs'))
-              |> fold_rev (curry Logic.mk_implies) (map shift gs @ gs')
-              |> fold_rev (fn (n,T) => fn b => Term.all T $ Abs(n,T,b)) (qs @ qs')
-              |> curry abstract_over fvar
-              |> curry subst_bound f
-          end
-    in
-      map mk_impl (unordered_pairs glrs)
-    end
-
-
-fun mk_completeness (Globals {x, Pbool, ...}) clauses qglrs =
-    let
-        fun mk_case (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) =
-            HOLogic.mk_Trueprop Pbool
-                     |> curry Logic.mk_implies (HOLogic.mk_Trueprop (mk_eq (x, lhs)))
-                     |> fold_rev (curry Logic.mk_implies) gs
-                     |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
-    in
-        HOLogic.mk_Trueprop Pbool
-                 |> fold_rev (curry Logic.mk_implies o mk_case) (clauses ~~ qglrs)
-                 |> mk_forall_rename ("x", x)
-                 |> mk_forall_rename ("P", Pbool)
-    end
-
-(** making a context with it's own local bindings **)
-
-fun mk_clause_context x ctxt (pre_qs,pre_gs,pre_lhs,pre_rhs) =
-    let
-      val (qs, ctxt') = Variable.variant_fixes (map fst pre_qs) ctxt
-                                           |>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs
-
-      val thy = ProofContext.theory_of ctxt'
-
-      fun inst t = subst_bounds (rev qs, t)
-      val gs = map inst pre_gs
-      val lhs = inst pre_lhs
-      val rhs = inst pre_rhs
-
-      val cqs = map (cterm_of thy) qs
-      val ags = map (assume o cterm_of thy) gs
-
-      val case_hyp = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (x, lhs))))
-    in
-      ClauseContext { ctxt = ctxt', qs = qs, gs = gs, lhs = lhs, rhs = rhs,
-                      cqs = cqs, ags = ags, case_hyp = case_hyp }
-    end
-
-
-(* lowlevel term function. FIXME: remove *)
-fun abstract_over_list vs body =
-  let
-    fun abs lev v tm =
-      if v aconv tm then Bound lev
-      else
-        (case tm of
-          Abs (a, T, t) => Abs (a, T, abs (lev + 1) v t)
-        | t $ u => abs lev v t $ abs lev v u
-        | t => t);
-  in
-    fold_index (fn (i, v) => fn t => abs i v t) vs body
-  end
-
-
-
-fun mk_clause_info globals G f no cdata qglr tree RCs GIntro_thm RIntro_thms =
-    let
-        val Globals {h, fvar, x, ...} = globals
-
-        val ClauseContext { ctxt, qs, cqs, ags, ... } = cdata
-        val cert = Thm.cterm_of (ProofContext.theory_of ctxt)
-
-        (* Instantiate the GIntro thm with "f" and import into the clause context. *)
-        val lGI = GIntro_thm
-                    |> forall_elim (cert f)
-                    |> fold forall_elim cqs
-                    |> fold Thm.elim_implies ags
-
-        fun mk_call_info (rcfix, rcassm, rcarg) RI =
-            let
-                val llRI = RI
-                             |> fold forall_elim cqs
-                             |> fold (forall_elim o cert o Free) rcfix
-                             |> fold Thm.elim_implies ags
-                             |> fold Thm.elim_implies rcassm
-
-                val h_assum =
-                    HOLogic.mk_Trueprop (G $ rcarg $ (h $ rcarg))
-                              |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
-                              |> fold_rev (Logic.all o Free) rcfix
-                              |> Pattern.rewrite_term (ProofContext.theory_of ctxt) [(f, h)] []
-                              |> abstract_over_list (rev qs)
-            in
-                RCInfo {RIvs=rcfix, rcarg=rcarg, CCas=rcassm, llRI=llRI, h_assum=h_assum}
-            end
-
-        val RC_infos = map2 mk_call_info RCs RIntro_thms
-    in
-        ClauseInfo
-            {
-             no=no,
-             cdata=cdata,
-             qglr=qglr,
-
-             lGI=lGI,
-             RCs=RC_infos,
-             tree=tree
-            }
-    end
-
-
-
-
-
-
-
-(* replace this by a table later*)
-fun store_compat_thms 0 thms = []
-  | store_compat_thms n thms =
-    let
-        val (thms1, thms2) = chop n thms
-    in
-        (thms1 :: store_compat_thms (n - 1) thms2)
-    end
-
-(* expects i <= j *)
-fun lookup_compat_thm i j cts =
-    nth (nth cts (i - 1)) (j - i)
-
-(* Returns "Gsi, Gsj, lhs_i = lhs_j |-- rhs_j_f = rhs_i_f" *)
-(* if j < i, then turn around *)
-fun get_compat_thm thy cts i j ctxi ctxj =
-    let
-      val ClauseContext {cqs=cqsi,ags=agsi,lhs=lhsi,...} = ctxi
-      val ClauseContext {cqs=cqsj,ags=agsj,lhs=lhsj,...} = ctxj
-
-      val lhsi_eq_lhsj = cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj)))
-    in if j < i then
-         let
-           val compat = lookup_compat_thm j i cts
-         in
-           compat         (* "!!qj qi. Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
-                |> fold forall_elim (cqsj @ cqsi) (* "Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
-                |> fold Thm.elim_implies agsj
-                |> fold Thm.elim_implies agsi
-                |> Thm.elim_implies ((assume lhsi_eq_lhsj) RS sym) (* "Gsj, Gsi, lhsi = lhsj |-- rhsj = rhsi" *)
-         end
-       else
-         let
-           val compat = lookup_compat_thm i j cts
-         in
-               compat        (* "!!qi qj. Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
-                 |> fold forall_elim (cqsi @ cqsj) (* "Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
-                 |> fold Thm.elim_implies agsi
-                 |> fold Thm.elim_implies agsj
-                 |> Thm.elim_implies (assume lhsi_eq_lhsj)
-                 |> (fn thm => thm RS sym) (* "Gsi, Gsj, lhsi = lhsj |-- rhsj = rhsi" *)
-         end
-    end
-
-
-
-
-(* Generates the replacement lemma in fully quantified form. *)
-fun mk_replacement_lemma thy h ih_elim clause =
-    let
-        val ClauseInfo {cdata=ClauseContext {qs, lhs, rhs, cqs, ags, case_hyp, ...}, RCs, tree, ...} = clause
-        local open Conv in
-        val ih_conv = arg1_conv o arg_conv o arg_conv
-        end
-
-        val ih_elim_case = Conv.fconv_rule (ih_conv (K (case_hyp RS eq_reflection))) ih_elim
-
-        val Ris = map (fn RCInfo {llRI, ...} => llRI) RCs
-        val h_assums = map (fn RCInfo {h_assum, ...} => assume (cterm_of thy (subst_bounds (rev qs, h_assum)))) RCs
-
-        val (eql, _) = FundefCtxTree.rewrite_by_tree thy h ih_elim_case (Ris ~~ h_assums) tree
-
-        val replace_lemma = (eql RS meta_eq_to_obj_eq)
-                                |> implies_intr (cprop_of case_hyp)
-                                |> fold_rev (implies_intr o cprop_of) h_assums
-                                |> fold_rev (implies_intr o cprop_of) ags
-                                |> fold_rev forall_intr cqs
-                                |> Thm.close_derivation
-    in
-      replace_lemma
-    end
-
-
-fun mk_uniqueness_clause thy globals f compat_store clausei clausej RLj =
-    let
-        val Globals {h, y, x, fvar, ...} = globals
-        val ClauseInfo {no=i, cdata=cctxi as ClauseContext {ctxt=ctxti, lhs=lhsi, case_hyp, ...}, ...} = clausei
-        val ClauseInfo {no=j, qglr=cdescj, RCs=RCsj, ...} = clausej
-
-        val cctxj as ClauseContext {ags = agsj', lhs = lhsj', rhs = rhsj', qs = qsj', cqs = cqsj', ...}
-            = mk_clause_context x ctxti cdescj
-
-        val rhsj'h = Pattern.rewrite_term thy [(fvar,h)] [] rhsj'
-        val compat = get_compat_thm thy compat_store i j cctxi cctxj
-        val Ghsj' = map (fn RCInfo {h_assum, ...} => assume (cterm_of thy (subst_bounds (rev qsj', h_assum)))) RCsj
-
-        val RLj_import =
-            RLj |> fold forall_elim cqsj'
-                |> fold Thm.elim_implies agsj'
-                |> fold Thm.elim_implies Ghsj'
-
-        val y_eq_rhsj'h = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (y, rhsj'h))))
-        val lhsi_eq_lhsj' = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj')))) (* lhs_i = lhs_j' |-- lhs_i = lhs_j' *)
-    in
-        (trans OF [case_hyp, lhsi_eq_lhsj']) (* lhs_i = lhs_j' |-- x = lhs_j' *)
-        |> implies_elim RLj_import (* Rj1' ... Rjk', lhs_i = lhs_j' |-- rhs_j'_h = rhs_j'_f *)
-        |> (fn it => trans OF [it, compat]) (* lhs_i = lhs_j', Gj', Rj1' ... Rjk' |-- rhs_j'_h = rhs_i_f *)
-        |> (fn it => trans OF [y_eq_rhsj'h, it]) (* lhs_i = lhs_j', Gj', Rj1' ... Rjk', y = rhs_j_h' |-- y = rhs_i_f *)
-        |> fold_rev (implies_intr o cprop_of) Ghsj'
-        |> fold_rev (implies_intr o cprop_of) agsj' (* lhs_i = lhs_j' , y = rhs_j_h' |-- Gj', Rj1'...Rjk' ==> y = rhs_i_f *)
-        |> implies_intr (cprop_of y_eq_rhsj'h)
-        |> implies_intr (cprop_of lhsi_eq_lhsj')
-        |> fold_rev forall_intr (cterm_of thy h :: cqsj')
-    end
-
-
-
-fun mk_uniqueness_case ctxt thy globals G f ihyp ih_intro G_cases compat_store clauses rep_lemmas clausei =
-    let
-        val Globals {x, y, ranT, fvar, ...} = globals
-        val ClauseInfo {cdata = ClauseContext {lhs, rhs, qs, cqs, ags, case_hyp, ...}, lGI, RCs, ...} = clausei
-        val rhsC = Pattern.rewrite_term thy [(fvar, f)] [] rhs
-
-        val ih_intro_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ih_intro
-
-        fun prep_RC (RCInfo {llRI, RIvs, CCas, ...}) = (llRI RS ih_intro_case)
-                                                            |> fold_rev (implies_intr o cprop_of) CCas
-                                                            |> fold_rev (forall_intr o cterm_of thy o Free) RIvs
-
-        val existence = fold (curry op COMP o prep_RC) RCs lGI
-
-        val P = cterm_of thy (mk_eq (y, rhsC))
-        val G_lhs_y = assume (cterm_of thy (HOLogic.mk_Trueprop (G $ lhs $ y)))
-
-        val unique_clauses = map2 (mk_uniqueness_clause thy globals f compat_store clausei) clauses rep_lemmas
-
-        val uniqueness = G_cases
-                           |> forall_elim (cterm_of thy lhs)
-                           |> forall_elim (cterm_of thy y)
-                           |> forall_elim P
-                           |> Thm.elim_implies G_lhs_y
-                           |> fold Thm.elim_implies unique_clauses
-                           |> implies_intr (cprop_of G_lhs_y)
-                           |> forall_intr (cterm_of thy y)
-
-        val P2 = cterm_of thy (lambda y (G $ lhs $ y)) (* P2 y := (lhs, y): G *)
-
-        val exactly_one =
-            ex1I |> instantiate' [SOME (ctyp_of thy ranT)] [SOME P2, SOME (cterm_of thy rhsC)]
-                 |> curry (op COMP) existence
-                 |> curry (op COMP) uniqueness
-                 |> simplify (HOL_basic_ss addsimps [case_hyp RS sym])
-                 |> implies_intr (cprop_of case_hyp)
-                 |> fold_rev (implies_intr o cprop_of) ags
-                 |> fold_rev forall_intr cqs
-
-        val function_value =
-            existence
-              |> implies_intr ihyp
-              |> implies_intr (cprop_of case_hyp)
-              |> forall_intr (cterm_of thy x)
-              |> forall_elim (cterm_of thy lhs)
-              |> curry (op RS) refl
-    in
-        (exactly_one, function_value)
-    end
-
-
-
-
-fun prove_stuff ctxt globals G f R clauses complete compat compat_store G_elim f_def =
-    let
-        val Globals {h, domT, ranT, x, ...} = globals
-        val thy = ProofContext.theory_of ctxt
-
-        (* Inductive Hypothesis: !!z. (z,x):R ==> EX!y. (z,y):G *)
-        val ihyp = Term.all domT $ Abs ("z", domT,
-                                   Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
-                                     HOLogic.mk_Trueprop (Const ("Ex1", (ranT --> boolT) --> boolT) $
-                                                             Abs ("y", ranT, G $ Bound 1 $ Bound 0))))
-                       |> cterm_of thy
-
-        val ihyp_thm = assume ihyp |> Thm.forall_elim_vars 0
-        val ih_intro = ihyp_thm RS (f_def RS ex1_implies_ex)
-        val ih_elim = ihyp_thm RS (f_def RS ex1_implies_un)
-                        |> instantiate' [] [NONE, SOME (cterm_of thy h)]
-
-        val _ = trace_msg (K "Proving Replacement lemmas...")
-        val repLemmas = map (mk_replacement_lemma thy h ih_elim) clauses
-
-        val _ = trace_msg (K "Proving cases for unique existence...")
-        val (ex1s, values) =
-            split_list (map (mk_uniqueness_case ctxt thy globals G f ihyp ih_intro G_elim compat_store clauses repLemmas) clauses)
-
-        val _ = trace_msg (K "Proving: Graph is a function")
-        val graph_is_function = complete
-                                  |> Thm.forall_elim_vars 0
-                                  |> fold (curry op COMP) ex1s
-                                  |> implies_intr (ihyp)
-                                  |> implies_intr (cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ x)))
-                                  |> forall_intr (cterm_of thy x)
-                                  |> (fn it => Drule.compose_single (it, 2, acc_induct_rule)) (* "EX! y. (?x,y):G" *)
-                                  |> (fn it => fold (forall_intr o cterm_of thy o Var) (Term.add_vars (prop_of it) []) it)
-
-        val goalstate =  Conjunction.intr graph_is_function complete
-                          |> Thm.close_derivation
-                          |> Goal.protect
-                          |> fold_rev (implies_intr o cprop_of) compat
-                          |> implies_intr (cprop_of complete)
-    in
-      (goalstate, values)
-    end
-
-
-fun define_graph Gname fvar domT ranT clauses RCss lthy =
-    let
-      val GT = domT --> ranT --> boolT
-      val Gvar = Free (the_single (Variable.variant_frees lthy [] [(Gname, GT)]))
-
-      fun mk_GIntro (ClauseContext {qs, gs, lhs, rhs, ...}) RCs =
-          let
-            fun mk_h_assm (rcfix, rcassm, rcarg) =
-                HOLogic.mk_Trueprop (Gvar $ rcarg $ (fvar $ rcarg))
-                          |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
-                          |> fold_rev (Logic.all o Free) rcfix
-          in
-            HOLogic.mk_Trueprop (Gvar $ lhs $ rhs)
-                      |> fold_rev (curry Logic.mk_implies o mk_h_assm) RCs
-                      |> fold_rev (curry Logic.mk_implies) gs
-                      |> fold_rev Logic.all (fvar :: qs)
-          end
-
-      val G_intros = map2 mk_GIntro clauses RCss
-
-      val (GIntro_thms, (G, G_elim, G_induct, lthy)) =
-          FundefInductiveWrap.inductive_def G_intros ((dest_Free Gvar, NoSyn), lthy)
-    in
-      ((G, GIntro_thms, G_elim, G_induct), lthy)
-    end
-
-
-
-fun define_function fdefname (fname, mixfix) domT ranT G default lthy =
-    let
-      val f_def =
-          Abs ("x", domT, Const (@{const_name FunDef.THE_default}, ranT --> (ranT --> boolT) --> ranT) $ (default $ Bound 0) $
-                                Abs ("y", ranT, G $ Bound 1 $ Bound 0))
-              |> Syntax.check_term lthy
-
-      val ((f, (_, f_defthm)), lthy) =
-        LocalTheory.define Thm.internalK ((Binding.name (function_name fname), mixfix), ((Binding.name fdefname, []), f_def)) lthy
-    in
-      ((f, f_defthm), lthy)
-    end
-
-
-fun define_recursion_relation Rname domT ranT fvar f qglrs clauses RCss lthy =
-    let
-
-      val RT = domT --> domT --> boolT
-      val Rvar = Free (the_single (Variable.variant_frees lthy [] [(Rname, RT)]))
-
-      fun mk_RIntro (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) (rcfix, rcassm, rcarg) =
-          HOLogic.mk_Trueprop (Rvar $ rcarg $ lhs)
-                    |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
-                    |> fold_rev (curry Logic.mk_implies) gs
-                    |> fold_rev (Logic.all o Free) rcfix
-                    |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
-                    (* "!!qs xs. CS ==> G => (r, lhs) : R" *)
-
-      val R_intross = map2 (map o mk_RIntro) (clauses ~~ qglrs) RCss
-
-      val (RIntro_thmss, (R, R_elim, _, lthy)) =
-          fold_burrow FundefInductiveWrap.inductive_def R_intross ((dest_Free Rvar, NoSyn), lthy)
-    in
-      ((R, RIntro_thmss, R_elim), lthy)
-    end
-
-
-fun fix_globals domT ranT fvar ctxt =
-    let
-      val ([h, y, x, z, a, D, P, Pbool],ctxt') =
-          Variable.variant_fixes ["h_fd", "y_fd", "x_fd", "z_fd", "a_fd", "D_fd", "P_fd", "Pb_fd"] ctxt
-    in
-      (Globals {h = Free (h, domT --> ranT),
-                y = Free (y, ranT),
-                x = Free (x, domT),
-                z = Free (z, domT),
-                a = Free (a, domT),
-                D = Free (D, domT --> boolT),
-                P = Free (P, domT --> boolT),
-                Pbool = Free (Pbool, boolT),
-                fvar = fvar,
-                domT = domT,
-                ranT = ranT
-               },
-       ctxt')
-    end
-
-
-
-fun inst_RC thy fvar f (rcfix, rcassm, rcarg) =
-    let
-      fun inst_term t = subst_bound(f, abstract_over (fvar, t))
-    in
-      (rcfix, map (assume o cterm_of thy o inst_term o prop_of) rcassm, inst_term rcarg)
-    end
-
-
-
-(**********************************************************
- *                   PROVING THE RULES
- **********************************************************)
-
-fun mk_psimps thy globals R clauses valthms f_iff graph_is_function =
-    let
-      val Globals {domT, z, ...} = globals
-
-      fun mk_psimp (ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {cqs, lhs, ags, ...}, ...}) valthm =
-          let
-            val lhs_acc = cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ lhs)) (* "acc R lhs" *)
-            val z_smaller = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ lhs)) (* "R z lhs" *)
-          in
-            ((assume z_smaller) RS ((assume lhs_acc) RS acc_downward))
-              |> (fn it => it COMP graph_is_function)
-              |> implies_intr z_smaller
-              |> forall_intr (cterm_of thy z)
-              |> (fn it => it COMP valthm)
-              |> implies_intr lhs_acc
-              |> asm_simplify (HOL_basic_ss addsimps [f_iff])
-              |> fold_rev (implies_intr o cprop_of) ags
-              |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
-          end
-    in
-      map2 mk_psimp clauses valthms
-    end
-
-
-(** Induction rule **)
-
-
-val acc_subset_induct = @{thm Orderings.predicate1I} RS @{thm accp_subset_induct}
-
-
-fun mk_partial_induct_rule thy globals R complete_thm clauses =
-    let
-      val Globals {domT, x, z, a, P, D, ...} = globals
-      val acc_R = mk_acc domT R
-
-      val x_D = assume (cterm_of thy (HOLogic.mk_Trueprop (D $ x)))
-      val a_D = cterm_of thy (HOLogic.mk_Trueprop (D $ a))
-
-      val D_subset = cterm_of thy (Logic.all x
-        (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x), HOLogic.mk_Trueprop (acc_R $ x))))
-
-      val D_dcl = (* "!!x z. [| x: D; (z,x):R |] ==> z:D" *)
-                    Logic.all x
-                    (Logic.all z (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x),
-                                                    Logic.mk_implies (HOLogic.mk_Trueprop (R $ z $ x),
-                                                                      HOLogic.mk_Trueprop (D $ z)))))
-                    |> cterm_of thy
-
-
-  (* Inductive Hypothesis: !!z. (z,x):R ==> P z *)
-      val ihyp = Term.all domT $ Abs ("z", domT,
-               Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
-                 HOLogic.mk_Trueprop (P $ Bound 0)))
-           |> cterm_of thy
-
-      val aihyp = assume ihyp
-
-  fun prove_case clause =
-      let
-    val ClauseInfo {cdata = ClauseContext {ctxt, qs, cqs, ags, gs, lhs, case_hyp, ...}, RCs,
-                    qglr = (oqs, _, _, _), ...} = clause
-
-    val case_hyp_conv = K (case_hyp RS eq_reflection)
-    local open Conv in
-    val lhs_D = fconv_rule (arg_conv (arg_conv (case_hyp_conv))) x_D
-    val sih = fconv_rule (More_Conv.binder_conv (K (arg1_conv (arg_conv (arg_conv case_hyp_conv)))) ctxt) aihyp
-    end
-
-    fun mk_Prec (RCInfo {llRI, RIvs, CCas, rcarg, ...}) =
-        sih |> forall_elim (cterm_of thy rcarg)
-            |> Thm.elim_implies llRI
-            |> fold_rev (implies_intr o cprop_of) CCas
-            |> fold_rev (forall_intr o cterm_of thy o Free) RIvs
-
-    val P_recs = map mk_Prec RCs   (*  [P rec1, P rec2, ... ]  *)
-
-    val step = HOLogic.mk_Trueprop (P $ lhs)
-            |> fold_rev (curry Logic.mk_implies o prop_of) P_recs
-            |> fold_rev (curry Logic.mk_implies) gs
-            |> curry Logic.mk_implies (HOLogic.mk_Trueprop (D $ lhs))
-            |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
-            |> cterm_of thy
-
-    val P_lhs = assume step
-           |> fold forall_elim cqs
-           |> Thm.elim_implies lhs_D
-           |> fold Thm.elim_implies ags
-           |> fold Thm.elim_implies P_recs
-
-    val res = cterm_of thy (HOLogic.mk_Trueprop (P $ x))
-           |> Conv.arg_conv (Conv.arg_conv case_hyp_conv)
-           |> symmetric (* P lhs == P x *)
-           |> (fn eql => equal_elim eql P_lhs) (* "P x" *)
-           |> implies_intr (cprop_of case_hyp)
-           |> fold_rev (implies_intr o cprop_of) ags
-           |> fold_rev forall_intr cqs
-      in
-        (res, step)
-      end
-
-  val (cases, steps) = split_list (map prove_case clauses)
-
-  val istep = complete_thm
-                |> Thm.forall_elim_vars 0
-                |> fold (curry op COMP) cases (*  P x  *)
-                |> implies_intr ihyp
-                |> implies_intr (cprop_of x_D)
-                |> forall_intr (cterm_of thy x)
-
-  val subset_induct_rule =
-      acc_subset_induct
-        |> (curry op COMP) (assume D_subset)
-        |> (curry op COMP) (assume D_dcl)
-        |> (curry op COMP) (assume a_D)
-        |> (curry op COMP) istep
-        |> fold_rev implies_intr steps
-        |> implies_intr a_D
-        |> implies_intr D_dcl
-        |> implies_intr D_subset
-
-  val subset_induct_all = fold_rev (forall_intr o cterm_of thy) [P, a, D] subset_induct_rule
-
-  val simple_induct_rule =
-      subset_induct_rule
-        |> forall_intr (cterm_of thy D)
-        |> forall_elim (cterm_of thy acc_R)
-        |> assume_tac 1 |> Seq.hd
-        |> (curry op COMP) (acc_downward
-                              |> (instantiate' [SOME (ctyp_of thy domT)]
-                                               (map (SOME o cterm_of thy) [R, x, z]))
-                              |> forall_intr (cterm_of thy z)
-                              |> forall_intr (cterm_of thy x))
-        |> forall_intr (cterm_of thy a)
-        |> forall_intr (cterm_of thy P)
-    in
-      simple_induct_rule
-    end
-
-
-
-(* FIXME: This should probably use fixed goals, to be more reliable and faster *)
-fun mk_domain_intro ctxt (Globals {domT, ...}) R R_cases clause =
-    let
-      val thy = ProofContext.theory_of ctxt
-      val ClauseInfo {cdata = ClauseContext {qs, gs, lhs, rhs, cqs, ...},
-                      qglr = (oqs, _, _, _), ...} = clause
-      val goal = HOLogic.mk_Trueprop (mk_acc domT R $ lhs)
-                          |> fold_rev (curry Logic.mk_implies) gs
-                          |> cterm_of thy
-    in
-      Goal.init goal
-      |> (SINGLE (resolve_tac [accI] 1)) |> the
-      |> (SINGLE (eresolve_tac [Thm.forall_elim_vars 0 R_cases] 1))  |> the
-      |> (SINGLE (auto_tac (clasimpset_of ctxt))) |> the
-      |> Goal.conclude
-      |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
-    end
-
-
-
-(** Termination rule **)
-
-val wf_induct_rule = @{thm Wellfounded.wfP_induct_rule};
-val wf_in_rel = @{thm FunDef.wf_in_rel};
-val in_rel_def = @{thm FunDef.in_rel_def};
-
-fun mk_nest_term_case thy globals R' ihyp clause =
-    let
-      val Globals {x, z, ...} = globals
-      val ClauseInfo {cdata = ClauseContext {qs,cqs,ags,lhs,rhs,case_hyp,...},tree,
-                      qglr=(oqs, _, _, _), ...} = clause
-
-      val ih_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ihyp
-
-      fun step (fixes, assumes) (_ $ arg) u (sub,(hyps,thms)) =
-          let
-            val used = map (fn (ctx,thm) => FundefCtxTree.export_thm thy ctx thm) (u @ sub)
-
-            val hyp = HOLogic.mk_Trueprop (R' $ arg $ lhs)
-                      |> fold_rev (curry Logic.mk_implies o prop_of) used (* additional hyps *)
-                      |> FundefCtxTree.export_term (fixes, assumes)
-                      |> fold_rev (curry Logic.mk_implies o prop_of) ags
-                      |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
-                      |> cterm_of thy
-
-            val thm = assume hyp
-                      |> fold forall_elim cqs
-                      |> fold Thm.elim_implies ags
-                      |> FundefCtxTree.import_thm thy (fixes, assumes)
-                      |> fold Thm.elim_implies used (*  "(arg, lhs) : R'"  *)
-
-            val z_eq_arg = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (z, arg))))
-
-            val acc = thm COMP ih_case
-            val z_acc_local = acc
-            |> Conv.fconv_rule (Conv.arg_conv (Conv.arg_conv (K (symmetric (z_eq_arg RS eq_reflection)))))
-
-            val ethm = z_acc_local
-                         |> FundefCtxTree.export_thm thy (fixes,
-                                                          z_eq_arg :: case_hyp :: ags @ assumes)
-                         |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
-
-            val sub' = sub @ [(([],[]), acc)]
-          in
-            (sub', (hyp :: hyps, ethm :: thms))
-          end
-        | step _ _ _ _ = raise Match
-    in
-      FundefCtxTree.traverse_tree step tree
-    end
-
-
-fun mk_nest_term_rule thy globals R R_cases clauses =
-    let
-      val Globals { domT, x, z, ... } = globals
-      val acc_R = mk_acc domT R
-
-      val R' = Free ("R", fastype_of R)
-
-      val Rrel = Free ("R", HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)))
-      val inrel_R = Const (@{const_name FunDef.in_rel}, HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)) --> fastype_of R) $ Rrel
-
-      val wfR' = cterm_of thy (HOLogic.mk_Trueprop (Const (@{const_name Wellfounded.wfP}, (domT --> domT --> boolT) --> boolT) $ R')) (* "wf R'" *)
-
-      (* Inductive Hypothesis: !!z. (z,x):R' ==> z : acc R *)
-      val ihyp = Term.all domT $ Abs ("z", domT,
-                                 Logic.mk_implies (HOLogic.mk_Trueprop (R' $ Bound 0 $ x),
-                                   HOLogic.mk_Trueprop (acc_R $ Bound 0)))
-                     |> cterm_of thy
-
-      val ihyp_a = assume ihyp |> Thm.forall_elim_vars 0
-
-      val R_z_x = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ x))
-
-      val (hyps,cases) = fold (mk_nest_term_case thy globals R' ihyp_a) clauses ([],[])
-    in
-      R_cases
-        |> forall_elim (cterm_of thy z)
-        |> forall_elim (cterm_of thy x)
-        |> forall_elim (cterm_of thy (acc_R $ z))
-        |> curry op COMP (assume R_z_x)
-        |> fold_rev (curry op COMP) cases
-        |> implies_intr R_z_x
-        |> forall_intr (cterm_of thy z)
-        |> (fn it => it COMP accI)
-        |> implies_intr ihyp
-        |> forall_intr (cterm_of thy x)
-        |> (fn it => Drule.compose_single(it,2,wf_induct_rule))
-        |> curry op RS (assume wfR')
-        |> forall_intr_vars
-        |> (fn it => it COMP allI)
-        |> fold implies_intr hyps
-        |> implies_intr wfR'
-        |> forall_intr (cterm_of thy R')
-        |> forall_elim (cterm_of thy (inrel_R))
-        |> curry op RS wf_in_rel
-        |> full_simplify (HOL_basic_ss addsimps [in_rel_def])
-        |> forall_intr (cterm_of thy Rrel)
-    end
-
-
-
-(* Tail recursion (probably very fragile)
- *
- * FIXME:
- * - Need to do forall_elim_vars on psimps: Unneccesary, if psimps would be taken from the same context.
- * - Must we really replace the fvar by f here?
- * - Splitting is not configured automatically: Problems with case?
- *)
-fun mk_trsimps octxt globals f G R f_def R_cases G_induct clauses psimps =
-    let
-      val Globals {domT, ranT, fvar, ...} = globals
-
-      val R_cases = Thm.forall_elim_vars 0 R_cases (* FIXME: Should be already in standard form. *)
-
-      val graph_implies_dom = (* "G ?x ?y ==> dom ?x"  *)
-          Goal.prove octxt ["x", "y"] [HOLogic.mk_Trueprop (G $ Free ("x", domT) $ Free ("y", ranT))]
-                     (HOLogic.mk_Trueprop (mk_acc domT R $ Free ("x", domT)))
-                     (fn {prems=[a], ...} =>
-                         ((rtac (G_induct OF [a]))
-                            THEN_ALL_NEW (rtac accI)
-                            THEN_ALL_NEW (etac R_cases)
-                            THEN_ALL_NEW (asm_full_simp_tac (simpset_of octxt))) 1)
-
-      val default_thm = (forall_intr_vars graph_implies_dom) COMP (f_def COMP fundef_default_value)
-
-      fun mk_trsimp clause psimp =
-          let
-            val ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {ctxt, cqs, qs, gs, lhs, rhs, ...}, ...} = clause
-            val thy = ProofContext.theory_of ctxt
-            val rhs_f = Pattern.rewrite_term thy [(fvar, f)] [] rhs
-
-            val trsimp = Logic.list_implies(gs, HOLogic.mk_Trueprop (HOLogic.mk_eq(f $ lhs, rhs_f))) (* "f lhs = rhs" *)
-            val lhs_acc = (mk_acc domT R $ lhs) (* "acc R lhs" *)
-            fun simp_default_tac ss = asm_full_simp_tac (ss addsimps [default_thm, Let_def])
-          in
-            Goal.prove ctxt [] [] trsimp
-                       (fn _ =>
-                           rtac (instantiate' [] [SOME (cterm_of thy lhs_acc)] case_split) 1
-                                THEN (rtac (Thm.forall_elim_vars 0 psimp) THEN_ALL_NEW assume_tac) 1
-                                THEN (simp_default_tac (simpset_of ctxt) 1)
-                                THEN (etac not_acc_down 1)
-                                THEN ((etac R_cases) THEN_ALL_NEW (simp_default_tac (simpset_of ctxt))) 1)
-              |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
-          end
-    in
-      map2 mk_trsimp clauses psimps
-    end
-
-
-fun prepare_fundef config defname [((fname, fT), mixfix)] abstract_qglrs lthy =
-    let
-      val FundefConfig {domintros, tailrec, default=default_str, ...} = config
-
-      val fvar = Free (fname, fT)
-      val domT = domain_type fT
-      val ranT = range_type fT
-
-      val default = Syntax.parse_term lthy default_str
-        |> TypeInfer.constrain fT |> Syntax.check_term lthy
-
-      val (globals, ctxt') = fix_globals domT ranT fvar lthy
-
-      val Globals { x, h, ... } = globals
-
-      val clauses = map (mk_clause_context x ctxt') abstract_qglrs
-
-      val n = length abstract_qglrs
-
-      fun build_tree (ClauseContext { ctxt, rhs, ...}) =
-            FundefCtxTree.mk_tree (fname, fT) h ctxt rhs
-
-      val trees = map build_tree clauses
-      val RCss = map find_calls trees
-
-      val ((G, GIntro_thms, G_elim, G_induct), lthy) =
-          PROFILE "def_graph" (define_graph (graph_name defname) fvar domT ranT clauses RCss) lthy
-
-      val ((f, f_defthm), lthy) =
-          PROFILE "def_fun" (define_function (defname ^ "_sumC_def") (fname, mixfix) domT ranT G default) lthy
-
-      val RCss = map (map (inst_RC (ProofContext.theory_of lthy) fvar f)) RCss
-      val trees = map (FundefCtxTree.inst_tree (ProofContext.theory_of lthy) fvar f) trees
-
-      val ((R, RIntro_thmss, R_elim), lthy) =
-          PROFILE "def_rel" (define_recursion_relation (rel_name defname) domT ranT fvar f abstract_qglrs clauses RCss) lthy
-
-      val (_, lthy) =
-          LocalTheory.abbrev Syntax.mode_default ((Binding.name (dom_name defname), NoSyn), mk_acc domT R) lthy
-
-      val newthy = ProofContext.theory_of lthy
-      val clauses = map (transfer_clause_ctx newthy) clauses
-
-      val cert = cterm_of (ProofContext.theory_of lthy)
-
-      val xclauses = PROFILE "xclauses" (map7 (mk_clause_info globals G f) (1 upto n) clauses abstract_qglrs trees RCss GIntro_thms) RIntro_thmss
-
-      val complete = mk_completeness globals clauses abstract_qglrs |> cert |> assume
-      val compat = mk_compat_proof_obligations domT ranT fvar f abstract_qglrs |> map (cert #> assume)
-
-      val compat_store = store_compat_thms n compat
-
-      val (goalstate, values) = PROFILE "prove_stuff" (prove_stuff lthy globals G f R xclauses complete compat compat_store G_elim) f_defthm
-
-      val mk_trsimps = mk_trsimps lthy globals f G R f_defthm R_elim G_induct xclauses
-
-      fun mk_partial_rules provedgoal =
-          let
-            val newthy = theory_of_thm provedgoal (*FIXME*)
-
-            val (graph_is_function, complete_thm) =
-                provedgoal
-                  |> Conjunction.elim
-                  |> apfst (Thm.forall_elim_vars 0)
-
-            val f_iff = graph_is_function RS (f_defthm RS ex1_implies_iff)
-
-            val psimps = PROFILE "Proving simplification rules" (mk_psimps newthy globals R xclauses values f_iff) graph_is_function
-
-            val simple_pinduct = PROFILE "Proving partial induction rule"
-                                                           (mk_partial_induct_rule newthy globals R complete_thm) xclauses
-
-
-            val total_intro = PROFILE "Proving nested termination rule" (mk_nest_term_rule newthy globals R R_elim) xclauses
-
-            val dom_intros = if domintros
-                             then SOME (PROFILE "Proving domain introduction rules" (map (mk_domain_intro lthy globals R R_elim)) xclauses)
-                             else NONE
-            val trsimps = if tailrec then SOME (mk_trsimps psimps) else NONE
-
-          in
-            FundefResult {fs=[f], G=G, R=R, cases=complete_thm,
-                          psimps=psimps, simple_pinducts=[simple_pinduct],
-                          termination=total_intro, trsimps=trsimps,
-                          domintros=dom_intros}
-          end
-    in
-      ((f, goalstate, mk_partial_rules), lthy)
-    end
-
-
-end
--- a/src/HOL/Tools/Function/fundef_datatype.ML	Fri Oct 23 13:23:18 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,178 +0,0 @@
-(*  Title:      HOL/Tools/Function/fundef_datatype.ML
-    Author:     Alexander Krauss, TU Muenchen
-
-A package for general recursive function definitions.
-A tactic to prove completeness of datatype patterns.
-*)
-
-signature FUNDEF_DATATYPE =
-sig
-    val add_fun : FundefCommon.fundef_config ->
-      (binding * typ option * mixfix) list -> (Attrib.binding * term) list ->
-      bool -> local_theory -> Proof.context
-    val add_fun_cmd : FundefCommon.fundef_config ->
-      (binding * string option * mixfix) list -> (Attrib.binding * string) list ->
-      bool -> local_theory -> Proof.context
-
-    val setup : theory -> theory
-end
-
-structure FundefDatatype : FUNDEF_DATATYPE =
-struct
-
-open FundefLib
-open FundefCommon
-
-
-fun check_pats ctxt geq =
-    let 
-      fun err str = error (cat_lines ["Malformed definition:",
-                                      str ^ " not allowed in sequential mode.",
-                                      Syntax.string_of_term ctxt geq])
-      val thy = ProofContext.theory_of ctxt
-                
-      fun check_constr_pattern (Bound _) = ()
-        | check_constr_pattern t =
-          let
-            val (hd, args) = strip_comb t
-          in
-            (((case Datatype.info_of_constr thy (dest_Const hd) of
-                 SOME _ => ()
-               | NONE => err "Non-constructor pattern")
-              handle TERM ("dest_Const", _) => err "Non-constructor patterns");
-             map check_constr_pattern args; 
-             ())
-          end
-          
-      val (fname, qs, gs, args, rhs) = split_def ctxt geq 
-                                       
-      val _ = if not (null gs) then err "Conditional equations" else ()
-      val _ = map check_constr_pattern args
-                  
-                  (* just count occurrences to check linearity *)
-      val _ = if fold (fold_aterms (fn Bound _ => Integer.add 1 | _ => I)) args 0 > length qs
-              then err "Nonlinear patterns" else ()
-    in
-      ()
-    end
-    
-val by_pat_completeness_auto =
-    Proof.global_future_terminal_proof
-      (Method.Basic Pat_Completeness.pat_completeness,
-       SOME (Method.Source_i (Args.src (("HOL.auto", []), Position.none))))
-
-fun termination_by method int =
-    Fundef.termination_proof NONE
-    #> Proof.global_future_terminal_proof (Method.Basic method, NONE) int
-
-fun mk_catchall fixes arity_of =
-    let
-      fun mk_eqn ((fname, fT), _) =
-          let 
-            val n = arity_of fname
-            val (argTs, rT) = chop n (binder_types fT)
-                                   |> apsnd (fn Ts => Ts ---> body_type fT) 
-                              
-            val qs = map Free (Name.invent_list [] "a" n ~~ argTs)
-          in
-            HOLogic.mk_eq(list_comb (Free (fname, fT), qs),
-                          Const ("HOL.undefined", rT))
-              |> HOLogic.mk_Trueprop
-              |> fold_rev Logic.all qs
-          end
-    in
-      map mk_eqn fixes
-    end
-
-fun add_catchall ctxt fixes spec =
-  let val fqgars = map (split_def ctxt) spec
-      val arity_of = map (fn (fname,_,_,args,_) => (fname, length args)) fqgars
-                     |> AList.lookup (op =) #> the
-  in
-    spec @ mk_catchall fixes arity_of
-  end
-
-fun warn_if_redundant ctxt origs tss =
-    let
-        fun msg t = "Ignoring redundant equation: " ^ quote (Syntax.string_of_term ctxt t)
-                    
-        val (tss', _) = chop (length origs) tss
-        fun check (t, []) = (warning (msg t); [])
-          | check (t, s) = s
-    in
-        (map check (origs ~~ tss'); tss)
-    end
-
-
-fun sequential_preproc (config as FundefConfig {sequential, ...}) ctxt fixes spec =
-      if sequential then
-        let
-          val (bnds, eqss) = split_list spec
-                            
-          val eqs = map the_single eqss
-                    
-          val feqs = eqs
-                      |> tap (check_defs ctxt fixes) (* Standard checks *)
-                      |> tap (map (check_pats ctxt))    (* More checks for sequential mode *)
-
-          val compleqs = add_catchall ctxt fixes feqs   (* Completion *)
-
-          val spliteqs = warn_if_redundant ctxt feqs
-                           (FundefSplit.split_all_equations ctxt compleqs)
-
-          fun restore_spec thms =
-              bnds ~~ Library.take (length bnds, Library.unflat spliteqs thms)
-              
-          val spliteqs' = flat (Library.take (length bnds, spliteqs))
-          val fnames = map (fst o fst) fixes
-          val indices = map (fn eq => find_index (curry op = (fname_of eq)) fnames) spliteqs'
-
-          fun sort xs = partition_list (fn i => fn (j,_) => i = j) 0 (length fnames - 1) (indices ~~ xs)
-                                       |> map (map snd)
-
-
-          val bnds' = bnds @ replicate (length spliteqs - length bnds) Attrib.empty_binding
-
-          (* using theorem names for case name currently disabled *)
-          val case_names = map_index (fn (i, (_, es)) => mk_case_names i "" (length es)) 
-                                     (bnds' ~~ spliteqs)
-                           |> flat
-        in
-          (flat spliteqs, restore_spec, sort, case_names)
-        end
-      else
-        FundefCommon.empty_preproc check_defs config ctxt fixes spec
-
-val setup =
-  Context.theory_map (FundefCommon.set_preproc sequential_preproc)
-
-
-val fun_config = FundefConfig { sequential=true, default="%x. undefined" (*FIXME dynamic scoping*), 
-                                domintros=false, tailrec=false }
-
-fun gen_fun add config fixes statements int lthy =
-  let val group = serial_string () in
-    lthy
-      |> LocalTheory.set_group group
-      |> add fixes statements config
-      |> by_pat_completeness_auto int
-      |> LocalTheory.restore
-      |> LocalTheory.set_group group
-      |> termination_by (FundefCommon.get_termination_prover lthy) int
-  end;
-
-val add_fun = gen_fun Fundef.add_fundef
-val add_fun_cmd = gen_fun Fundef.add_fundef_cmd
-
-
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val _ =
-  OuterSyntax.local_theory' "fun" "define general recursive functions (short version)" K.thy_decl
-  (fundef_parser fun_config
-     >> (fn ((config, fixes), statements) => add_fun_cmd config fixes statements));
-
-end
-
-end
--- a/src/HOL/Tools/Function/fundef_lib.ML	Fri Oct 23 13:23:18 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,179 +0,0 @@
-(*  Title:      HOL/Tools/Function/fundef_lib.ML
-    Author:     Alexander Krauss, TU Muenchen
-
-A package for general recursive function definitions. 
-Some fairly general functions that should probably go somewhere else... 
-*)
-
-structure FundefLib = struct
-
-fun map_option f NONE = NONE 
-  | map_option f (SOME x) = SOME (f x);
-
-fun fold_option f NONE y = y
-  | fold_option f (SOME x) y = f x y;
-
-fun fold_map_option f NONE y = (NONE, y)
-  | fold_map_option f (SOME x) y = apfst SOME (f x y);
-
-(* Ex: "The variable" ^ plural " is" "s are" vs *)
-fun plural sg pl [x] = sg
-  | plural sg pl _ = pl
-
-(* lambda-abstracts over an arbitrarily nested tuple
-  ==> hologic.ML? *)
-fun tupled_lambda vars t =
-    case vars of
-      (Free v) => lambda (Free v) t
-    | (Var v) => lambda (Var v) t
-    | (Const ("Pair", Type ("fun", [Ta, Type ("fun", [Tb, _])]))) $ us $ vs =>  
-      (HOLogic.split_const (Ta,Tb, fastype_of t)) $ (tupled_lambda us (tupled_lambda vs t))
-    | _ => raise Match
-                 
-                 
-fun dest_all (Const ("all", _) $ Abs (a as (_,T,_))) =
-    let
-      val (n, body) = Term.dest_abs a
-    in
-      (Free (n, T), body)
-    end
-  | dest_all _ = raise Match
-                         
-
-(* Removes all quantifiers from a term, replacing bound variables by frees. *)
-fun dest_all_all (t as (Const ("all",_) $ _)) = 
-    let
-      val (v,b) = dest_all t
-      val (vs, b') = dest_all_all b
-    in
-      (v :: vs, b')
-    end
-  | dest_all_all t = ([],t)
-                     
-
-(* FIXME: similar to Variable.focus *)
-fun dest_all_all_ctx ctx (Const ("all", _) $ Abs (a as (n,T,b))) =
-    let
-      val [(n', _)] = Variable.variant_frees ctx [] [(n,T)]
-      val (_, ctx') = ProofContext.add_fixes [(Binding.name n', SOME T, NoSyn)] ctx
-
-      val (n'', body) = Term.dest_abs (n', T, b) 
-      val _ = (n' = n'') orelse error "dest_all_ctx"
-      (* Note: We assume that n' does not occur in the body. Otherwise it would be fixed. *)
-
-      val (ctx'', vs, bd) = dest_all_all_ctx ctx' body
-    in
-      (ctx'', (n', T) :: vs, bd)
-    end
-  | dest_all_all_ctx ctx t = 
-    (ctx, [], t)
-
-
-fun map3 _ [] [] [] = []
-  | map3 f (x :: xs) (y :: ys) (z :: zs) = f x y z :: map3 f xs ys zs
-  | map3 _ _ _ _ = raise Library.UnequalLengths;
-
-fun map4 _ [] [] [] [] = []
-  | map4 f (x :: xs) (y :: ys) (z :: zs) (u :: us) = f x y z u :: map4 f xs ys zs us
-  | map4 _ _ _ _ _ = raise Library.UnequalLengths;
-
-fun map6 _ [] [] [] [] [] [] = []
-  | map6 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (w :: ws) = f x y z u v w :: map6 f xs ys zs us vs ws
-  | map6 _ _ _ _ _ _ _ = raise Library.UnequalLengths;
-
-fun map7 _ [] [] [] [] [] [] [] = []
-  | map7 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (w :: ws) (b :: bs) = f x y z u v w b :: map7 f xs ys zs us vs ws bs
-  | map7 _ _ _ _ _ _ _ _ = raise Library.UnequalLengths;
-
-
-
-(* forms all "unordered pairs": [1, 2, 3] ==> [(1, 1), (1, 2), (1, 3), (2, 2), (2, 3), (3, 3)] *)
-(* ==> library *)
-fun unordered_pairs [] = []
-  | unordered_pairs (x::xs) = map (pair x) (x::xs) @ unordered_pairs xs
-
-
-(* Replaces Frees by name. Works with loose Bounds. *)
-fun replace_frees assoc =
-    map_aterms (fn c as Free (n, _) => the_default c (AList.lookup (op =) assoc n)
-                 | t => t)
-
-
-fun rename_bound n (Q $ Abs(_, T, b)) = (Q $ Abs(n, T, b))
-  | rename_bound n _ = raise Match
-
-fun mk_forall_rename (n, v) =
-    rename_bound n o Logic.all v 
-
-fun forall_intr_rename (n, cv) thm =
-    let
-      val allthm = forall_intr cv thm
-      val (_ $ abs) = prop_of allthm
-    in
-      Thm.rename_boundvars abs (Abs (n, dummyT, Term.dummy_pattern dummyT)) allthm
-    end
-
-
-(* Returns the frees in a term in canonical order, excluding the fixes from the context *)
-fun frees_in_term ctxt t =
-    Term.add_frees t []
-    |> filter_out (Variable.is_fixed ctxt o fst)
-    |> rev
-
-
-datatype proof_attempt = Solved of thm | Stuck of thm | Fail
-
-fun try_proof cgoal tac = 
-    case SINGLE tac (Goal.init cgoal) of
-      NONE => Fail
-    | SOME st =>
-        if Thm.no_prems st
-        then Solved (Goal.finish (Syntax.init_pretty_global (Thm.theory_of_cterm cgoal)) st)
-        else Stuck st
-
-
-fun dest_binop_list cn (t as (Const (n, _) $ a $ b)) = 
-    if cn = n then dest_binop_list cn a @ dest_binop_list cn b else [ t ]
-  | dest_binop_list _ t = [ t ]
-
-
-(* separate two parts in a +-expression:
-   "a + b + c + d + e" --> "(a + b + d) + (c + e)"
-
-   Here, + can be any binary operation that is AC.
-
-   cn - The name of the binop-constructor (e.g. @{const_name Un})
-   ac - the AC rewrite rules for cn
-   is - the list of indices of the expressions that should become the first part
-        (e.g. [0,1,3] in the above example)
-*)
-
-fun regroup_conv neu cn ac is ct =
- let
-   val mk = HOLogic.mk_binop cn
-   val t = term_of ct
-   val xs = dest_binop_list cn t
-   val js = subtract (op =) is (0 upto (length xs) - 1)
-   val ty = fastype_of t
-   val thy = theory_of_cterm ct
- in
-   Goal.prove_internal []
-     (cterm_of thy
-       (Logic.mk_equals (t,
-          if is = []
-          then mk (Const (neu, ty), foldr1 mk (map (nth xs) js))
-          else if js = []
-            then mk (foldr1 mk (map (nth xs) is), Const (neu, ty))
-            else mk (foldr1 mk (map (nth xs) is), foldr1 mk (map (nth xs) js)))))
-     (K (rewrite_goals_tac ac
-         THEN rtac Drule.reflexive_thm 1))
- end
-
-(* instance for unions *)
-fun regroup_union_conv t = regroup_conv @{const_name Set.empty} @{const_name Lattices.sup}
-  (map (fn t => t RS eq_reflection) (@{thms Un_ac} @
-                                     @{thms Un_empty_right} @
-                                     @{thms Un_empty_left})) t
-
-
-end
--- a/src/HOL/Tools/Function/induction_scheme.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/Function/induction_scheme.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -13,10 +13,10 @@
 end
 
 
-structure InductionScheme : INDUCTION_SCHEME =
+structure Induction_Scheme : INDUCTION_SCHEME =
 struct
 
-open FundefLib
+open Function_Lib
 
 
 type rec_call_info = int * (string * typ) list * term list * term list
--- a/src/HOL/Tools/Function/inductive_wrap.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/Function/inductive_wrap.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -6,17 +6,17 @@
 the introduction and elimination rules.
 *)
 
-signature FUNDEF_INDUCTIVE_WRAP =
+signature FUNCTION_INDUCTIVE_WRAP =
 sig
   val inductive_def :  term list 
                        -> ((bstring * typ) * mixfix) * local_theory
                        -> thm list * (term * thm * thm * local_theory)
 end
 
-structure FundefInductiveWrap: FUNDEF_INDUCTIVE_WRAP =
+structure Function_Inductive_Wrap: FUNCTION_INDUCTIVE_WRAP =
 struct
 
-open FundefLib
+open Function_Lib
 
 fun requantify ctxt lfix orig_def thm =
     let
--- a/src/HOL/Tools/Function/lexicographic_order.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/Function/lexicographic_order.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -13,10 +13,10 @@
   val setup: theory -> theory
 end
 
-structure LexicographicOrder : LEXICOGRAPHIC_ORDER =
+structure Lexicographic_Order : LEXICOGRAPHIC_ORDER =
 struct
 
-open FundefLib
+open Function_Lib
 
 (** General stuff **)
 
@@ -58,7 +58,7 @@
 
 fun dest_term (t : term) =
     let
-      val (vars, prop) = FundefLib.dest_all_all t
+      val (vars, prop) = Function_Lib.dest_all_all t
       val (prems, concl) = Logic.strip_horn prop
       val (lhs, rhs) = concl
                          |> HOLogic.dest_Trueprop
@@ -215,9 +215,9 @@
     end
 
 fun lexicographic_order_tac ctxt =
-  TRY (FundefCommon.apply_termination_rule ctxt 1)
+  TRY (Function_Common.apply_termination_rule ctxt 1)
   THEN lex_order_tac ctxt
-    (auto_tac (clasimpset_of ctxt addsimps2 FundefCommon.Termination_Simps.get ctxt))
+    (auto_tac (clasimpset_of ctxt addsimps2 Function_Common.Termination_Simps.get ctxt))
 
 val lexicographic_order = SIMPLE_METHOD o lexicographic_order_tac
 
@@ -225,7 +225,7 @@
   Method.setup @{binding lexicographic_order}
     (Method.sections clasimp_modifiers >> (K lexicographic_order))
     "termination prover for lexicographic orderings"
-  #> Context.theory_map (FundefCommon.set_termination_prover lexicographic_order)
+  #> Context.theory_map (Function_Common.set_termination_prover lexicographic_order)
 
 end;
 
--- a/src/HOL/Tools/Function/mutual.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/Function/mutual.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -5,29 +5,26 @@
 Tools for mutual recursive definitions.
 *)
 
-signature FUNDEF_MUTUAL =
+signature FUNCTION_MUTUAL =
 sig
 
-  val prepare_fundef_mutual : FundefCommon.fundef_config
+  val prepare_function_mutual : Function_Common.function_config
                               -> string (* defname *)
                               -> ((string * typ) * mixfix) list
                               -> term list
                               -> local_theory
                               -> ((thm (* goalstate *)
-                                   * (thm -> FundefCommon.fundef_result) (* proof continuation *)
+                                   * (thm -> Function_Common.function_result) (* proof continuation *)
                                   ) * local_theory)
 
 end
 
 
-structure FundefMutual: FUNDEF_MUTUAL =
+structure Function_Mutual: FUNCTION_MUTUAL =
 struct
 
-open FundefLib
-open FundefCommon
-
-
-
+open Function_Lib
+open Function_Common
 
 type qgar = string * (string * typ) list * term list * term list * term
 
@@ -268,7 +265,7 @@
 fun mk_partial_rules_mutual lthy inner_cont (m as Mutual {parts, fqgars, ...}) proof =
     let
       val result = inner_cont proof
-      val FundefResult {fs=[f], G, R, cases, psimps, trsimps, simple_pinducts=[simple_pinduct],
+      val FunctionResult {fs=[f], G, R, cases, psimps, trsimps, simple_pinducts=[simple_pinduct],
                         termination,domintros} = result
                                                                                                                
       val (all_f_defs, fs) = map (fn MutualPart {f_defthm = SOME f_def, f = SOME f, cargTs, ...} =>
@@ -288,20 +285,20 @@
       val mtermination = full_simplify rew_ss termination
       val mdomintros = map_option (map (full_simplify rew_ss)) domintros
     in
-      FundefResult { fs=fs, G=G, R=R,
+      FunctionResult { fs=fs, G=G, R=R,
                      psimps=mpsimps, simple_pinducts=minducts,
                      cases=cases, termination=mtermination,
                      domintros=mdomintros,
                      trsimps=mtrsimps}
     end
       
-fun prepare_fundef_mutual config defname fixes eqss lthy =
+fun prepare_function_mutual config defname fixes eqss lthy =
     let
       val mutual = analyze_eqs lthy defname (map fst fixes) (map Envir.beta_eta_contract eqss)
       val Mutual {fsum_var=(n, T), qglrs, ...} = mutual
           
       val ((fsum, goalstate, cont), lthy') =
-          FundefCore.prepare_fundef config defname [((n, T), NoSyn)] qglrs lthy
+          Function_Core.prepare_function config defname [((n, T), NoSyn)] qglrs lthy
           
       val (mutual', lthy'') = define_projections fixes mutual fsum lthy'
 
--- a/src/HOL/Tools/Function/pat_completeness.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/Function/pat_completeness.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -17,8 +17,8 @@
 structure Pat_Completeness : PAT_COMPLETENESS =
 struct
 
-open FundefLib
-open FundefCommon
+open Function_Lib
+open Function_Common
 
 
 fun mk_argvar i T = Free ("_av" ^ (string_of_int i), T)
--- a/src/HOL/Tools/Function/pattern_split.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/Function/pattern_split.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -8,7 +8,7 @@
 
 *)
 
-signature FUNDEF_SPLIT =
+signature FUNCTION_SPLIT =
 sig
   val split_some_equations :
       Proof.context -> (bool * term) list -> term list list
@@ -17,10 +17,10 @@
       Proof.context -> term list -> term list list
 end
 
-structure FundefSplit : FUNDEF_SPLIT =
+structure Function_Split : FUNCTION_SPLIT =
 struct
 
-open FundefLib
+open Function_Lib
 
 (* We use proof context for the variable management *)
 (* FIXME: no __ *)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/relation.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -0,0 +1,36 @@
+(*  Title:      HOL/Tools/Function/relation.ML
+    Author:     Alexander Krauss, TU Muenchen
+
+A package for general recursive function definitions.
+Method "relation" to commence a termination proof using a user-specified relation.
+*)
+
+signature FUNCTION_RELATION =
+sig
+  val relation_tac: Proof.context -> term -> int -> tactic
+  val setup: theory -> theory
+end
+
+structure Function_Relation : FUNCTION_RELATION =
+struct
+
+fun inst_thm ctxt rel st =
+    let
+      val cert = Thm.cterm_of (ProofContext.theory_of ctxt)
+      val rel' = cert (singleton (Variable.polymorphic ctxt) rel)
+      val st' = Thm.incr_indexes (#maxidx (Thm.rep_cterm rel') + 1) st
+      val Rvar = cert (Var (the_single (Term.add_vars (prop_of st') [])))
+    in 
+      Drule.cterm_instantiate [(Rvar, rel')] st' 
+    end
+
+fun relation_tac ctxt rel i = 
+    TRY (Function_Common.apply_termination_rule ctxt i)
+    THEN PRIMITIVE (inst_thm ctxt rel)
+
+val setup =
+  Method.setup @{binding relation}
+    (Args.term >> (fn rel => fn ctxt => SIMPLE_METHOD' (relation_tac ctxt rel)))
+    "proves termination using a user-specified wellfounded relation"
+
+end
--- a/src/HOL/Tools/Function/scnp_reconstruct.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/Function/scnp_reconstruct.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -38,8 +38,8 @@
 structure ScnpReconstruct : SCNP_RECONSTRUCT =
 struct
 
-val PROFILE = FundefCommon.PROFILE
-fun TRACE x = if ! FundefCommon.profile then tracing x else ()
+val PROFILE = Function_Common.PROFILE
+fun TRACE x = if ! Function_Common.profile then tracing x else ()
 
 open ScnpSolve
 
@@ -64,7 +64,7 @@
    reduction_pair : thm
   }
 
-structure MultisetSetup = TheoryDataFun
+structure Multiset_Setup = TheoryDataFun
 (
   type T = multiset_setup option
   val empty = NONE
@@ -73,10 +73,10 @@
   fun merge _ (v1, v2) = if is_some v2 then v2 else v1
 )
 
-val multiset_setup = MultisetSetup.put o SOME
+val multiset_setup = Multiset_Setup.put o SOME
 
 fun undef x = error "undef"
-fun get_multiset_setup thy = MultisetSetup.get thy
+fun get_multiset_setup thy = Multiset_Setup.get thy
   |> the_default (Multiset
 { msetT = undef, mk_mset=undef,
   mset_regroup_conv=undef, mset_member_tac = undef,
@@ -287,7 +287,7 @@
         |> cterm_of thy
     in
       PROFILE "Proof Reconstruction"
-        (CONVERSION (Conv.arg_conv (Conv.arg_conv (FundefLib.regroup_union_conv sl))) 1
+        (CONVERSION (Conv.arg_conv (Conv.arg_conv (Function_Lib.regroup_union_conv sl))) 1
          THEN (rtac @{thm reduction_pair_lemma} 1)
          THEN (rtac @{thm rp_inv_image_rp} 1)
          THEN (rtac (order_rpair ms_rp label) 1)
@@ -350,7 +350,7 @@
 
 fun single_scnp_tac use_tags orders ctxt cont err_cont D = Termination.CALLS (fn (cs, i) =>
   let
-    val ms_configured = is_some (MultisetSetup.get (ProofContext.theory_of ctxt))
+    val ms_configured = is_some (Multiset_Setup.get (ProofContext.theory_of ctxt))
     val orders' = if ms_configured then orders
                   else filter_out (curry op = MS) orders
     val gp = gen_probl D cs
@@ -395,7 +395,7 @@
   end
 
 fun gen_sizechange_tac orders autom_tac ctxt err_cont =
-  TRY (FundefCommon.apply_termination_rule ctxt 1)
+  TRY (Function_Common.apply_termination_rule ctxt 1)
   THEN TRY (Termination.wf_union_tac ctxt)
   THEN
    (rtac @{thm wf_empty} 1
@@ -406,7 +406,7 @@
 
 fun decomp_scnp orders ctxt =
   let
-    val extra_simps = FundefCommon.Termination_Simps.get ctxt
+    val extra_simps = Function_Common.Termination_Simps.get ctxt
     val autom_tac = auto_tac (clasimpset_of ctxt addsimps2 extra_simps)
   in
     SIMPLE_METHOD
--- a/src/HOL/Tools/Function/scnp_solve.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/Function/scnp_solve.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -73,7 +73,7 @@
 (* SAT solving *)
 val solver = Unsynchronized.ref "auto";
 fun sat_solver x =
-  FundefCommon.PROFILE "sat_solving..." (SatSolver.invoke_solver (!solver)) x
+  Function_Common.PROFILE "sat_solving..." (SatSolver.invoke_solver (!solver)) x
 
 (* "Virtual constructors" for various propositional variables *)
 fun var_constrs (gp as GP (arities, gl)) =
--- a/src/HOL/Tools/Function/termination.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/Function/termination.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -48,7 +48,7 @@
 structure Termination : TERMINATION =
 struct
 
-open FundefLib
+open Function_Lib
 
 val term2_ord = prod_ord TermOrd.fast_term_ord TermOrd.fast_term_ord
 structure Term2tab = Table(type key = term * term val ord = term2_ord);
@@ -145,7 +145,7 @@
 
 fun mk_sum_skel rel =
   let
-    val cs = FundefLib.dest_binop_list @{const_name Lattices.sup} rel
+    val cs = Function_Lib.dest_binop_list @{const_name Lattices.sup} rel
     fun collect_pats (Const (@{const_name Collect}, _) $ Abs (_, _, c)) =
       let
         val (Const ("op &", _) $ (Const ("op =", _) $ _ $ (Const ("Pair", _) $ r $ l)) $ Gam)
@@ -233,7 +233,7 @@
 fun CALLS tac i st =
   if Thm.no_prems st then all_tac st
   else case Thm.term_of (Thm.cprem_of st i) of
-    (_ $ (_ $ rel)) => tac (FundefLib.dest_binop_list @{const_name Lattices.sup} rel, i) st
+    (_ $ (_ $ rel)) => tac (Function_Lib.dest_binop_list @{const_name Lattices.sup} rel, i) st
   |_ => no_tac st
 
 type ttac = (data -> int -> tactic) -> (data -> int -> tactic) -> data -> int -> tactic
@@ -251,7 +251,7 @@
 local
 fun dest_term (t : term) = (* FIXME, cf. Lexicographic order *)
     let
-      val (vars, prop) = FundefLib.dest_all_all t
+      val (vars, prop) = Function_Lib.dest_all_all t
       val (prems, concl) = Logic.strip_horn prop
       val (lhs, rhs) = concl
                          |> HOLogic.dest_Trueprop
--- a/src/HOL/Tools/Predicate_Compile/pred_compile_aux.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/Predicate_Compile/pred_compile_aux.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -6,8 +6,16 @@
 structure Predicate_Compile_Aux =
 struct
 
+(* general syntactic functions *)
+
+(*Like dest_conj, but flattens conjunctions however nested*)
+fun conjuncts_aux (Const ("op &", _) $ t $ t') conjs = conjuncts_aux t (conjuncts_aux t' conjs)
+  | conjuncts_aux t conjs = t::conjs;
+
+fun conjuncts t = conjuncts_aux t [];
+
 (* syntactic functions *)
- 
+
 fun is_equationlike_term (Const ("==", _) $ _ $ _) = true
   | is_equationlike_term (Const ("Trueprop", _) $ (Const ("op =", _) $ _ $ _)) = true
   | is_equationlike_term _ = false
@@ -69,26 +77,52 @@
   in ((ps', t''), nctxt') end;
 
 
+(* introduction rule combinators *)
 
+(* combinators to apply a function to all literals of an introduction rules *)
 
-(*
 fun map_atoms f intro = 
-fun fold_atoms f intro =
-*)
+  let
+    val (literals, head) = Logic.strip_horn intro
+    fun appl t = (case t of
+        (@{term "Not"} $ t') => HOLogic.mk_not (f t')
+      | _ => f t)
+  in
+    Logic.list_implies
+      (map (HOLogic.mk_Trueprop o appl o HOLogic.dest_Trueprop) literals, head)
+  end
+
+fun fold_atoms f intro s =
+  let
+    val (literals, head) = Logic.strip_horn intro
+    fun appl t s = (case t of
+      (@{term "Not"} $ t') => f t' s
+      | _ => f t s)
+  in fold appl (map HOLogic.dest_Trueprop literals) s end
+
 fun fold_map_atoms f intro s =
   let
     val (literals, head) = Logic.strip_horn intro
     fun appl t s = (case t of
-      (@{term "Not"} $ t') =>
-        let
-          val (t'', s') = f t' s
-        in (@{term "Not"} $ t'', s') end
+      (@{term "Not"} $ t') => apfst HOLogic.mk_not (f t' s)
       | _ => f t s)
     val (literals', s') = fold_map appl (map HOLogic.dest_Trueprop literals) s
   in
     (Logic.list_implies (map HOLogic.mk_Trueprop literals', head), s')
   end;
+
+fun maps_premises f intro =
+  let
+    val (premises, head) = Logic.strip_horn intro
+  in
+    Logic.list_implies (maps f premises, head)
+  end
   
+(* lifting term operations to theorems *)
+
+fun map_term thy f th =
+  Skip_Proof.make_thm thy (f (prop_of th))
+
 (*
 fun equals_conv lhs_cv rhs_cv ct =
   case Thm.term_of ct of
@@ -96,5 +130,107 @@
   | _ => error "equals_conv"  
 *)
 
+(* Different options for compiler *)
+
+datatype options = Options of {  
+  expected_modes : (string * int list list) option,
+  show_steps : bool,
+  show_mode_inference : bool,
+  show_proof_trace : bool,
+  show_intermediate_results : bool,
+  show_compilation : bool,
+  skip_proof : bool,
+
+  inductify : bool,
+  rpred : bool,
+  depth_limited : bool
+};
+
+fun expected_modes (Options opt) = #expected_modes opt
+fun show_steps (Options opt) = #show_steps opt
+fun show_mode_inference (Options opt) = #show_mode_inference opt
+fun show_intermediate_results (Options opt) = #show_intermediate_results opt
+fun show_proof_trace (Options opt) = #show_proof_trace opt
+fun show_compilation (Options opt) = #show_compilation opt
+fun skip_proof (Options opt) = #skip_proof opt
+
+fun is_inductify (Options opt) = #inductify opt
+fun is_rpred (Options opt) = #rpred opt
+fun is_depth_limited (Options opt) = #depth_limited opt
+
+val default_options = Options {
+  expected_modes = NONE,
+  show_steps = false,
+  show_intermediate_results = false,
+  show_proof_trace = false,
+  show_mode_inference = false,
+  show_compilation = false,
+  skip_proof = false,
+  
+  inductify = false,
+  rpred = false,
+  depth_limited = false
+}
+
+
+fun print_step options s =
+  if show_steps options then tracing s else ()
+
+(* tuple processing *)
+
+fun expand_tuples thy intro =
+  let
+    fun rewrite_args [] (pats, intro_t, ctxt) = (pats, intro_t, ctxt)
+      | rewrite_args (arg::args) (pats, intro_t, ctxt) = 
+      (case HOLogic.strip_tupleT (fastype_of arg) of
+        (Ts as _ :: _ :: _) =>
+        let
+          fun rewrite_arg' (Const ("Pair", _) $ _ $ t2, Type ("*", [_, T2]))
+            (args, (pats, intro_t, ctxt)) = rewrite_arg' (t2, T2) (args, (pats, intro_t, ctxt))
+            | rewrite_arg' (t, Type ("*", [T1, T2])) (args, (pats, intro_t, ctxt)) =
+              let
+                val ([x, y], ctxt') = Variable.variant_fixes ["x", "y"] ctxt
+                val pat = (t, HOLogic.mk_prod (Free (x, T1), Free (y, T2)))
+                val intro_t' = Pattern.rewrite_term thy [pat] [] intro_t
+                val args' = map (Pattern.rewrite_term thy [pat] []) args
+              in
+                rewrite_arg' (Free (y, T2), T2) (args', (pat::pats, intro_t', ctxt'))
+              end
+            | rewrite_arg' _ (args, (pats, intro_t, ctxt)) = (args, (pats, intro_t, ctxt))
+          val (args', (pats, intro_t', ctxt')) = rewrite_arg' (arg, fastype_of arg)
+            (args, (pats, intro_t, ctxt))
+        in
+          rewrite_args args' (pats, intro_t', ctxt')
+        end
+      | _ => rewrite_args args (pats, intro_t, ctxt))
+    fun rewrite_prem atom =
+      let
+        val (_, args) = strip_comb atom
+      in rewrite_args args end
+    val ctxt = ProofContext.init thy
+    val (((T_insts, t_insts), [intro']), ctxt1) = Variable.import false [intro] ctxt
+    val intro_t = prop_of intro'
+    val concl = Logic.strip_imp_concl intro_t
+    val (p, args) = strip_comb (HOLogic.dest_Trueprop concl)
+    val (pats', intro_t', ctxt2) = rewrite_args args ([], intro_t, ctxt1)
+    val (pats', intro_t', ctxt3) = 
+      fold_atoms rewrite_prem intro_t' (pats', intro_t', ctxt2)
+    fun rewrite_pat (ct1, ct2) =
+      (ct1, cterm_of thy (Pattern.rewrite_term thy pats' [] (term_of ct2)))
+    val t_insts' = map rewrite_pat t_insts
+    val intro'' = Thm.instantiate (T_insts, t_insts') intro
+    val [intro'''] = Variable.export ctxt3 ctxt [intro'']
+    val intro'''' = Simplifier.full_simplify
+      (HOL_basic_ss addsimps [@{thm fst_conv}, @{thm snd_conv}, @{thm Pair_eq}])
+      intro'''
+    (* splitting conjunctions introduced by Pair_eq*)
+    fun split_conj prem =
+      map HOLogic.mk_Trueprop (conjuncts (HOLogic.dest_Trueprop prem))
+    val intro''''' = map_term thy (maps_premises split_conj) intro''''
+  in
+    intro'''''
+  end
+
+
 
 end;
--- a/src/HOL/Tools/Predicate_Compile/pred_compile_data.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/Predicate_Compile/pred_compile_data.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -3,16 +3,16 @@
 Book-keeping datastructure for the predicate compiler
 
 *)
-signature PRED_COMPILE_DATA =
+signature PREDICATE_COMPILE_DATA =
 sig
   type specification_table;
   val make_const_spec_table : theory -> specification_table
   val get_specification :  specification_table -> string -> thm list
-  val obtain_specification_graph : specification_table -> string -> thm list Graph.T
+  val obtain_specification_graph : theory -> specification_table -> string -> thm list Graph.T
   val normalize_equation : theory -> thm -> thm
 end;
 
-structure Pred_Compile_Data : PRED_COMPILE_DATA =
+structure Predicate_Compile_Data : PREDICATE_COMPILE_DATA =
 struct
 
 open Predicate_Compile_Aux;
@@ -81,11 +81,13 @@
     Const ("Trueprop", _) $ (Const ("op =", _) $ _ $ _) => th RS @{thm eq_reflection}
   | _ => th
 
+val meta_fun_cong = @{lemma "f == g ==> f x == g x" by simp}
+
 fun full_fun_cong_expand th =
   let
     val (f, args) = strip_comb (fst (Logic.dest_equals (prop_of th)))
     val i = length (binder_types (fastype_of f)) - length args
-  in funpow i (fn th => th RS @{thm meta_fun_cong}) th end;
+  in funpow i (fn th => th RS meta_fun_cong) th end;
 
 fun declare_names s xs ctxt =
   let
@@ -117,7 +119,7 @@
 
 fun normalize_equation thy th =
   mk_meta_equation th
-  |> Pred_Compile_Set.unfold_set_notation
+  |> Predicate_Compile_Set.unfold_set_notation
   |> full_fun_cong_expand
   |> split_all_pairs thy
   |> tap check_equation_format
@@ -127,19 +129,18 @@
 
 fun store_thm_in_table ignore_consts thy th=
   let
-    val th = AxClass.unoverload thy th
+    val th = th
       |> inline_equations thy
+      |> Predicate_Compile_Set.unfold_set_notation
+      |> AxClass.unoverload thy
     val (const, th) =
       if is_equationlike th then
         let
-          val _ = priority "Normalizing definition..."
           val eq = normalize_equation thy th
         in
           (defining_const_of_equation eq, eq)
         end
-      else if (is_introlike th) then
-        let val th = Pred_Compile_Set.unfold_set_notation th
-        in (defining_const_of_introrule th, th) end
+      else if (is_introlike th) then (defining_const_of_introrule th, th)
       else error "store_thm: unexpected definition format"
   in
     if not (member (op =) ignore_consts const) then
@@ -147,18 +148,6 @@
     else I
   end
 
-(*
-fun make_const_spec_table_warning thy =
-  fold
-    (fn th => fn thy => case try (store_thm th) thy of
-      SOME thy => thy
-    | NONE => (warning ("store_thm fails for " ^ Display.string_of_thm_global thy th) ; thy))
-      (Predicate_Compile_Preproc_Const_Defs.get (ProofContext.init thy)) thy
-
-fun make_const_spec_table thy =
-  fold store_thm (Predicate_Compile_Preproc_Const_Defs.get (ProofContext.init thy)) thy
-  |> (fn thy => fold store_thm (Nitpick_Simps.get (ProofContext.init thy)) thy)
-*)
 fun make_const_spec_table thy =
   let
     fun store ignore_const f = fold (store_thm_in_table ignore_const thy) (map (Thm.transfer thy) (f (ProofContext.init thy)))
@@ -166,35 +155,27 @@
       |> store [] Predicate_Compile_Alternative_Defs.get
     val ignore_consts = Symtab.keys table
   in
-    table   
+    table
     |> store ignore_consts Predicate_Compile_Preproc_Const_Defs.get
     |> store ignore_consts Nitpick_Simps.get
     |> store ignore_consts Nitpick_Intros.get
   end
-  (*
-fun get_specification thy constname =
-  case Symtab.lookup (#const_spec_table (Data.get thy)) constname of
-    SOME thms => thms
-  | NONE => error ("get_specification: lookup of constant " ^ quote constname ^ " failed")
-  *)
+
 fun get_specification table constname =
   case Symtab.lookup table constname of
-  SOME thms =>
-    let
-      val _ = tracing ("Looking up specification of " ^ constname ^ ": "
-        ^ (commas (map Display.string_of_thm_without_context thms)))
-    in thms end
+    SOME thms => thms
   | NONE => error ("get_specification: lookup of constant " ^ quote constname ^ " failed")
 
 val logic_operator_names =
-  [@{const_name "=="}, @{const_name "op ="}, @{const_name "op -->"}, @{const_name "All"}, @{const_name "op &"}]
+  [@{const_name "=="}, @{const_name "op ="}, @{const_name "op -->"}, @{const_name "All"}, @{const_name "Ex"}, 
+   @{const_name "op &"}]
 
-val special_cases = member (op =) [@{const_name "Suc"}, @{const_name Nat.zero_nat_inst.zero_nat},
+val special_cases = member (op =) [
+    @{const_name "False"},
+    @{const_name "Suc"}, @{const_name Nat.zero_nat_inst.zero_nat},
     @{const_name Nat.one_nat_inst.one_nat},
 @{const_name "HOL.ord_class.less"}, @{const_name "HOL.ord_class.less_eq"}, @{const_name "HOL.zero_class.zero"},
 @{const_name "HOL.one_class.one"},  @{const_name HOL.plus_class.plus},
-@{const_name "Nat.nat.nat_case"}, @{const_name "List.list.list_case"},
-@{const_name "Option.option.option_case"},
 @{const_name Nat.ord_nat_inst.less_eq_nat},
 @{const_name number_nat_inst.number_of_nat},
   @{const_name Int.Bit0},
@@ -203,13 +184,19 @@
 @{const_name "Int.zero_int_inst.zero_int"},
 @{const_name "List.filter"}]
 
-fun obtain_specification_graph table constname =
+fun case_consts thy s = is_some (Datatype.info_of_case thy s)
+
+fun obtain_specification_graph thy table constname =
   let
     fun is_nondefining_constname c = member (op =) logic_operator_names c
     val is_defining_constname = member (op =) (Symtab.keys table)
+    fun has_code_pred_intros c = is_some (try (Predicate_Compile_Core.intros_of thy) c)
     fun defiants_of specs =
       fold (Term.add_const_names o prop_of) specs []
       |> filter is_defining_constname
+      |> filter_out is_nondefining_constname
+      |> filter_out has_code_pred_intros
+      |> filter_out (case_consts thy)
       |> filter_out special_cases
     fun extend constname =
       let
--- a/src/HOL/Tools/Predicate_Compile/pred_compile_fun.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/Predicate_Compile/pred_compile_fun.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -5,9 +5,10 @@
 
 signature PREDICATE_COMPILE_FUN =
 sig
-  val define_predicates : (string * thm list) list -> theory -> theory
+val define_predicates : (string * thm list) list -> theory -> (string * thm list) list * theory
   val rewrite_intro : theory -> thm -> thm list
   val setup_oracle : theory -> theory
+  val pred_of_function : theory -> string -> string option
 end;
 
 structure Predicate_Compile_Fun : PREDICATE_COMPILE_FUN =
@@ -63,6 +64,8 @@
   fun merge _ = Symtab.merge (op =);
 )
 
+fun pred_of_function thy name = Symtab.lookup (Pred_Compile_Preproc.get thy) name
+
 fun defined thy = Symtab.defined (Pred_Compile_Preproc.get thy) 
 
 
@@ -99,23 +102,29 @@
       (Free (Long_Name.base_name name ^ "P", pred_type T))
   end
 
-fun mk_param lookup_pred (t as Free (v, _)) = lookup_pred t
-  | mk_param lookup_pred t =
+fun mk_param thy lookup_pred (t as Free (v, _)) = lookup_pred t
+  | mk_param thy lookup_pred t =
   let
-    val (vs, body) = strip_abs t
-    val names = Term.add_free_names body []
-    val vs_names = Name.variant_list names (map fst vs)
-    val vs' = map2 (curry Free) vs_names (map snd vs)
-    val body' = subst_bounds (rev vs', body)
-    val (f, args) = strip_comb body'
-    val resname = Name.variant (vs_names @ names) "res"
-    val resvar = Free (resname, body_type (fastype_of body'))
-    val P = lookup_pred f
-    val pred_body = list_comb (P, args @ [resvar])
-    val param = fold_rev lambda (vs' @ [resvar]) pred_body
-  in param end;
-
-
+  val _ = tracing ("called param with " ^ (Syntax.string_of_term_global thy t))
+  in if Predicate_Compile_Aux.is_predT (fastype_of t) then
+    t
+  else
+    let
+      val (vs, body) = strip_abs t
+      val names = Term.add_free_names body []
+      val vs_names = Name.variant_list names (map fst vs)
+      val vs' = map2 (curry Free) vs_names (map snd vs)
+      val body' = subst_bounds (rev vs', body)
+      val (f, args) = strip_comb body'
+      val resname = Name.variant (vs_names @ names) "res"
+      val resvar = Free (resname, body_type (fastype_of body'))
+      (*val P = case try lookup_pred f of SOME P => P | NONE => error "mk_param"
+      val pred_body = list_comb (P, args @ [resvar])
+      *)
+      val pred_body = HOLogic.mk_eq (body', resvar)
+      val param = fold_rev lambda (vs' @ [resvar]) pred_body
+    in param end
+  end
 (* creates the list of premises for every intro rule *)
 (* theory -> term -> (string list, term list list) *)
 
@@ -210,10 +219,14 @@
   let
     fun mk_prems' (t as Const (name, T)) (names, prems) =
       if is_constr thy name orelse (is_none (try lookup_pred t)) then
-        [(t ,(names, prems))]
+        [(t, (names, prems))]
       else [(lookup_pred t, (names, prems))]
     | mk_prems' (t as Free (f, T)) (names, prems) = 
       [(lookup_pred t, (names, prems))]
+    | mk_prems' (t as Abs _) (names, prems) =
+      if Predicate_Compile_Aux.is_predT (fastype_of t) then
+      [(t, (names, prems))] else error "mk_prems': Abs "
+      (* mk_param *)
     | mk_prems' t (names, prems) =
       if Predicate_Compile_Aux.is_constrt thy t then
         [(t, (names, prems))]
@@ -243,8 +256,10 @@
             maps mk_prems_of_assm assms
           end
         else
-          let 
+          let
             val (f, args) = strip_comb t
+            (* TODO: special procedure for higher-order functions: split arguments in
+              simple types and function types *)
             val resname = Name.variant names "res"
             val resvar = Free (resname, body_type (fastype_of t))
             val names' = resname :: names
@@ -261,8 +276,7 @@
                   val pred = lookup_pred t
                   val nparams = get_nparams pred
                   val (params, args) = chop nparams args
-                  val _ = tracing ("mk_prems'': " ^ (Syntax.string_of_term_global thy t) ^ " has " ^ string_of_int nparams ^ " parameters.")
-                  val params' = map (mk_param lookup_pred) params
+                  val params' = map (mk_param thy lookup_pred) params
                 in
                   folds_map mk_prems' args (names', prems)
                   |> map (fn (argvs, (names'', prems')) =>
@@ -281,7 +295,8 @@
                        val prem = HOLogic.mk_Trueprop (list_comb (pred, argvs @ [resvar]))
                      in (names', prem :: prems') end)
                 end
-            | mk_prems'' t = error ("Invalid term: " ^ Syntax.string_of_term_global thy t)
+            | mk_prems'' t =
+              error ("Invalid term: " ^ Syntax.string_of_term_global thy t)
           in
             map (pair resvar) (mk_prems'' f)
           end
@@ -292,7 +307,7 @@
 (* assumption: mutual recursive predicates all have the same parameters. *)  
 fun define_predicates specs thy =
   if forall (fn (const, _) => member (op =) (Symtab.keys (Pred_Compile_Preproc.get thy)) const) specs then
-    thy
+    ([], thy)
   else
   let
     val consts = map fst specs
@@ -307,16 +322,14 @@
     val funnames = map (fst o dest_Const) funs
     val fun_pred_names = (funnames ~~ prednames)  
       (* mapping from term (Free or Const) to term *)
-    fun lookup_pred (Const (@{const_name Cons}, T)) =
-      Const ("Preprocessing.ConsP", pred_type T) (* FIXME: temporary - Cons lookup *)
-      | lookup_pred (Const (name, T)) =
+    fun lookup_pred (Const (name, T)) =
       (case (Symtab.lookup (Pred_Compile_Preproc.get thy) name) of
           SOME c => Const (c, pred_type T)
         | NONE =>
           (case AList.lookup op = fun_pred_names name of
             SOME f => Free (f, pred_type T)
           | NONE => Const (name, T)))
-      | lookup_pred  (Free (name, T)) =
+      | lookup_pred (Free (name, T)) =
         if member op = (map fst pnames) name then
           Free (name, transform_ho_typ T)
         else
@@ -347,16 +360,14 @@
             Logic.list_implies (prems, HOLogic.mk_Trueprop (list_comb (pred, args @ [resultt]))))
         end
     fun mk_rewr_thm (func, pred) = @{thm refl}
-  in    
+  in
     case try (maps mk_intros) ((funs ~~ preds) ~~ (argss' ~~ rhss)) of
-      NONE => thy 
-    | SOME intr_ts => let
-        val _ = map (tracing o (Syntax.string_of_term_global thy)) intr_ts      
-      in
+      NONE => ([], thy) 
+    | SOME intr_ts =>
         if is_some (try (map (cterm_of thy)) intr_ts) then
           let
             val (ind_result, thy') =
-              Inductive.add_inductive_global (serial_string ())
+              Inductive.add_inductive_global (serial ())
                 {quiet_mode = false, verbose = false, kind = Thm.internalK,
                   alt_name = Binding.empty, coind = false, no_elim = false,
                   no_ind = false, skip_mono = false, fork_mono = false}
@@ -367,10 +378,15 @@
             val prednames = map (fst o dest_Const) (#preds ind_result)
             (* val rewr_thms = map mk_rewr_eq ((distinct (op =) funs) ~~ (#preds ind_result)) *)
             (* add constants to my table *)
-          in Pred_Compile_Preproc.map (fold Symtab.update_new (consts ~~ prednames)) thy' end
+            val specs = map (fn predname => (predname, filter (Predicate_Compile_Aux.is_intro predname) (#intrs ind_result))) prednames
+            val thy'' = Pred_Compile_Preproc.map (fold Symtab.update_new (consts ~~ prednames)) thy'
+          in
+            (specs, thy'')
+          end
         else
-          thy
-      end
+          let
+            val _ = tracing "Introduction rules of function_predicate are not welltyped"
+          in ([], thy) end
   end
 
 (* preprocessing intro rules - uses oracle *)
@@ -391,7 +407,6 @@
     | get_nparams t = error ("No parameters for " ^ (Syntax.string_of_term_global thy t))
     
     val intro_t = (Logic.unvarify o prop_of) intro
-    val _ = tracing (Syntax.string_of_term_global thy intro_t)
     val (prems, concl) = Logic.strip_horn intro_t
     val frees = map fst (Term.add_frees intro_t [])
     fun rewrite prem names =
@@ -415,8 +430,6 @@
         rewrite concl frees'
         |> map (fn (concl'::conclprems, _) =>
           Logic.list_implies ((flat prems') @ conclprems, concl')))
-    val _ = tracing ("intro_ts': " ^
-      commas (map (Syntax.string_of_term_global thy) intro_ts'))
   in
     map (Drule.standard o the_oracle () o cterm_of thy) intro_ts'
   end; 
--- a/src/HOL/Tools/Predicate_Compile/pred_compile_pred.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/Predicate_Compile/pred_compile_pred.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -123,16 +123,67 @@
       else
         error ("unexpected specification for constant " ^ quote constname ^ ":\n"
           ^ commas (map (quote o Display.string_of_thm_global thy) specs))
-    val _ = tracing ("Introduction rules of definitions before flattening: "
-      ^ commas (map (Display.string_of_thm ctxt) intros))
-    val _ = tracing "Defining local predicates and their intro rules..."
     val (intros', (local_defs, thy')) = flatten_intros constname intros thy
     val (intross, thy'') = fold_map preprocess local_defs thy'
   in
-    (intros' :: flat intross,thy'')
+    ((constname, intros') :: flat intross,thy'')
   end;
 
 fun preprocess_term t thy = error "preprocess_pred_term: to implement" 
-  
-  
+
+fun is_Abs (Abs _) = true
+  | is_Abs _       = false
+
+fun flat_higher_order_arguments (intross, thy) =
+  let
+    fun process constname atom (new_defs, thy) =
+      let
+        val (pred, args) = strip_comb atom
+        val abs_args = filter is_Abs args
+        fun replace_abs_arg (abs_arg as Abs _ ) (new_defs, thy) =
+          let
+            val _ = tracing ("Introduce new constant for " ^
+              Syntax.string_of_term_global thy abs_arg)
+            val vars = map Var (Term.add_vars abs_arg [])
+            val abs_arg' = Logic.unvarify abs_arg
+            val frees = map Free (Term.add_frees abs_arg' [])
+            val constname = Name.variant (map (Long_Name.base_name o fst) new_defs)
+              ((Long_Name.base_name constname) ^ "_hoaux")
+            val full_constname = Sign.full_bname thy constname
+            val constT = map fastype_of frees ---> (fastype_of abs_arg')
+            val const = Const (full_constname, constT)
+            val lhs = list_comb (const, frees)
+            val def = Logic.mk_equals (lhs, abs_arg')
+            val _ = tracing (Syntax.string_of_term_global thy def)
+            val ([definition], thy') = thy
+              |> Sign.add_consts_i [(Binding.name constname, constT, NoSyn)]
+              |> PureThy.add_defs false [((Binding.name (constname ^ "_def"), def), [])]
+          in
+            (list_comb (Logic.varify const, vars), ((full_constname, [definition])::new_defs, thy'))
+          end
+        | replace_abs_arg arg (new_defs, thy) = (arg, (new_defs, thy))
+        val (args', (new_defs', thy')) = fold_map replace_abs_arg args (new_defs, thy)
+      in
+        (list_comb (pred, args'), (new_defs', thy'))
+      end
+    fun flat_intro intro (new_defs, thy) =
+      let
+        val constname = fst (dest_Const (fst (strip_comb
+          (HOLogic.dest_Trueprop (Logic.strip_imp_concl (prop_of intro))))))
+        val (intro_ts, (new_defs, thy)) = fold_map_atoms (process constname) (prop_of intro) (new_defs, thy)
+        val th = Skip_Proof.make_thm thy intro_ts
+      in
+        (th, (new_defs, thy))
+      end
+    fun fold_map_spec f [] s = ([], s)
+      | fold_map_spec f ((c, ths) :: specs) s =
+        let
+          val (ths', s') = f ths s
+          val (specs', s'') = fold_map_spec f specs s'
+        in ((c, ths') :: specs', s'') end
+    val (intross', (new_defs, thy')) = fold_map_spec (fold_map flat_intro) intross ([], thy)
+  in
+    (intross', (new_defs, thy'))
+  end
+
 end;
--- a/src/HOL/Tools/Predicate_Compile/pred_compile_quickcheck.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/Predicate_Compile/pred_compile_quickcheck.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -3,18 +3,19 @@
 A quickcheck generator based on the predicate compiler
 *)
 
-signature PRED_COMPILE_QUICKCHECK =
+signature PREDICATE_COMPILE_QUICKCHECK =
 sig
   val quickcheck : Proof.context -> term -> int -> term list option
   val test_ref :
     ((unit -> int -> int * int -> term list Predicate.pred * (int * int)) option) Unsynchronized.ref
 end;
 
-structure Pred_Compile_Quickcheck : PRED_COMPILE_QUICKCHECK =
+structure Predicate_Compile_Quickcheck : PREDICATE_COMPILE_QUICKCHECK =
 struct
 
 val test_ref =
   Unsynchronized.ref (NONE : (unit -> int -> int * int -> term list Predicate.pred * (int * int)) option)
+
 val target = "Quickcheck"
 
 fun dest_compfuns (Predicate_Compile_Core.CompilationFuns funs) = funs
@@ -63,21 +64,21 @@
     val _ = tracing (Display.string_of_thm ctxt' intro)
     val thy'' = thy'
       |> Context.theory_map (Predicate_Compile_Preproc_Const_Defs.add_thm intro)
-      |> Predicate_Compile.preprocess full_constname
-      |> Predicate_Compile_Core.add_equations [full_constname]
-      |> Predicate_Compile_Core.add_sizelim_equations [full_constname]
-      |> Predicate_Compile_Core.add_quickcheck_equations [full_constname]
-    val sizelim_modes = Predicate_Compile_Core.sizelim_modes_of thy'' full_constname
+      |> Predicate_Compile.preprocess Predicate_Compile_Aux.default_options full_constname
+      (* |> Predicate_Compile_Core.add_equations Predicate_Compile_Aux.default_options [full_constname]*)
+      (*  |> Predicate_Compile_Core.add_depth_limited_equations Predicate_Compile_Aux.default_options [full_constname]*)
+      (* |> Predicate_Compile_Core.add_quickcheck_equations Predicate_Compile_Aux.default_options [full_constname] *)
+    val depth_limited_modes = Predicate_Compile_Core.depth_limited_modes_of thy'' full_constname
     val modes = Predicate_Compile_Core.generator_modes_of thy'' full_constname  
     val prog =
       if member (op =) modes ([], []) then
         let
           val name = Predicate_Compile_Core.generator_name_of thy'' full_constname ([], [])
-          val T = @{typ code_numeral} --> (mk_rpredT (HOLogic.mk_tupleT (map snd vs')))
-        in Const (name, T) $ Bound 0 end
-      else if member (op =) sizelim_modes ([], []) then
+          val T = [@{typ bool}, @{typ code_numeral}] ---> (mk_rpredT (HOLogic.mk_tupleT (map snd vs')))
+          in Const (name, T) $ @{term True} $ Bound 0 end
+      else if member (op =) depth_limited_modes ([], []) then
         let
-          val name = Predicate_Compile_Core.sizelim_function_name_of thy'' full_constname ([], [])
+          val name = Predicate_Compile_Core.depth_limited_function_name_of thy'' full_constname ([], [])
           val T = @{typ code_numeral} --> (mk_predT (HOLogic.mk_tupleT (map snd vs')))
         in lift_pred (Const (name, T) $ Bound 0) end
       else error "Predicate Compile Quickcheck failed"
@@ -85,7 +86,7 @@
       mk_split_lambda (map Free vs') (mk_return (HOLogic.mk_list @{typ term}
       (map2 HOLogic.mk_term_of (map snd vs') (map Free vs'))))))
     val _ = tracing (Syntax.string_of_term ctxt' qc_term)
-    val compile = Code_ML.eval (SOME target) ("Pred_Compile_Quickcheck.test_ref", test_ref)
+    val compile = Code_ML.eval (SOME target) ("Predicate_Compile_Quickcheck.test_ref", test_ref)
       (fn proc => fn g => fn s => g s #>> (Predicate.map o map) proc)
       thy'' qc_term []
   in
--- a/src/HOL/Tools/Predicate_Compile/pred_compile_set.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/Predicate_Compile/pred_compile_set.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -3,7 +3,7 @@
 Preprocessing sets to predicates
 *)
 
-signature PRED_COMPILE_SET =
+signature PREDICATE_COMPILE_SET =
 sig
 (*
   val preprocess_intro : thm -> theory -> thm * theory
@@ -12,10 +12,11 @@
   val unfold_set_notation : thm -> thm;
 end;
 
-structure Pred_Compile_Set : PRED_COMPILE_SET =
+structure Predicate_Compile_Set : PREDICATE_COMPILE_SET =
 struct
 (*FIXME: unfolding Ball in pretty adhoc here *)
-val unfold_set_lemmas = [@{thm Collect_def}, @{thm mem_def}, @{thm Ball_def}]
+val unfold_set_lemmas = [@{thm Collect_def}, @{thm mem_def},
+@{thm Ball_def}, @{thm Bex_def}]
 
 val unfold_set_notation = Simplifier.rewrite_rule unfold_set_lemmas
 
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -4,87 +4,170 @@
 signature PREDICATE_COMPILE =
 sig
   val setup : theory -> theory
-  val preprocess : string -> theory -> theory
+  val preprocess : Predicate_Compile_Aux.options -> string -> theory -> theory
 end;
 
 structure Predicate_Compile : PREDICATE_COMPILE =
 struct
 
+(* options *)
+val fail_safe_mode = true
+
 open Predicate_Compile_Aux;
 
 val priority = tracing;
 
 (* Some last processing *)
+
 fun remove_pointless_clauses intro =
   if Logic.strip_imp_prems (prop_of intro) = [@{prop "False"}] then
     []
   else [intro]
 
-fun preprocess_strong_conn_constnames gr constnames thy =
+fun tracing s = ()
+
+fun print_intross options thy msg intross =
+  if show_intermediate_results options then
+   Output.tracing (msg ^ 
+    (space_implode "\n" (map 
+      (fn (c, intros) => "Introduction rule(s) of " ^ c ^ ":\n" ^
+         commas (map (Display.string_of_thm_global thy) intros)) intross)))
+  else ()
+      
+fun print_specs thy specs =
+  map (fn (c, thms) => "Constant " ^ c ^ " has specification:\n"
+    ^ (space_implode "\n" (map (Display.string_of_thm_global thy) thms)) ^ "\n") specs
+
+fun map_specs f specs =
+  map (fn (s, ths) => (s, f ths)) specs
+
+fun process_specification options specs thy' =
+  let
+    val _ = print_step options "Compiling predicates to flat introrules..."
+    val specs = map (apsnd (map
+      (fn th => if is_equationlike th then Predicate_Compile_Data.normalize_equation thy' th else th))) specs
+    val (intross1, thy'') = apfst flat (fold_map Predicate_Compile_Pred.preprocess specs thy')
+    val _ = print_intross options thy'' "Flattened introduction rules: " intross1
+    val _ = print_step options "Replacing functions in introrules..."
+    val intross2 =
+      if fail_safe_mode then
+        case try (map_specs (maps (Predicate_Compile_Fun.rewrite_intro thy''))) intross1 of
+          SOME intross => intross
+        | NONE => let val _ = warning "Function replacement failed!" in intross1 end
+      else map_specs (maps (Predicate_Compile_Fun.rewrite_intro thy'')) intross1
+    val _ = print_intross options thy'' "Introduction rules with replaced functions: " intross2
+    val _ = print_step options "Introducing new constants for abstractions at higher-order argument positions..."
+    val (intross3, (new_defs, thy''')) = Predicate_Compile_Pred.flat_higher_order_arguments (intross2, thy'')
+    val (new_intross, thy'''')  =
+      if not (null new_defs) then
+      let
+        val _ = print_step options "Recursively obtaining introduction rules for new definitions..."
+      in process_specification options new_defs thy''' end
+    else ([], thy''')
+  in
+    (intross3 @ new_intross, thy'''')
+  end
+
+
+fun preprocess_strong_conn_constnames options gr constnames thy =
   let
     val get_specs = map (fn k => (k, Graph.get_node gr k))
-    val _ = priority ("Preprocessing scc of " ^ commas constnames)
+    val _ = print_step options ("Preprocessing scc of " ^ commas constnames)
     val (prednames, funnames) = List.partition (is_pred thy) constnames
     (* untangle recursion by defining predicates for all functions *)
-    val _ = priority "Compiling functions to predicates..."
-    val _ = tracing ("funnames: " ^ commas funnames)
-    val thy' =
-      thy |> not (null funnames) ? Predicate_Compile_Fun.define_predicates
-      (get_specs funnames)
-    val _ = priority "Compiling predicates to flat introrules..."
-    val (intross, thy'') = apfst flat (fold_map Predicate_Compile_Pred.preprocess
-      (get_specs prednames) thy')
-    val _ = tracing ("Flattened introduction rules: " ^
-      commas (map (Display.string_of_thm_global thy'') (flat intross)))
-    val _ = priority "Replacing functions in introrules..."
-      (*  val _ = burrow (maps (Predicate_Compile_Fun.rewrite_intro thy'')) intross  *)
-    val intross' =
-      case try (burrow (maps (Predicate_Compile_Fun.rewrite_intro thy''))) intross of
-        SOME intross' => intross'
-      | NONE => let val _ = warning "Function replacement failed!" in intross end
-    val _ = tracing ("Introduction rules with replaced functions: " ^
-      commas (map (Display.string_of_thm_global thy'') (flat intross')))
-    val intross'' = burrow (maps remove_pointless_clauses) intross'
-    val intross'' = burrow (map (AxClass.overload thy'')) intross''
-    val _ = priority "Registering intro rules..."
-    val thy''' = fold Predicate_Compile_Core.register_intros intross'' thy''
+    val _ = print_step options
+      ("Compiling functions (" ^ commas funnames ^ ") to predicates...")
+    val (fun_pred_specs, thy') =
+      if not (null funnames) then Predicate_Compile_Fun.define_predicates
+      (get_specs funnames) thy else ([], thy)
+    val _ = print_specs thy' fun_pred_specs
+    val specs = (get_specs prednames) @ fun_pred_specs
+    val (intross3, thy''') = process_specification options specs thy'
+    val _ = print_intross options thy''' "Introduction rules with new constants: " intross3
+    val intross4 = map_specs (maps remove_pointless_clauses) intross3
+    val _ = print_intross options thy''' "After removing pointless clauses: " intross4
+      (*val intross5 = map (fn s, ths) => ( s, map (AxClass.overload thy''') ths)) intross4*)
+    val intross6 = map_specs (map (expand_tuples thy''')) intross4
+    val _ = print_intross options thy''' "introduction rules before registering: " intross6
+    val _ = print_step options "Registering introduction rules..."
+    val thy'''' = fold Predicate_Compile_Core.register_intros intross6 thy'''
   in
-    thy'''
+    thy''''
   end;
 
-fun preprocess const thy =
+fun preprocess options const thy =
   let
-    val _ = tracing ("Fetching definitions from theory...")
-    val table = Pred_Compile_Data.make_const_spec_table thy
-    val gr = Pred_Compile_Data.obtain_specification_graph table const
-    val _ = tracing (commas (Graph.all_succs gr [const]))
+    val _ = print_step options "Fetching definitions from theory..."
+    val table = Predicate_Compile_Data.make_const_spec_table thy
+    val gr = Predicate_Compile_Data.obtain_specification_graph thy table const
     val gr = Graph.subgraph (member (op =) (Graph.all_succs gr [const])) gr
-  in fold_rev (preprocess_strong_conn_constnames gr)
+  in fold_rev (preprocess_strong_conn_constnames options gr)
     (Graph.strong_conn gr) thy
   end
 
-fun code_pred_cmd ((inductify_all, rpred), raw_const) lthy =
-  if inductify_all then
-    let
-      val thy = ProofContext.theory_of lthy
-      val const = Code.read_const thy raw_const
-      val lthy' = LocalTheory.theory (preprocess const) lthy
-        |> LocalTheory.checkpoint
-      val _ = tracing "Starting Predicate Compile Core..."
-    in Predicate_Compile_Core.code_pred_cmd rpred raw_const lthy' end
-  else
-    Predicate_Compile_Core.code_pred_cmd rpred raw_const lthy
+fun extract_options ((modes, raw_options), const) =
+  let
+    fun chk s = member (op =) raw_options s
+  in
+    Options {
+      expected_modes = Option.map (pair const) modes,
+      show_steps = chk "show_steps",
+      show_intermediate_results = chk "show_intermediate_results",
+      show_proof_trace = chk "show_proof_trace",
+      show_mode_inference = chk "show_mode_inference",
+      show_compilation = chk "show_compilation",
+      skip_proof = chk "skip_proof",
+      inductify = chk "inductify",
+      rpred = chk "rpred",
+      depth_limited = chk "depth_limited"
+    }
+  end
+
+fun code_pred_cmd ((modes, raw_options), raw_const) lthy =
+  let
+     val thy = ProofContext.theory_of lthy
+     val const = Code.read_const thy raw_const
+     val options = extract_options ((modes, raw_options), const)
+  in
+    if (is_inductify options) then
+      let
+        val lthy' = LocalTheory.theory (preprocess options const) lthy
+          |> LocalTheory.checkpoint
+        val const = case Predicate_Compile_Fun.pred_of_function (ProofContext.theory_of lthy') const of
+            SOME c => c
+          | NONE => const
+        val _ = print_step options "Starting Predicate Compile Core..."
+      in
+        Predicate_Compile_Core.code_pred options const lthy'
+      end
+    else
+      Predicate_Compile_Core.code_pred_cmd options raw_const lthy
+  end
 
 val setup = Predicate_Compile_Fun.setup_oracle #> Predicate_Compile_Core.setup
 
-val _ = List.app OuterKeyword.keyword ["inductify_all", "rpred"]
+val bool_options = ["show_steps", "show_intermediate_results", "show_proof_trace",
+  "show_mode_inference", "show_compilation", "skip_proof", "inductify", "rpred", "depth_limited"]
 
 local structure P = OuterParse
 in
 
+val opt_modes =
+  Scan.optional (P.$$$ "(" |-- Args.$$$ "mode" |-- P.$$$ ":" |--
+   P.enum1 "," (P.$$$ "[" |-- P.enum "," P.nat --| P.$$$ "]")
+  --| P.$$$ ")" >> SOME) NONE
+
+val scan_params =
+  let
+    val scan_bool_param = foldl1 (op ||) (map Args.$$$ bool_options)
+  in
+    Scan.optional (P.$$$ "[" |-- P.enum1 "," scan_bool_param --| P.$$$ "]") []
+  end
+
 val _ = OuterSyntax.local_theory_to_proof "code_pred"
   "prove equations for predicate specified by intro/elim rules"
-  OuterKeyword.thy_goal (P.opt_keyword "inductify_all" -- P.opt_keyword "rpred" -- P.term_group >> code_pred_cmd)
+  OuterKeyword.thy_goal (opt_modes -- scan_params -- P.term_group >>
+    code_pred_cmd)
 
 end
 
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -7,26 +7,25 @@
 signature PREDICATE_COMPILE_CORE =
 sig
   val setup: theory -> theory
-  val code_pred: bool -> string -> Proof.context -> Proof.state
-  val code_pred_cmd: bool -> string -> Proof.context -> Proof.state
+  val code_pred: Predicate_Compile_Aux.options -> string -> Proof.context -> Proof.state
+  val code_pred_cmd: Predicate_Compile_Aux.options -> string -> Proof.context -> Proof.state
   type smode = (int * int list option) list
   type mode = smode option list * smode
   datatype tmode = Mode of mode * smode * tmode option list;
-  (*val add_equations_of: bool -> string list -> theory -> theory *)
-  val register_predicate : (thm list * thm * int) -> theory -> theory
-  val register_intros : thm list -> theory -> theory
+  val register_predicate : (string * thm list * thm * int) -> theory -> theory
+  val register_intros : string * thm list -> theory -> theory
   val is_registered : theory -> string -> bool
- (* 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 sizelim_modes_of: theory -> string -> mode list
-  val sizelim_function_name_of : theory -> string -> mode -> string
+  val depth_limited_modes_of: theory -> string -> mode list
+  val depth_limited_function_name_of : theory -> string -> mode -> string
   val generator_modes_of: theory -> string -> mode list
   val generator_name_of : theory -> string -> mode -> string
+  val all_modes_of : theory -> (string * mode list) list
+  val all_generator_modes_of : theory -> (string * mode list) list
   val string_of_mode : mode -> string
   val intros_of: theory -> string -> thm list
   val nparams_of: theory -> string -> int
@@ -35,24 +34,12 @@
   val set_nparams : string -> int -> theory -> theory
   val print_stored_rules: theory -> unit
   val print_all_modes: theory -> unit
-  val do_proofs: bool Unsynchronized.ref
-  val mk_casesrule : Proof.context -> int -> thm list -> term
-  val analyze_compr: theory -> term -> term
-  val eval_ref: (unit -> term Predicate.pred) option Unsynchronized.ref
-  val add_equations : string list -> theory -> theory
+  val mk_casesrule : Proof.context -> term -> int -> thm list -> term
+  val eval_ref : (unit -> term Predicate.pred) option Unsynchronized.ref
+  val random_eval_ref : (unit -> int * int -> term Predicate.pred * (int * int)) option Unsynchronized.ref
   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,
@@ -64,66 +51,39 @@
     mk_not : term -> term,
     mk_map : typ -> typ -> term -> term -> term,
     lift_pred : term -> term
-  };  
-  type moded_clause = term list * (indprem * tmode) list
-  type 'a pred_mode_table = (string * (mode * 'a) list) list
-  val infer_modes : theory -> (string * mode list) list
-    -> (string * mode list) list
-    -> string list
-    -> (string * (term list * indprem list) list) list
-    -> (moded_clause list) pred_mode_table
-  val infer_modes_with_generator : theory -> (string * mode list) list
-    -> (string * mode list) 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 pred_compfuns : compilation_funs
   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
-  val cprods_subset : 'a list list -> 'a list list
+    (*  val add_equations : Predicate_Compile_Aux.options -> string list -> theory -> theory
+  val add_quickcheck_equations : Predicate_Compile_Aux.options -> string list -> theory -> theory
+  val add_depth_limited_equations : Predicate_Compile_Aux.options -> string list -> theory -> theory
+  *)
 end;
 
 structure Predicate_Compile_Core : PREDICATE_COMPILE_CORE =
 struct
 
 open Predicate_Compile_Aux;
+
 (** auxiliary **)
 
 (* debug stuff *)
 
-fun tracing s = (if ! Toplevel.debug then tracing s else ());
+fun tracing s = (if ! Toplevel.debug then Output.tracing s else ());
+
+fun print_tac s = Seq.single;
+
+fun print_tac' options s = 
+  if show_proof_trace options then Tactical.print_tac s else Seq.single;
 
-fun print_tac s = Seq.single; (*Tactical.print_tac s;*) (* (if ! Toplevel.debug then Tactical.print_tac s else Seq.single); *)
-fun debug_tac msg = Seq.single; (* (fn st => (tracing msg; Seq.single st)); *)
+fun debug_tac msg = Seq.single; (* (fn st => (Output.tracing msg; Seq.single st)); *)
 
-val do_proofs = Unsynchronized.ref true;
+datatype assertion = Max_number_of_subgoals of int
+fun assert_tac (Max_number_of_subgoals i) st =
+  if (nprems_of st <= i) then Seq.single st
+  else error ("assert_tac: Numbers of subgoals mismatch at goal state :"
+    ^ "\n" ^ Pretty.string_of (Pretty.chunks
+      (Goal_Display.pretty_goals_without_context (! Goal_Display.goals_limit) st)));
 
 (* reference to preprocessing of InductiveSet package *)
 
@@ -139,20 +99,6 @@
             HOLogic.mk_eq (Free (a, fastype_of b), b) :: mk_eqs a cs
   in mk_eqs x xs end;
 
-fun mk_tupleT [] = HOLogic.unitT
-  | mk_tupleT Ts = foldr1 HOLogic.mk_prodT Ts;
-
-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 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_scomp (t, u) =
   let
     val T = fastype_of t
@@ -189,6 +135,13 @@
 
 (** data structures **)
 
+(* new datatype for modes: *)
+(*
+datatype instantiation = Input | Output
+type arg_mode = Tuple of instantiation list | Atom of instantiation | HigherOrderMode of mode
+type mode = arg_mode list
+type tmode = Mode of mode * 
+*)
 type smode = (int * int list option) list
 type mode = smode option list * smode;
 datatype tmode = Mode of mode * smode * tmode option list;
@@ -241,35 +194,8 @@
   (if null param_modes then "" else
     "; " ^ "params: " ^ commas (map (the_default "NONE" o Option.map string_of_tmode) param_modes))
 
-(* generation of case rules from user-given introduction rules *)
-
-fun mk_casesrule ctxt nparams introrules =
-  let
-    val ((_, intros_th), ctxt1) = Variable.import false introrules ctxt
-    val intros = map prop_of intros_th
-    val (pred, (params, args)) = strip_intro_concl nparams (hd intros)
-    val ([propname], ctxt2) = Variable.variant_fixes ["thesis"] ctxt1
-    val prop = HOLogic.mk_Trueprop (Free (propname, HOLogic.boolT))
-    val (argnames, ctxt3) = Variable.variant_fixes
-      (map (fn i => "a" ^ string_of_int i) (1 upto (length args))) ctxt2
-    val argvs = map2 (curry Free) argnames (map fastype_of args)
-    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)
-        val frees = (fold o fold_aterms)
-          (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
-    val assm = HOLogic.mk_Trueprop (list_comb (pred, params @ argvs))
-    val cases = map mk_case intros
-  in Logic.list_implies (assm :: cases, prop) end;
-    
-
-datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term |
-  GeneratorPrem of term list * term | Generator of (string * typ);
+datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term
+  | Generator of (string * typ);
 
 type moded_clause = term list * (indprem * tmode) list
 type 'a pred_mode_table = (string * (mode * 'a) list) list
@@ -300,25 +226,25 @@
   nparams : int,
   functions : (mode * predfun_data) list,
   generators : (mode * function_data) list,
-  sizelim_functions : (mode * function_data) list 
+  depth_limited_functions : (mode * function_data) list 
 };
 
 fun rep_pred_data (PredData data) = data;
-fun mk_pred_data ((intros, elim, nparams), (functions, generators, sizelim_functions)) =
+fun mk_pred_data ((intros, elim, nparams), (functions, generators, depth_limited_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)))
-  
+    functions = functions, generators = generators, depth_limited_functions = depth_limited_functions}
+fun map_pred_data f (PredData {intros, elim, nparams, functions, generators, depth_limited_functions}) =
+  mk_pred_data (f ((intros, elim, nparams), (functions, generators, depth_limited_functions)))
+
 fun eq_option eq (NONE, NONE) = true
   | eq_option eq (SOME x, SOME y) = eq (x, y)
   | eq_option eq _ = false
-  
+
 fun eq_pred_data (PredData d1, PredData d2) = 
   eq_list (Thm.eq_thm) (#intros d1, #intros d2) andalso
   eq_option (Thm.eq_thm) (#elim d1, #elim d2) andalso
   #nparams d1 = #nparams d2
-  
+
 structure PredData = TheoryDataFun
 (
   type T = pred_data Graph.T;
@@ -353,7 +279,7 @@
 
 val modes_of = (map fst) o #functions oo the_pred_data
 
-val sizelim_modes_of = (map fst) o #sizelim_functions oo the_pred_data
+val depth_limited_modes_of = (map fst) o #depth_limited_functions oo the_pred_data
 
 val rpred_modes_of = (map fst) o #generators oo the_pred_data
   
@@ -380,7 +306,7 @@
 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
@@ -392,24 +318,25 @@
 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 =
+fun lookup_depth_limited_function_data thy name mode =
   Option.map rep_function_data (AList.lookup (op =)
-  (#sizelim_functions (the_pred_data thy name)) mode)
+  (#depth_limited_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
+fun the_depth_limited_function_data thy name mode = case lookup_depth_limited_function_data thy name mode
+  of NONE => error ("No depth-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 depth_limited_function_name_of = #name ooo the_depth_limited_function_data
 
 (*val generator_modes_of = (map fst) o #generators oo the_pred_data*)
-     
+
 (* diagnostic display functions *)
 
-fun print_modes modes = tracing ("Inferred modes:\n" ^
-  cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
-    string_of_mode ms)) modes));
+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
@@ -417,15 +344,20 @@
       ^ (string_of_entry pred mode entry)  
     fun print_pred (pred, modes) =
       "predicate " ^ pred ^ ": " ^ cat_lines (map (print_mode pred) modes)
-    val _ = tracing (cat_lines (map print_pred pred_mode_table))
+    val _ = Output.tracing (cat_lines (map print_pred pred_mode_table))
   in () end;
 
+fun string_of_prem thy (Prem (ts, p)) =
+    (Syntax.string_of_term_global thy (list_comb (p, ts))) ^ "(premise)"
+  | string_of_prem thy (Negprem (ts, p)) =
+    (Syntax.string_of_term_global thy (HOLogic.mk_not (list_comb (p, ts)))) ^ "(negative premise)"
+  | string_of_prem thy (Sidecond t) =
+    (Syntax.string_of_term_global thy t) ^ "(sidecondition)"
+  | string_of_prem thy _ = error "string_of_prem: unexpected input"
+
 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, _)) =
@@ -435,18 +367,25 @@
     (Syntax.string_of_term_global thy t) ^
     "(sidecond mode: " ^ string_of_smode is ^ ")"    
   | string_of_moded_prem _ _ = error "string_of_moded_prem: unimplemented"
-     
+
 fun print_moded_clauses thy =
-  let        
+  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 string_of_clause thy pred (ts, prems) =
+  (space_implode " --> "
+  (map (string_of_prem thy) prems)) ^ " --> " ^ pred ^ " "
+   ^ (space_implode " " (map (Syntax.string_of_term_global thy) ts))
+
+fun print_compiled_terms options thy =
+  if show_compilation options then
+    print_pred_mode_table (fn _ => fn _ => Syntax.string_of_term_global thy) thy
+  else K ()
+
 fun print_stored_rules thy =
   let
     val preds = (Graph.keys o PredData.get) thy
@@ -477,7 +416,105 @@
   in
     fold print (all_modes_of thy) ()
   end
-  
+
+(* validity checks *)
+
+fun check_expected_modes (options : Predicate_Compile_Aux.options) modes =
+  case expected_modes options of
+    SOME (s, ms) => (case AList.lookup (op =) modes s of
+      SOME modes =>
+        if not (eq_set (op =) (map (map (rpair NONE)) ms, map snd modes)) then
+          error ("expected modes were not inferred:\n"
+          ^ "inferred modes for " ^ s ^ ": "
+          ^ commas (map ((enclose "[" "]") o string_of_smode o snd) modes))
+        else ()
+      | NONE => ())
+  | NONE => ()
+
+(* importing introduction rules *)
+
+fun unify_consts thy cs intr_ts =
+  (let
+     val add_term_consts_2 = fold_aterms (fn Const c => insert (op =) c | _ => I);
+     fun varify (t, (i, ts)) =
+       let val t' = map_types (Logic.incr_tvar (i + 1)) (#2 (Type.varify [] t))
+       in (maxidx_of_term t', t'::ts) end;
+     val (i, cs') = List.foldr varify (~1, []) cs;
+     val (i', intr_ts') = List.foldr varify (i, []) intr_ts;
+     val rec_consts = fold add_term_consts_2 cs' [];
+     val intr_consts = fold add_term_consts_2 intr_ts' [];
+     fun unify (cname, cT) =
+       let val consts = map snd (List.filter (fn c => fst c = cname) intr_consts)
+       in fold (Sign.typ_unify thy) ((replicate (length consts) cT) ~~ consts) end;
+     val (env, _) = fold unify rec_consts (Vartab.empty, i');
+     val subst = map_types (Envir.norm_type env)
+   in (map subst cs', map subst intr_ts')
+   end) handle Type.TUNIFY =>
+     (warning "Occurrences of recursive constant have non-unifiable types"; (cs, intr_ts));
+
+fun import_intros inp_pred nparams [] ctxt =
+  let
+    val ([outp_pred], ctxt') = Variable.import_terms false [inp_pred] ctxt
+    val (paramTs, _) = chop nparams (binder_types (fastype_of outp_pred))
+    val (param_names, ctxt'') = Variable.variant_fixes (map (fn i => "p" ^ (string_of_int i))
+      (1 upto nparams)) ctxt'
+    val params = map Free (param_names ~~ paramTs)
+    in (((outp_pred, params), []), ctxt') end
+  | import_intros inp_pred nparams (th :: ths) ctxt =
+    let
+      val ((_, [th']), ctxt') = Variable.import false [th] ctxt
+      val thy = ProofContext.theory_of ctxt'
+      val (pred, (params, args)) = strip_intro_concl nparams (prop_of th')
+      val ho_args = filter (is_predT o fastype_of) args
+      fun subst_of (pred', pred) =
+        let
+          val subst = Sign.typ_match thy (fastype_of pred', fastype_of pred) Vartab.empty
+        in map (fn (indexname, (s, T)) => ((indexname, s), T)) (Vartab.dest subst) end
+      fun instantiate_typ th =
+        let
+          val (pred', _) = strip_intro_concl 0 (prop_of th)
+          val _ = if not (fst (dest_Const pred) = fst (dest_Const pred')) then
+            error "Trying to instantiate another predicate" else ()
+        in Thm.certify_instantiate (subst_of (pred', pred), []) th end;
+      fun instantiate_ho_args th =
+        let
+          val (_, (params', args')) = strip_intro_concl nparams (prop_of th)
+          val ho_args' = map dest_Var (filter (is_predT o fastype_of) args')
+        in Thm.certify_instantiate ([], map dest_Var params' ~~ params) th end
+      val outp_pred =
+        Term_Subst.instantiate (subst_of (inp_pred, pred), []) inp_pred
+      val ((_, ths'), ctxt1) =
+        Variable.import false (map (instantiate_typ #> instantiate_ho_args) ths) ctxt'
+    in
+      (((outp_pred, params), th' :: ths'), ctxt1)
+    end
+
+(* generation of case rules from user-given introduction rules *)
+
+fun mk_casesrule ctxt pred nparams introrules =
+  let
+    val (((pred, params), intros_th), ctxt1) = import_intros pred nparams introrules ctxt
+    val intros = map prop_of intros_th
+    val ([propname], ctxt2) = Variable.variant_fixes ["thesis"] ctxt1
+    val prop = HOLogic.mk_Trueprop (Free (propname, HOLogic.boolT))
+    val (_, argsT) = chop nparams (binder_types (fastype_of pred))
+    val (argnames, ctxt3) = Variable.variant_fixes
+      (map (fn i => "a" ^ string_of_int i) (1 upto length argsT)) ctxt2
+    val argvs = map2 (curry Free) argnames argsT
+    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)
+        val frees = (fold o fold_aterms)
+          (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
+    val assm = HOLogic.mk_Trueprop (list_comb (pred, params @ argvs))
+    val cases = map mk_case intros
+  in Logic.list_implies (assm :: cases, prop) end;
+
 (** preprocessing rules **)  
 
 fun imp_prems_conv cv ct =
@@ -498,39 +535,30 @@
 
 fun preprocess_elim thy nparams elimrule =
   let
-    val _ = tracing ("Preprocessing elimination rule "
-      ^ (Display.string_of_thm_global thy elimrule))
     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 ctxt = ProofContext.init thy
+    val ((_, [elimrule]), ctxt') = Variable.import false [elimrule] ctxt
     val prems = Thm.prems_of elimrule
     val nargs = length (snd (strip_comb (HOLogic.dest_Trueprop (hd prems)))) - nparams
     fun preprocess_case t =
-     let
+      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
+      in
        list_all (params, Logic.list_implies (assums_hyp', Logic.strip_assums_concl t))
-     end
+      end
     val cases' = map preprocess_case (tl prems)
     val elimrule' = Logic.list_implies ((hd prems) :: cases', Thm.concl_of elimrule)
-    (*val _ =  tracing ("elimrule': "^ (Syntax.string_of_term_global thy elimrule'))*)
     val bigeq = (Thm.symmetric (Conv.implies_concl_conv
       (MetaSimplifier.rewrite true [@{thm Predicate.eq_is_eq}])
         (cterm_of thy elimrule')))
-    (*
-    val _ = tracing ("bigeq:" ^ (Display.string_of_thm_global thy bigeq))   
-    val res = 
-    Thm.equal_elim bigeq elimrule
-    *)
-    (*
-    val t = (fn {...} => mycheat_tac thy 1)
-    val eq = Goal.prove (ProofContext.init thy) [] [] (Logic.mk_equals ((Thm.prop_of elimrule), elimrule')) t
-    *)
-    val _ = tracing "Preprocessed elimination rule"
+    val tac = (fn _ => Skip_Proof.cheat_tac thy)    
+    val eq = Goal.prove ctxt' [] [] (Logic.mk_equals ((Thm.prop_of elimrule), elimrule')) tac
   in
-    Thm.equal_elim bigeq elimrule
+    Thm.equal_elim eq elimrule |> singleton (Variable.export ctxt' ctxt)
   end;
 
 (* special case: predicate with no introduction rule *)
@@ -552,6 +580,8 @@
   ([intro], elim)
 end
 
+fun expand_tuples_elim th = th
+
 fun fetch_pred_data thy name =
   case try (Inductive.the_inductive (ProofContext.init thy)) name of
     SOME (info as (_, result)) => 
@@ -560,17 +590,23 @@
           let
             val (const, _) = strip_comb (HOLogic.dest_Trueprop (concl_of intro))
           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 intros = ind_set_codegen_preproc thy
+          (map (expand_tuples thy #> preprocess_intro thy) (filter is_intro_of (#intrs result)))
+        val index = find_index (fn s => s = name) (#names (fst info))
+        val pre_elim = nth (#elims result) index
+        val pred = nth (#preds result) index
         val nparams = length (Inductive.params_of (#raw_induct result))
-        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)
+        (*val elim = singleton (ind_set_codegen_preproc thy) (preprocess_elim thy nparams 
+          (expand_tuples_elim pre_elim))*)
+        val elim =
+          (Drule.standard o Skip_Proof.make_thm thy)
+          (mk_casesrule (ProofContext.init thy) pred nparams intros)
+        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 apfst3 f (x, y, z) =  (f x, y, z)
@@ -605,6 +641,7 @@
     (data, keys)
   end;
 *)
+
 (* guessing number of parameters *)
 fun find_indexes pred xs =
   let
@@ -624,7 +661,7 @@
    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
+         (apfst (fn (intros, elim, nparams) => (thm::intros, elim, nparams)))) gr
      | NONE =>
        let
          val nparams = the_default (guess_nparams T)  (try (#nparams o rep_pred_data o (fetch_pred_data thy)) name)
@@ -640,30 +677,34 @@
 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 =
+
+fun register_predicate (constname, 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
-    if not (member (op =) (Graph.keys (PredData.get thy)) name) then
+    if not (member (op =) (Graph.keys (PredData.get thy)) constname) then
       PredData.map
-        (Graph.new_node (name, mk_pred_data ((intros, SOME elim, nparams), ([], [], [])))) thy
+        (Graph.new_node (constname, mk_pred_data ((intros, SOME elim, nparams), ([], [], [])))) thy
     else thy
   end
 
-fun register_intros pre_intros thy =
+fun register_intros (constname, pre_intros) thy =
   let
-    val (c, T) = dest_Const (fst (strip_intro_concl 0 (prop_of (hd pre_intros))))
-    val _ = tracing ("Registering introduction rules of " ^ c)
-    val _ = tracing (commas (map (Display.string_of_thm_global thy) pre_intros))
+    val T = Sign.the_const_type thy constname
+    fun constname_of_intro intr = fst (dest_Const (fst (strip_intro_concl 0 (prop_of intr))))
+    val _ = if not (forall (fn intr => constname_of_intro intr = constname) pre_intros) then
+      error ("register_intros: Introduction rules of different constants are used\n" ^
+        "expected rules for " ^ constname ^ ", but received rules for " ^
+          commas (map constname_of_intro pre_intros))
+      else ()
+    val pred = Const (constname, T)
     val nparams = guess_nparams T
     val pre_elim = 
       (Drule.standard o Skip_Proof.make_thm thy)
-      (mk_casesrule (ProofContext.init thy) nparams pre_intros)
-  in register_predicate (pre_intros, pre_elim, nparams) thy end
+      (mk_casesrule (ProofContext.init thy) pred nparams pre_intros)
+  in register_predicate (constname, pre_intros, pre_elim, nparams) thy end
 
 fun set_generator_name pred mode name = 
   let
@@ -672,7 +713,7 @@
     PredData.map (Graph.map_node pred (map_pred_data set))
   end
 
-fun set_sizelim_function_name pred mode name = 
+fun set_depth_limited_function_name pred mode name = 
   let
     val set = (apsnd o aptrd3 o cons) (mode, mk_function_data (name, NONE))
   in
@@ -694,8 +735,6 @@
   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
 };
@@ -708,8 +747,6 @@
 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
 
@@ -719,7 +756,7 @@
     val (paramTs, (inargTs, outargTs)) = split_modeT (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))
+    (paramTs' @ inargTs) ---> (mk_predT compfuns (HOLogic.mk_tupleT outargTs))
   end;
 
 fun mk_fun_of compfuns thy (name, T) mode = 
@@ -809,9 +846,11 @@
 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_not t = let val T = mk_rpredT HOLogic.unitT
+  in Const (@{const_name RPred.not_rpred}, T --> T) $ t end
 
-fun mk_map t = error "FIXME" (*FIXME*)
+fun mk_map T1 T2 tf tp = Const (@{const_name RPred.map},
+  (T1 --> T2) --> mk_rpredT T1 --> mk_rpredT T2) $ tf $ tp
 
 fun lift_pred t =
   let
@@ -839,20 +878,21 @@
       RPredCompFuns.mk_rpredT T) $ random
   end;
 
-fun sizelim_funT_of compfuns (iss, is) T =
+fun depth_limited_funT_of compfuns (iss, is) T =
   let
     val Ts = binder_types T
     val (paramTs, (inargTs, outargTs)) = split_modeT (iss, is) Ts
-    val paramTs' = map2 (fn SOME is => sizelim_funT_of PredicateCompFuns.compfuns ([], is) | NONE => I) iss paramTs 
+    val paramTs' = map2 (fn SOME is => depth_limited_funT_of compfuns ([], is) | NONE => I) iss paramTs 
   in
-    (paramTs' @ inargTs @ [@{typ "code_numeral"}]) ---> (mk_predT compfuns (mk_tupleT outargTs))
+    (paramTs' @ inargTs @ [@{typ bool}, @{typ "code_numeral"}])
+      ---> (mk_predT compfuns (HOLogic.mk_tupleT outargTs))
   end;  
 
-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_depth_limited_fun_of compfuns thy (name, T) mode =
+  Const (depth_limited_function_name_of thy name mode, depth_limited_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)
+  Const (generator_name_of thy name mode, depth_limited_funT_of compfuns mode T)
 
 (* Mode analysis *)
 
@@ -882,16 +922,16 @@
 fun term_vTs tm =
   fold_aterms (fn Free xT => cons xT | _ => I) tm [];
 
-(*FIXME this function should not be named merge... make it local instead*)
-fun merge xs [] = xs
-  | merge [] ys = ys
-  | merge (x::xs) (y::ys) = if length x >= length y then x::merge xs (y::ys)
-      else y::merge (x::xs) ys;
-
-fun subsets i j = if i <= j then
-       let val is = subsets (i+1) j
-       in merge (map (fn ks => i::ks) is) is end
-     else [[]];
+fun subsets i j =
+  if i <= j then
+    let
+      fun merge xs [] = xs
+        | merge [] ys = ys
+        | merge (x::xs) (y::ys) = if length x >= length y then x::merge xs (y::ys)
+            else y::merge (x::xs) ys;
+      val is = subsets (i+1) j
+    in merge (map (fn ks => i::ks) is) is end
+  else [[]];
      
 (* FIXME: should be in library - cprod = map_prod I *)
 fun cprod ([], ys) = []
@@ -905,46 +945,9 @@
     val yss = (cprods_subset xss)
   in maps (fn ys => map (fn x => cons x ys) xs) yss @ yss end
   
-(*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
-    fun mk_modes name args = Option.map (maps (fn (m as (iss, is)) =>
-        let
-          val (args1, args2) =
-            if length args < length iss then
-              error ("Too few arguments for inductive predicate " ^ name)
-            else chop (length iss) args;
-          val k = length args2;
-          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
-          let
-            val is' = 
-            (fold_index (fn (i, j) => if j mem is then cons (i + 1) else I) perm [])
-            |> fold (fn i => if i > k then cons (i - k + b) else I) is
-              
-           val res = map (fn x => Mode (m, is', x)) (cprods (map
-            (fn (NONE, _) => [NONE]
-              | (SOME js, arg) => map SOME (filter
-                  (fn Mode (_, js', _) => js=js') (modes_of_term modes arg)))
-                    (iss ~~ args1)))
-          in res end
-        end)) (AList.lookup op = modes name)
-  in case strip_comb 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)
-    | _ => default end
-  
-and
-*)
 fun modes_of_term modes t =
   let
-    val ks = map_index (fn (i, T) => (i, NONE)) (binder_types (fastype_of t));
+    val ks = map_index (fn (i, T) => (i + 1, NONE)) (binder_types (fastype_of t));
     val default = [Mode (([], ks), ks, [])];
     fun mk_modes name args = Option.map (maps (fn (m as (iss, is)) =>
         let
@@ -956,7 +959,7 @@
           val prfx = map (rpair NONE) (1 upto k)
         in
           if not (is_prefix op = prfx is) then [] else
-          let val is' = List.drop (is, k)
+          let val is' = map (fn (i, t) => (i - k, t)) (List.drop (is, k))
           in map (fn x => Mode (m, is', x)) (cprods (map
             (fn (NONE, _) => [NONE]
               | (SOME js, arg) => map SOME (filter
@@ -992,7 +995,7 @@
             (modes_of_term modes t handle Option =>
                error ("Bad predicate: " ^ Syntax.string_of_term_global thy t))
       | Negprem (us, t) => find_first (fn Mode (_, is, _) =>
-            length us = length is andalso
+            is = map (rpair NONE) (1 upto length us) andalso
             subset (op =) (terms_vs us, vs) andalso
             subset (op =) (term_vs t, vs))
             (modes_of_term modes t handle Option =>
@@ -1015,24 +1018,8 @@
     (Generator (v, T), Mode (([], []), [], []))
   end;
 
-fun gen_prem (Prem (us, t)) = GeneratorPrem (us, t)
-  | gen_prem (Negprem (us, t)) = error "it is a negated prem"
-  | gen_prem (Sidecond t) = error "it is a sidecond"
-  | 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 _ = tracing ("param_vs:" ^ commas param_vs)
-  val _ = tracing ("iss:" ^
-    commas (map (fn is => case is of SOME is => string_of_smode is | NONE => "NONE") iss))
-    *)
     val modes' = modes @ map_filter
       (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
         (param_vs ~~ iss);
@@ -1046,7 +1033,7 @@
           NONE =>
             (if with_generator then
               (case select_mode_prem thy gen_modes' vs ps of
-                SOME (p as Prem _, SOME mode) => check_mode_prems ((gen_prem p, mode) :: acc_ps) 
+                SOME (p as Prem _, SOME mode) => check_mode_prems ((p, mode) :: acc_ps) 
                   (case p of Prem (us, _) => union (op =) vs (terms_vs us) | _ => vs)
                   (filter_out (equal p) ps)
               | _ =>
@@ -1057,14 +1044,11 @@
                       (select_mode_prem thy modes' (union (op =) vs generator_vs) ps)) all_generator_vs) of
                       SOME generator_vs => check_mode_prems ((map (generator vTs) generator_vs) @ acc_ps)
                         (union (op =) vs generator_vs) ps
-                    | NONE => let
-                    val _ = tracing ("ps:" ^ (commas
-                    (map (fn p => string_of_moded_prem thy (p, Mode (([], []), [], []))) ps)))
-                  in (*error "mode analysis failed"*)NONE end
+                    | 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) 
+        | SOME (p, SOME mode) => check_mode_prems ((p, mode) :: acc_ps) 
             (case p of Prem (us, _) => union (op =) vs (terms_vs us) | _ => vs)
             (filter_out (equal p) ps))
     val (in_ts, in_ts') = List.partition (is_constrt thy) (fst (split_smode is ts));
@@ -1083,33 +1067,41 @@
     else NONE
   end;
 
-fun check_modes_pred with_generator thy param_vs clauses modes gen_modes (p, ms) =
-  let val SOME rs = AList.lookup (op =) clauses p
+fun print_failed_mode options thy modes p m rs i =
+  if show_mode_inference options then
+    let
+      val _ = Output.tracing ("Clause " ^ string_of_int (i + 1) ^ " of " ^
+      p ^ " violates mode " ^ string_of_mode m)
+      val _ = Output.tracing (string_of_clause thy p (nth rs i))
+    in () end
+  else ()
+
+fun check_modes_pred options with_generator thy param_vs clauses modes gen_modes (p, ms) =
+  let
+    val rs = case AList.lookup (op =) clauses p of SOME rs => rs | NONE => []
   in (p, List.filter (fn m => case find_index
     (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 " ^
-      p ^ " violates mode " ^ string_of_mode m);
-        tracing (commas (map (Syntax.string_of_term_global thy) (fst (nth rs i)))); false)) ms)
+    | i => (print_failed_mode options thy modes p m rs i; false)) ms)
   end;
 
 fun get_modes_pred with_generator thy param_vs clauses modes gen_modes (p, ms) =
   let
-    val SOME rs = AList.lookup (op =) clauses p 
+    val rs = case AList.lookup (op =) clauses p of SOME rs => rs | NONE => []
   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 all_modes param_vs clauses =
+fun infer_modes options thy extra_modes all_modes param_vs clauses =
   let
     val modes =
       fixp (fn modes =>
-        map (check_modes_pred false thy param_vs clauses (modes @ extra_modes) []) modes)
+        map (check_modes_pred options false thy param_vs clauses (modes @ extra_modes) []) modes)
           all_modes
   in
     map (get_modes_pred false thy param_vs clauses (modes @ extra_modes) []) modes
@@ -1122,19 +1114,24 @@
     | SOME vs' => (k, subtract (op =) vs' vs))
     :: remove_from rem xs
     
-fun infer_modes_with_generator thy extra_modes all_modes param_vs clauses =
+fun infer_modes_with_generator options thy extra_modes all_modes param_vs clauses =
   let
     val prednames = map fst clauses
-    val extra_modes = all_modes_of thy
+    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 all_modes 
+    val starting_modes = remove_from extra_modes' all_modes
+    fun eq_mode (m1, m2) = (m1 = m2)
     val modes =
       fixp (fn modes =>
-        map (check_modes_pred true thy param_vs clauses extra_modes (gen_modes @ modes)) modes)
-         starting_modes 
+        map (check_modes_pred options true thy param_vs clauses extra_modes' (gen_modes @ modes)) modes)
+         starting_modes
   in
-    map (get_modes_pred true thy param_vs clauses extra_modes (gen_modes @ modes)) modes
+    AList.join (op =)
+    (fn _ => fn ((mps1, mps2)) =>
+      merge (fn ((m1, _), (m2, _)) => eq_mode (m1, m2)) (mps1, mps2))
+    (infer_modes options thy extra_modes all_modes param_vs clauses,
+    map (get_modes_pred true thy param_vs clauses extra_modes (gen_modes @ modes)) modes)
   end;
 
 (* term construction *)
@@ -1157,136 +1154,10 @@
       in (t' $ u', nvs'') end
   | distinct_v x nvs = (x, nvs);
 
-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 = mk_tupleT (map fastype_of out_ts);
-    val U = fastype_of success_t;
-    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) DatatypeCase.Quiet [] v
-      [(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_bot compfuns U'),
-       (v', mk_bot compfuns U')]))
-  end;
-
-(*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) = split_mode is vs    
-        val (f, args) = strip_comb u
-        val (params, args') = chop (length ms) args
-        val (inargs, outargs) = split_mode is' args'
-        val b = length vs
-        val perm = map (fn i => (find_index_eq (Bound (b - i)) args') + 1) (1 upto b)
-        val outp_perm =
-          snd (split_mode is perm)
-          |> map (fn i => i - length (filter (fn x => x < i) is'))
-        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
-                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, 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 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 compfuns (f_app, match_t))
-      end
-  | compile_param_ext _ _ _ _ = error "compile params"
-*)
-
-fun compile_param neg_in_sizelim size thy compfuns (NONE, t) = t
-  | compile_param neg_in_sizelim 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 neg_in_sizelim 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) =>
-         case neg_in_sizelim of
-           SOME _ =>  Free (name, sizelim_funT_of compfuns (iss, is') T)
-         | NONE => Free (name, funT_of compfuns (iss, is') T)
-           
-       | _ => error ("PredicateCompiler: illegal parameter term")
-   in
-     (case neg_in_sizelim of SOME size_t =>
-       (fn t =>
-       let
-         val Ts = fst (split_last (binder_types (fastype_of t)))
-         val names = map (fn i => "x" ^ string_of_int i) (1 upto length Ts)
-       in
-         list_abs (names ~~ Ts, list_comb (t, (map Bound ((length Ts) - 1 downto 0)) @ [size_t]))
-       end)
-     | NONE => I)
-     (list_comb (f', params' @ args'))
-   end
-
-fun compile_expr neg_in_sizelim size thy ((Mode (mode, is, ms)), t) =
-  case strip_comb t of
-    (Const (name, T), params) =>
-       let
-         val params' = map (compile_param neg_in_sizelim 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) inargs =
-  case strip_comb t of
-    (Const (name, T), params) =>
-      let
-        val params' = map (compile_param NONE size thy PredicateCompFuns.compfuns) (ms ~~ params)
-      in
-        list_comb (mk_generator_of compfuns thy (name, T) mode, params' @ inargs)
-      end
-    | (Free (name, T), params) =>
-    lift_pred compfuns
-    (list_comb (Free (name, sizelim_funT_of PredicateCompFuns.compfuns ([], is) T), params @ inargs))
-      
-          
 (** specific rpred functions -- move them to the correct place in this file *)
 
-fun mk_Eval_of size ((x, T), NONE) names = (x, names)
-  | mk_Eval_of size ((x, T), SOME mode) names =
+fun mk_Eval_of additional_arguments ((x, T), NONE) names = (x, names)
+  | mk_Eval_of additional_arguments ((x, T), SOME mode) names =
 	let
     val Ts = binder_types T
     (*val argnames = Name.variant_list names
@@ -1331,32 +1202,99 @@
 			end
 		val (inoutargs, args) = split_list (map mk_arg (1 upto (length Ts) ~~ Ts))
     val (inargs, outargs) = pairself flat (split_list inoutargs)
-    val size_t = case size of NONE => [] | SOME size_t => [size_t]
-		val r = PredicateCompFuns.mk_Eval (list_comb (x, inargs @ size_t), mk_tuple outargs)
+		val r = PredicateCompFuns.mk_Eval 
+      (list_comb (x, inargs @ additional_arguments), HOLogic.mk_tuple outargs)
     val t = fold_rev mk_split_lambda args r
   in
     (t, names)
   end;
 
-fun compile_arg size thy param_vs iss arg = 
+fun compile_arg compilation_modifiers compfuns additional_arguments thy param_vs iss arg = 
   let
-    val funT_of = case size of NONE => funT_of | SOME _ => sizelim_funT_of
     fun map_params (t as Free (f, T)) =
       if member (op =) param_vs f then
         case (the (AList.lookup (op =) (param_vs ~~ iss) f)) of
-          SOME is => let val T' = funT_of PredicateCompFuns.compfuns ([], is) T
-            in fst (mk_Eval_of size ((Free (f, T'), T), SOME is) []) end
+          SOME is =>
+            let
+              val T' = #funT_of compilation_modifiers compfuns ([], is) T
+            in fst (mk_Eval_of additional_arguments ((Free (f, T'), T), SOME is) []) end
         | NONE => t
       else t
       | map_params t = t
     in map_aterms map_params arg end
-  
-fun compile_clause compfuns size final_term thy all_vs param_vs (iss, is) inp (ts, moded_ps) =
+
+fun compile_match compilation_modifiers compfuns additional_arguments param_vs iss thy eqs eqs' out_ts success_t =
+  let
+    val eqs'' = maps mk_eq eqs @ eqs'
+    val eqs'' =
+      map (compile_arg compilation_modifiers compfuns additional_arguments thy param_vs iss) 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 U = fastype_of success_t;
+    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) DatatypeCase.Quiet [] v
+      [(HOLogic.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_bot compfuns U'),
+       (v', mk_bot compfuns U')]))
+  end;
+
+(*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 compilation_modifiers compfuns thy (NONE, t) = t
+  | compile_param compilation_modifiers compfuns thy (m as SOME (Mode (mode, _, ms)), t) =
+   let
+     val (f, args) = strip_comb (Envir.eta_contract t)
+     val (params, args') = chop (length ms) args
+     val params' = map (compile_param compilation_modifiers compfuns thy) (ms ~~ params)
+     val f' =
+       case f of
+         Const (name, T) => Const (#const_name_of compilation_modifiers thy name mode,
+           #funT_of compilation_modifiers compfuns mode T)
+       | Free (name, T) => Free (name, #funT_of compilation_modifiers compfuns mode T)
+       | _ => error ("PredicateCompiler: illegal parameter term")
+   in
+     list_comb (f', params' @ args')
+   end
+
+fun compile_expr compilation_modifiers compfuns thy ((Mode (mode, _, ms)), t) inargs additional_arguments =
+  case strip_comb t of
+    (Const (name, T), params) =>
+       let
+         val params' = map (compile_param compilation_modifiers compfuns thy) (ms ~~ params)
+           (*val mk_fun_of = if depth_limited then mk_depth_limited_fun_of else mk_fun_of*)
+         val name' = #const_name_of compilation_modifiers thy name mode
+         val T' = #funT_of compilation_modifiers compfuns mode T
+       in
+         (list_comb (Const (name', T'), params' @ inargs @ additional_arguments))
+       end
+  | (Free (name, T), params) =>
+    list_comb (Free (name, #funT_of compilation_modifiers compfuns mode T), params @ inargs @ additional_arguments)
+
+fun compile_clause compilation_modifiers compfuns thy all_vs param_vs additional_arguments (iss, is) inp (ts, moded_ps) =
+  let
+    val compile_match = compile_match compilation_modifiers compfuns additional_arguments param_vs iss thy
     fun check_constrt t (names, eqs) =
       if is_constrt thy t then (t, (names, eqs)) else
         let
-          val s = Name.variant names "x";
+          val s = Name.variant names "x"
           val v = Free (s, fastype_of t)
         in (v, (s::names, HOLogic.mk_eq (v, t)::eqs)) end;
 
@@ -1371,12 +1309,8 @@
             val (out_ts''', (names'', constr_vs)) = fold_map distinct_v
               out_ts'' (names', map (rpair []) vs);
           in
-          (* 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)
+            compile_match constr_vs (eqs @ eqs') out_ts'''
+              (mk_single compfuns (HOLogic.mk_tuple out_ts))
           end
       | compile_prems out_ts vs names ((p, mode as Mode ((_, is), _, _)) :: ps) =
           let
@@ -1385,16 +1319,16 @@
               fold_map check_constrt out_ts (names, [])
             val (out_ts'', (names'', constr_vs')) = fold_map distinct_v
               out_ts' ((names', map (rpair []) vs))
+            val additional_arguments' =
+              #transform_additional_arguments compilation_modifiers p additional_arguments
             val (compiled_clause, rest) = case p of
                Prem (us, t) =>
                  let
                    val (in_ts, out_ts''') = split_smode is us;
-                   val in_ts = map (compile_arg size thy param_vs iss) in_ts
-                   val args = case size of
-                     NONE => in_ts
-                   | SOME size_t => in_ts @ [size_t]
-                   val u = lift_pred compfuns
-                     (list_comb (compile_expr NONE size thy (mode, t), args))                     
+                   val in_ts = map (compile_arg compilation_modifiers compfuns additional_arguments
+                     thy param_vs iss) in_ts
+                   val u =
+                     compile_expr compilation_modifiers compfuns thy (mode, t) in_ts additional_arguments'
                    val rest = compile_prems out_ts''' vs' names'' ps
                  in
                    (u, rest)
@@ -1402,38 +1336,32 @@
              | Negprem (us, t) =>
                  let
                    val (in_ts, out_ts''') = split_smode is us
-                   val u = lift_pred compfuns
-                     (mk_not PredicateCompFuns.compfuns (list_comb (compile_expr size NONE thy (mode, t), in_ts)))
+                   val in_ts = map (compile_arg compilation_modifiers compfuns additional_arguments
+                     thy param_vs iss) in_ts
+                   val u = mk_not compfuns
+                     (compile_expr compilation_modifiers compfuns thy (mode, t) in_ts additional_arguments')
                    val rest = compile_prems out_ts''' vs' names'' ps
                  in
                    (u, rest)
                  end
              | Sidecond t =>
                  let
+                   val t = compile_arg compilation_modifiers compfuns additional_arguments
+                     thy param_vs iss t
                    val rest = compile_prems [] vs' names'' ps;
                  in
                    (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 = 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 (the size))
+                   val [size] = additional_arguments
+                   val u = lift_random (HOLogic.mk_random T size)
                    val rest = compile_prems [Free (v, T)]  vs' names'' ps;
                  in
                    (u, rest)
                  end
           in
-            compile_match thy compfuns constr_vs' eqs out_ts'' 
+            compile_match constr_vs' eqs out_ts''
               (mk_bind compfuns (compiled_clause, rest))
           end
     val prem_t = compile_prems in_ts' param_vs all_vs' moded_ps;
@@ -1441,14 +1369,13 @@
     mk_bind compfuns (mk_single compfuns inp, prem_t)
   end
 
-fun compile_pred compfuns mk_fun_of use_size thy all_vs param_vs s T mode moded_cls =
+fun compile_pred compilation_modifiers compfuns thy all_vs param_vs s T mode moded_cls =
   let
 	  val (Ts1, Ts2) = chop (length (fst mode)) (binder_types T)
     val (Us1, Us2) = split_smodeT (snd mode) Ts2
-    val funT_of = if use_size then sizelim_funT_of else funT_of
-    val Ts1' = map2 (fn NONE => I | SOME is => funT_of PredicateCompFuns.compfuns ([], is)) (fst mode) Ts1
-    val size_name = Name.variant (all_vs @ param_vs) "size"
-  	fun mk_input_term (i, NONE) =
+    val Ts1' =
+      map2 (fn NONE => I | SOME is => #funT_of compilation_modifiers compfuns ([], is)) (fst mode) Ts1
+    fun mk_input_term (i, NONE) =
 		    [Free (Name.variant (all_vs @ param_vs) ("x" ^ string_of_int i), nth Ts2 (i - 1))]
 		  | mk_input_term (i, SOME pis) = case HOLogic.strip_tupleT (nth Ts2 (i - 1)) of
 						   [] => error "strange unit input"
@@ -1461,30 +1388,22 @@
 						   else [HOLogic.mk_tuple (map Free (vnames ~~ map (fn j => nth Ts (j - 1)) pis))] end
 		val in_ts = maps mk_input_term (snd mode)
     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 additional_arguments = #additional_arguments compilation_modifiers (all_vs @ param_vs)
     val cl_ts =
-      map (compile_clause compfuns decr_size (fn out_ts => mk_single compfuns (mk_tuple out_ts))
-        thy all_vs param_vs mode (mk_tuple in_ts)) 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 @ in_ts @ [size]), size_t)
-    else
-      (list_comb (fun_const, params @ in_ts), t)
+      map (compile_clause compilation_modifiers compfuns
+        thy all_vs param_vs additional_arguments mode (HOLogic.mk_tuple in_ts)) moded_cls;
+    val compilation = #wrap_compilation compilation_modifiers compfuns s T mode additional_arguments
+      (if null cl_ts then
+        mk_bot compfuns (HOLogic.mk_tupleT Us2)
+      else foldr1 (mk_sup compfuns) cl_ts)
+    val fun_const =
+      Const (#const_name_of compilation_modifiers thy s mode,
+        #funT_of compilation_modifiers compfuns mode T)
   in
-    HOLogic.mk_Trueprop (HOLogic.mk_eq eq)
+    HOLogic.mk_Trueprop
+      (HOLogic.mk_eq (list_comb (fun_const, params @ in_ts @ additional_arguments), compilation))
   end;
-  
+
 (* special setup for simpset *)                  
 val HOL_basic_ss' = HOL_basic_ss addsimps (@{thms "HOL.simp_thms"} @ [@{thm Pair_eq}])
   setSolver (mk_solver "all_tac_solver" (fn _ => fn _ => all_tac))
@@ -1531,15 +1450,15 @@
   val param_names' = Name.variant_list (param_names @ argnames)
     (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 NONE) ((params ~~ Ts1) ~~ iss) []
+  val (params', names) = fold_map (mk_Eval_of []) ((params ~~ Ts1) ~~ iss) []
   val predpropI = HOLogic.mk_Trueprop (list_comb (pred, param_vs @ args))
   val predpropE = HOLogic.mk_Trueprop (list_comb (pred, params' @ args))
   val param_eqs = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (param_vs ~~ params')
   val funargs = params @ inargs
   val funpropE = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, funargs),
-                  if null outargs then Free("y", HOLogic.unitT) else mk_tuple outargs))
+                  if null outargs then Free("y", HOLogic.unitT) else HOLogic.mk_tuple outargs))
   val funpropI = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, funargs),
-                   mk_tuple outargs))
+                   HOLogic.mk_tuple outargs))
   val introtrm = Logic.list_implies (predpropI :: param_eqs, funpropI)
   val simprules = [defthm, @{thm eval_pred},
 	  @{thm "split_beta"}, @{thm "fst_conv"}, @{thm "snd_conv"}, @{thm pair_collapse}]
@@ -1601,7 +1520,7 @@
       val (Ts1, Ts2) = chop (length iss) Ts
       val (Us1, Us2) =  split_smodeT 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 funT = (Ts1' @ Us1) ---> (mk_predT compfuns (HOLogic.mk_tupleT Us2))
       val names = Name.variant_list []
         (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
 			(* old *)
@@ -1634,7 +1553,7 @@
    	  val (xinoutargs, names) = fold_map mk_vars ((1 upto (length Ts2)) ~~ Ts2) param_names
       val (xinout, xargs) = split_list xinoutargs
 			val (xins, xouts) = pairself flat (split_list xinout)
-			val (xparams', names') = fold_map (mk_Eval_of NONE) ((xparams ~~ Ts1) ~~ iss) names
+			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 =
@@ -1663,16 +1582,16 @@
     fold create_definition modes thy
   end;
 
-fun sizelim_create_definitions preds (name, modes) thy =
+fun create_definitions_of_depth_limited_functions 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
+        val mode_cname = create_constname_of_mode thy "depth_limited_" name mode
+        val funT = depth_limited_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 
+        |> set_depth_limited_function_name name mode mode_cname 
       end;
   in
     fold create_definition modes thy
@@ -1682,9 +1601,10 @@
   let
     val Ts = binder_types T
     val (paramTs, (inargTs, outargTs)) = split_modeT (iss, is) Ts
-    val paramTs' = map2 (fn SOME is => sizelim_funT_of PredicateCompFuns.compfuns ([], is) | NONE => I) iss paramTs 
+    val paramTs' = map2 (fn SOME is => generator_funT_of ([], is) | NONE => I) iss paramTs
   in
-    (paramTs' @ inargTs @ [@{typ "code_numeral"}]) ---> (mk_predT RPredCompFuns.compfuns (mk_tupleT outargTs))
+    (paramTs' @ inargTs @ [@{typ code_numeral}]) --->
+      (mk_predT RPredCompFuns.compfuns (HOLogic.mk_tupleT outargTs))
   end
 
 fun rpred_create_definitions preds (name, modes) thy =
@@ -1696,7 +1616,7 @@
         val funT = generator_funT_of 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 
+        |> set_generator_name name mode mode_cname
       end;
   in
     fold create_definition modes thy
@@ -1818,7 +1738,7 @@
     (* need better control here! *)
   end
 
-fun prove_clause thy nargs modes (iss, is) (_, clauses) (ts, moded_ps) =
+fun prove_clause options 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 [] =
@@ -1874,7 +1794,8 @@
       end;
     val prems_tac = prove_prems in_ts moded_ps
   in
-    rtac @{thm bindI} 1
+    print_tac' options "Proving clause..."
+    THEN rtac @{thm bindI} 1
     THEN rtac @{thm singleI} 1
     THEN prems_tac
   end;
@@ -1883,20 +1804,20 @@
   | select_sup _ 1 = [rtac @{thm supI1}]
   | select_sup n i = (rtac @{thm supI2})::(select_sup (n - 1) (i - 1));
 
-fun prove_one_direction thy clauses preds modes pred mode moded_clauses =
+fun prove_one_direction options 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 print_tac "before applying elim rule"
+		THEN print_tac' options "before applying elim rule"
     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 (EVERY (map2 (prove_clause options thy nargs modes mode) clauses moded_clauses))
     THEN print_tac "proved one direction"
   end;
 
@@ -1914,15 +1835,20 @@
           | _ => PureThy.get_thms thy (((fst o dest_Type o fastype_of) t) ^ ".split_asm")
         val (_, ts) = strip_comb t
       in
-        (Splitter.split_asm_tac split_rules 1)
-(*        THEN (Simplifier.asm_full_simp_tac HOL_basic_ss 1)
-          THEN (DETERM (TRY (etac @{thm Pair_inject} 1))) *)
-        THEN (REPEAT_DETERM_N (num_of_constrs - 1) (etac @{thm botE} 1 ORELSE etac @{thm botE} 2))
+        (print_tac ("Term " ^ (Syntax.string_of_term_global thy t) ^ 
+          "splitting with rules \n" ^
+        commas (map (Display.string_of_thm_global thy) split_rules)))
+        THEN TRY ((Splitter.split_asm_tac split_rules 1)
+        THEN (print_tac "after splitting with split_asm rules")
+        (* THEN (Simplifier.asm_full_simp_tac HOL_basic_ss 1)
+          THEN (DETERM (TRY (etac @{thm Pair_inject} 1)))*)
+          THEN (REPEAT_DETERM_N (num_of_constrs - 1) (etac @{thm botE} 1 ORELSE etac @{thm botE} 2)))
+        THEN (assert_tac (Max_number_of_subgoals 2))
         THEN (EVERY (map split_term_tac ts))
       end
     else all_tac
   in
-    split_term_tac (mk_tuple out_ts)
+    split_term_tac (HOLogic.mk_tuple out_ts)
     THEN (DETERM (TRY ((Splitter.split_asm_tac [@{thm "split_if_asm"}] 1) THEN (etac @{thm botE} 2))))
   end
 
@@ -2053,7 +1979,7 @@
     THEN prems_tac
   end;
  
-fun prove_other_direction thy modes pred mode moded_clauses =
+fun prove_other_direction options 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)
@@ -2063,26 +1989,28 @@
      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))))
+     THEN (if null moded_clauses then
+         etac @{thm botE} 1
+       else EVERY (map2 prove_clause moded_clauses (1 upto (length moded_clauses))))
   end;
 
 (** proof procedure **)
 
-fun prove_pred thy clauses preds modes pred mode (moded_clauses, compiled_term) =
+fun prove_pred options thy clauses preds modes pred mode (moded_clauses, compiled_term) =
   let
     val ctxt = ProofContext.init thy
-    val clauses = the (AList.lookup (op =) clauses pred)
+    val clauses = case AList.lookup (op =) clauses pred of SOME rs => rs | NONE => []
   in
     Goal.prove ctxt (Term.add_free_names compiled_term []) [] compiled_term
-      (if !do_proofs then
+      (if not (skip_proof options) then
         (fn _ =>
         rtac @{thm pred_iffI} 1
-        THEN print_tac "after pred_iffI"
-        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 _ => Skip_Proof.cheat_tac thy)
+				THEN print_tac' options "after pred_iffI"
+        THEN prove_one_direction options thy clauses preds modes pred mode moded_clauses
+        THEN print_tac' options "proved one direction"
+        THEN prove_other_direction options thy modes pred mode moded_clauses
+        THEN print_tac' options "proved other direction")
+      else (fn _ => Skip_Proof.cheat_tac thy))
   end;
 
 (* composition of mode inference, definition, compilation and proof *)
@@ -2101,48 +2029,57 @@
   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)
+fun compile_preds comp_modifiers compfuns thy all_vs param_vs preds moded_clauses =
+  map_preds_modes (fn pred => compile_pred comp_modifiers compfuns thy all_vs param_vs pred
+      (the (AList.lookup (op =) preds pred))) moded_clauses
+
+fun prove options thy clauses preds modes moded_clauses compiled_terms =
+  map_preds_modes (prove_pred options thy clauses preds modes)
     (join_preds_modes moded_clauses compiled_terms)
 
-fun prove_by_skip thy _ _ _ _ compiled_terms =
+fun prove_by_skip options thy _ _ _ _ compiled_terms =
   map_preds_modes (fn pred => fn mode => fn t => Drule.standard (Skip_Proof.make_thm thy t))
     compiled_terms
+
+fun dest_prem thy params 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 thy params 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))
+  | (c as Const (s, _), ts) =>
+    if is_registered thy s then
+      let val (ts1, ts2) = chop (nparams_of thy s) ts
+      in Prem (ts2, list_comb (c, ts1)) end
+    else Sidecond t
+  | _ => Sidecond t)
     
-fun prepare_intrs thy prednames =
+fun prepare_intrs thy prednames intros =
   let
-    val intrs = maps (intros_of thy) prednames
-      |> map (Logic.unvarify o prop_of)
+    val intrs = map prop_of intros
     val nparams = nparams_of thy (hd prednames)
+    val preds = map (fn c => Const (c, Sign.the_const_type thy c)) prednames
+    val (preds, intrs) = unify_consts thy preds intrs
+    val ([preds, intrs], _) = fold_burrow (Variable.import_terms false) [preds, intrs] (ProofContext.init thy)
+    val preds = map dest_Const preds
     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 _ $ u = Logic.strip_imp_concl (hd intrs);
-    val params = List.take (snd (strip_comb u), nparams);
+    val params = case intrs of
+        [] =>
+          let
+            val (paramTs, _) = chop nparams (binder_types (snd (hd preds)))
+            val param_names = Name.variant_list [] (map (fn i => "p" ^ string_of_int i) (1 upto length paramTs))
+          in map Free (param_names ~~ paramTs) end
+      | intr :: _ => fst (chop nparams
+        (snd (strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl intr)))))
     val param_vs = maps term_vs params
     val all_vs = terms_vs intrs
-    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          
-          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))
-      | (c as Const (s, _), ts) =>
-        if is_registered thy s then
-          let val (ts1, ts2) = chop (nparams_of thy s) ts
-          in Prem (ts2, list_comb (c, ts1)) end
-        else Sidecond t
-      | _ => Sidecond t)
     fun add_clause intr (clauses, arities) =
     let
       val _ $ t = Logic.strip_imp_concl intr;
       val (Const (name, T), ts) = strip_comb t;
       val (ts1, ts2) = chop nparams ts;
-      val prems = map (dest_prem o HOLogic.dest_Trueprop) (Logic.strip_imp_prems intr);
+      val prems = map (dest_prem thy params o HOLogic.dest_Trueprop) (Logic.strip_imp_prems intr);
       val (Ts, Us) = chop nparams (binder_types T)
     in
       (AList.update op = (name, these (AList.lookup op = clauses name) @
@@ -2177,31 +2114,74 @@
     val all_modes = map (fn (s, T) => (s, modes_of_typ T)) preds
   in (preds, nparams, all_vs, param_vs, extra_modes, clauses, all_modes) end;
 
+fun check_format_of_intro_rule thy intro =
+  let
+    val concl = Logic.strip_imp_concl (prop_of intro)
+    val (p, args) = strip_comb (HOLogic.dest_Trueprop concl)
+    val params = List.take (args, nparams_of thy (fst (dest_Const p)))
+    fun check_arg arg = case HOLogic.strip_tupleT (fastype_of arg) of
+      (Ts as _ :: _ :: _) =>
+        if (length (HOLogic.strip_tuple arg) = length Ts) then true
+        else
+        error ("Format of introduction rule is invalid: tuples must be expanded:"
+        ^ (Syntax.string_of_term_global thy arg) ^ " in " ^
+        (Display.string_of_thm_global thy intro)) 
+      | _ => true
+    val prems = Logic.strip_imp_prems (prop_of intro)
+    fun check_prem (Prem (args, _)) = forall check_arg args
+      | check_prem (Negprem (args, _)) = forall check_arg args
+      | check_prem _ = true
+  in
+    forall check_arg args andalso
+    forall (check_prem o dest_prem thy params o HOLogic.dest_Trueprop) prems
+  end
+
+(*
+fun check_intros_elim_match thy prednames =
+  let
+    fun check predname =
+      let
+        val intros = intros_of thy predname
+        val elim = the_elim_of thy predname
+        val nparams = nparams_of thy predname
+        val elim' =
+          (Drule.standard o (Skip_Proof.make_thm thy))
+          (mk_casesrule (ProofContext.init thy) nparams intros)
+      in
+        if not (Thm.equiv_thm (elim, elim')) then
+          error "Introduction and elimination rules do not match!"
+        else true
+      end
+  in forall check prednames end
+*)
+
 (** main function of predicate compiler **)
 
-fun add_equations_of steps prednames thy =
+fun add_equations_of steps options prednames thy =
   let
-    val _ = tracing ("Starting predicate compiler for predicates " ^ commas prednames ^ "...")
+    val _ = print_step options ("Starting predicate compiler for predicates " ^ commas prednames ^ "...")
     val _ = tracing (commas (map (Display.string_of_thm_global thy) (maps (intros_of thy) prednames)))
+      (*val _ = check_intros_elim_match thy prednames*)
+      (*val _ = map (check_format_of_intro_rule thy) (maps (intros_of thy) prednames)*)
     val (preds, nparams, all_vs, param_vs, extra_modes, clauses, all_modes) =
-      prepare_intrs thy prednames
-    val _ = tracing "Infering modes..."
-    val moded_clauses = #infer_modes steps thy extra_modes all_modes param_vs clauses 
+      prepare_intrs thy prednames (maps (intros_of thy) prednames)
+    val _ = print_step options "Infering modes..."
+    val moded_clauses = #infer_modes steps options thy extra_modes all_modes param_vs clauses 
     val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses
+    val _ = check_expected_modes options modes
     val _ = print_modes modes
-    val _ = print_moded_clauses thy moded_clauses
-    val _ = tracing "Defining executable functions..."
+      (*val _ = print_moded_clauses thy moded_clauses*)
+    val _ = print_step options "Defining executable functions..."
     val thy' = fold (#create_definitions steps preds) modes thy
       |> Theory.checkpoint
-    val _ = tracing "Compiling equations..."
+    val _ = print_step options "Compiling equations..."
     val compiled_terms =
       (#compile_preds steps) thy' all_vs param_vs preds moded_clauses
-    val _ = print_compiled_terms thy' compiled_terms
-    val _ = tracing "Proving equations..."
-    val result_thms = #prove steps thy' clauses preds (extra_modes @ modes)
+    val _ = print_compiled_terms options thy' compiled_terms
+    val _ = print_step options "Proving equations..."
+    val result_thms = #prove steps options 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
@@ -2226,7 +2206,7 @@
 
 fun extend value_of edges_of key G = fst (extend' value_of edges_of key (G, [])) 
   
-fun gen_add_equations steps names thy =
+fun gen_add_equations steps options names thy =
   let
     val thy' = PredData.map (fold (extend (fetch_pred_data thy) (depending_preds_of thy)) names) thy
       |> Theory.checkpoint;
@@ -2235,33 +2215,83 @@
     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)
+        if #are_not_defined steps thy preds then
+          add_equations_of steps options preds thy else thy)
       scc thy' |> Theory.checkpoint
   in thy'' end
 
 (* different instantiantions of the predicate compiler *)
 
+val predicate_comp_modifiers =
+  {const_name_of = predfun_name_of,
+  funT_of = funT_of,
+  additional_arguments = K [],
+  wrap_compilation = K (K (K (K (K I)))),
+  transform_additional_arguments = K I
+  }
+
+val depth_limited_comp_modifiers =
+  {const_name_of = depth_limited_function_name_of,
+  funT_of = depth_limited_funT_of,
+  additional_arguments = fn names =>
+    let
+      val [depth_name, polarity_name] = Name.variant_list names ["depth", "polarity"]
+    in [Free (polarity_name, @{typ "bool"}), Free (depth_name, @{typ "code_numeral"})] end,
+  wrap_compilation =
+    fn compfuns => fn s => fn T => fn mode => fn additional_arguments => fn compilation =>
+    let
+      val [polarity, depth] = additional_arguments
+      val (_, Ts2) = chop (length (fst mode)) (binder_types T)
+      val (_, Us2) = split_smodeT (snd mode) Ts2
+      val T' = mk_predT compfuns (HOLogic.mk_tupleT Us2)
+      val if_const = Const (@{const_name "If"}, @{typ bool} --> T' --> T' --> T')
+      val full_mode = null Us2
+    in
+      if_const $ HOLogic.mk_eq (depth, @{term "0 :: code_numeral"})
+        $ (if_const $ polarity $ mk_bot compfuns (dest_predT compfuns T')
+          $ (if full_mode then mk_single compfuns HOLogic.unit else Const (@{const_name undefined}, T')))
+        $ compilation
+    end,
+  transform_additional_arguments =
+    fn prem => fn additional_arguments =>
+    let
+      val [polarity, depth] = additional_arguments
+      val polarity' = (case prem of Prem _ => I | Negprem _ => HOLogic.mk_not | _ => I) polarity
+      val depth' =
+        Const ("HOL.minus_class.minus", @{typ "code_numeral => code_numeral => code_numeral"})
+          $ depth $ Const ("HOL.one_class.one", @{typ "Code_Numeral.code_numeral"})
+    in [polarity', depth'] end
+  }
+
+val rpred_comp_modifiers =
+  {const_name_of = generator_name_of,
+  funT_of = K generator_funT_of,
+  additional_arguments = fn names => [Free (Name.variant names "size", @{typ code_numeral})],
+  wrap_compilation = K (K (K (K (K I)))),
+  transform_additional_arguments = K I
+  }
+
+
 val add_equations = gen_add_equations
   {infer_modes = infer_modes,
   create_definitions = create_definitions,
-  compile_preds = compile_preds PredicateCompFuns.compfuns mk_fun_of false,
+  compile_preds = compile_preds predicate_comp_modifiers PredicateCompFuns.compfuns,
   prove = prove,
   are_not_defined = fn thy => forall (null o modes_of thy),
   qname = "equation"}
 
-val add_sizelim_equations = gen_add_equations
+val add_depth_limited_equations = gen_add_equations
   {infer_modes = infer_modes,
-  create_definitions = sizelim_create_definitions,
-  compile_preds = compile_preds PredicateCompFuns.compfuns mk_sizelim_fun_of true,
+  create_definitions = create_definitions_of_depth_limited_functions,
+  compile_preds = compile_preds depth_limited_comp_modifiers PredicateCompFuns.compfuns,
   prove = prove_by_skip,
-  are_not_defined = fn thy => forall (null o sizelim_modes_of thy),
-  qname = "sizelim_equation"
-  }
+  are_not_defined = fn thy => forall (null o depth_limited_modes_of thy),
+  qname = "depth_limited_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,
+  compile_preds = compile_preds rpred_comp_modifiers RPredCompFuns.compfuns,
   prove = prove_by_skip,
   are_not_defined = fn thy => forall (null o rpred_modes_of thy),
   qname = "rpred_equation"}
@@ -2283,15 +2313,11 @@
 val setup = PredData.put (Graph.empty) #>
   Attrib.setup @{binding code_pred_intros} (Scan.succeed (attrib add_intro))
     "adding alternative introduction rules for code generation of inductive predicates"
-(*  Attrib.setup @{binding code_ind_cases} (Scan.succeed add_elim_attrib)
-    "adding alternative elimination rules for code generation of inductive predicates";
-    *)
   (*FIXME name discrepancy in attribs and ML code*)
   (*FIXME intros should be better named intro*)
-  (*FIXME why distinguished attribute for cases?*)
 
 (* TODO: make TheoryDataFun to GenericDataFun & remove duplication of local theory and theory *)
-fun generic_code_pred prep_const rpred raw_const lthy =
+fun generic_code_pred prep_const options raw_const lthy =
   let
     val thy = ProofContext.theory_of lthy
     val const = prep_const thy raw_const
@@ -2302,9 +2328,11 @@
     val preds = Graph.all_preds (PredData.get thy') [const] |> filter_out (has_elim thy')
     fun mk_cases const =
       let
+        val T = Sign.the_const_type thy const
+        val pred = Const (const, T)
         val nparams = nparams_of thy' const
         val intros = intros_of thy' const
-      in mk_casesrule lthy' nparams intros end  
+      in mk_casesrule lthy' pred nparams intros end  
     val cases_rules = map mk_cases preds
     val cases =
       map (fn case_rule => RuleCases.Case {fixes = [],
@@ -2320,11 +2348,14 @@
           (ProofContext.init (ProofContext.theory_of goal_ctxt)) (map the_single thms)
       in
         goal_ctxt |> LocalTheory.theory (fold set_elim global_thms #>
-          (if rpred then
-            (add_equations [const] #>
-             add_sizelim_equations [const] #> add_quickcheck_equations [const])
-        else add_equations [const]))
-      end  
+          (if is_rpred options then
+            (add_equations options [const] #>
+            add_quickcheck_equations options [const])
+           else if is_depth_limited options then
+             add_depth_limited_equations options [const]
+           else
+             add_equations options [const]))
+      end
   in
     Proof.theorem_i NONE after_qed (map (single o (rpair [])) cases_rules) lthy''
   end;
@@ -2335,9 +2366,11 @@
 (* transformation for code generation *)
 
 val eval_ref = Unsynchronized.ref (NONE : (unit -> term Predicate.pred) option);
+val random_eval_ref = Unsynchronized.ref (NONE : (unit -> int * int -> term Predicate.pred * (int * int)) option);
 
 (*FIXME turn this into an LCF-guarded preprocessor for comprehensions*)
-fun analyze_compr thy t_compr =
+(* TODO: *)
+fun analyze_compr thy compfuns (depth_limit, random) t_compr =
   let
     val split = case t_compr of (Const (@{const_name Collect}, _) $ t) => t
       | _ => error ("Not a set comprehension: " ^ Syntax.string_of_term_global thy t_compr);
@@ -2348,6 +2381,8 @@
       (fn (i, t) => case t of Bound j => if j < length Ts then NONE
         else SOME (i+1) | _ => SOME (i+1)) args); (*FIXME dangling bounds should not occur*)
     val user_mode' = map (rpair NONE) user_mode
+    val all_modes_of = if random then all_generator_modes_of else all_modes_of
+      (*val compile_expr = if random then compile_gen_expr else compile_expr*)
     val modes = filter (fn Mode (_, is, _) => is = user_mode')
       (modes_of_term (all_modes_of thy) (list_comb (pred, params)));
     val m = case modes
@@ -2357,7 +2392,17 @@
       | m :: _ :: _ => (warning ("Multiple modes possible for comprehension "
                 ^ Syntax.string_of_term_global thy t_compr); m);
     val (inargs, outargs) = split_smode user_mode' args;
-    val t_pred = list_comb (compile_expr NONE NONE thy (m, list_comb (pred, params)), inargs);
+    val additional_arguments =
+      case depth_limit of
+        NONE => (if random then [@{term "5 :: code_numeral"}] else [])
+      | SOME d => [@{term "True"}, HOLogic.mk_number @{typ "code_numeral"} d]
+    val comp_modifiers =
+      case depth_limit of NONE => 
+      (if random then rpred_comp_modifiers else predicate_comp_modifiers) | SOME _ => depth_limited_comp_modifiers
+    val mk_fun_of = if random then mk_generator_of else
+      if (is_some depth_limit) then mk_depth_limited_fun_of else mk_fun_of
+    val t_pred = compile_expr comp_modifiers compfuns thy
+      (m, list_comb (pred, params)) inargs additional_arguments;
     val t_eval = if null outargs then t_pred else
       let
         val outargs_bounds = map (fn Bound i => i) outargs;
@@ -2370,22 +2415,30 @@
         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_map PredicateCompFuns.compfuns T_pred T_compr arrange t_pred end
+      in mk_map compfuns T_pred T_compr arrange t_pred end
   in t_eval end;
 
-fun eval thy t_compr =
+fun eval thy (options as (depth_limit, random)) t_compr =
   let
-    val t = analyze_compr thy t_compr;
-    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_Core.eval_ref", eval_ref) Predicate.map thy t' []) end;
+    val compfuns = if random then RPredCompFuns.compfuns else PredicateCompFuns.compfuns
+    val t = analyze_compr thy compfuns options t_compr;
+    val T = dest_predT compfuns (fastype_of t);
+    val t' = mk_map compfuns T HOLogic.termT (HOLogic.term_of_const T) t;
+    val eval =
+      if random then
+        Code_ML.eval NONE ("Predicate_Compile_Core.random_eval_ref", random_eval_ref)
+            (fn proc => fn g => fn s => g s |>> Predicate.map proc) thy t' []
+          |> Random_Engine.run
+      else
+        Code_ML.eval NONE ("Predicate_Compile_Core.eval_ref", eval_ref) Predicate.map thy t' []
+  in (T, eval) end;
 
-fun values ctxt k t_compr =
+fun values ctxt options k t_compr =
   let
     val thy = ProofContext.theory_of ctxt;
-    val (T, t) = eval thy t_compr;
+    val (T, ts) = eval thy options t_compr;
+    val (ts, _) = Predicate.yieldn k ts;
     val setT = HOLogic.mk_setT T;
-    val (ts, _) = Predicate.yieldn k t;
     val elemsT = HOLogic.mk_set T ts;
   in if k = ~1 orelse length ts < k then elemsT
     else Const (@{const_name Set.union}, setT --> setT --> setT) $ elemsT $ t_compr
@@ -2398,11 +2451,11 @@
   in
   end;
   *)
-fun values_cmd modes k raw_t state =
+fun values_cmd modes options k raw_t state =
   let
     val ctxt = Toplevel.context_of state;
     val t = Syntax.read_term ctxt raw_t;
-    val t' = values ctxt k t;
+    val t' = values ctxt options k t;
     val ty' = Term.type_of t';
     val ctxt' = Variable.auto_fixes t' ctxt;
     val p = PrintMode.with_modes modes (fn () =>
@@ -2410,15 +2463,24 @@
         Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) ();
   in Pretty.writeln p end;
 
-
 local structure P = OuterParse in
 
 val opt_modes = Scan.optional (P.$$$ "(" |-- P.!!! (Scan.repeat1 P.xname --| P.$$$ ")")) [];
 
+val _ = List.app OuterKeyword.keyword ["depth_limit", "random"]
+
+val options =
+  let
+    val depth_limit = Scan.optional (P.$$$ "depth_limit" |-- P.$$$ "=" |-- P.nat >> SOME) NONE
+    val random = Scan.optional (P.$$$ "random" >> K true) false
+  in
+    Scan.optional (P.$$$ "[" |-- depth_limit -- random --| P.$$$ "]") (NONE, false)
+  end
+
 val _ = OuterSyntax.improper_command "values" "enumerate and print comprehensions" OuterKeyword.diag
-  (opt_modes -- Scan.optional P.nat ~1 -- P.term
-    >> (fn ((modes, k), t) => Toplevel.no_timing o Toplevel.keep
-        (values_cmd modes k t)));
+  (opt_modes -- options -- Scan.optional P.nat ~1 -- P.term
+    >> (fn (((modes, options), k), t) => Toplevel.no_timing o Toplevel.keep
+        (values_cmd modes options k t)));
 
 end;
 
--- a/src/HOL/Tools/inductive.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/inductive.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -47,7 +47,7 @@
     (Attrib.binding * string) list ->
     (Facts.ref * Attrib.src list) list ->
     bool -> local_theory -> inductive_result * local_theory
-  val add_inductive_global: string -> inductive_flags ->
+  val add_inductive_global: serial -> inductive_flags ->
     ((binding * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list ->
     thm list -> theory -> inductive_result * theory
   val arities_of: thm -> (string * int) list
@@ -144,7 +144,7 @@
     val (tab, monos) = get_inductives ctxt;
     val space = Consts.space_of (ProofContext.consts_of ctxt);
   in
-    [Pretty.strs ("(co)inductives:" :: map #1 (NameSpace.extern_table (space, tab))),
+    [Pretty.strs ("(co)inductives:" :: map #1 (Name_Space.extern_table (space, tab))),
      Pretty.big_list "monotonicity rules:" (map (Display.pretty_thm ctxt) monos)]
     |> Pretty.chunks |> Pretty.writeln
   end;
@@ -859,7 +859,7 @@
       skip_mono = false, fork_mono = not int};
   in
     lthy
-    |> LocalTheory.set_group (serial_string ())
+    |> LocalTheory.set_group (serial ())
     |> gen_add_inductive_i mk_def flags cs (map (apfst Binding.name_of o fst) ps) intrs monos
   end;
 
--- a/src/HOL/Tools/inductive_realizer.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/inductive_realizer.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -351,7 +351,7 @@
     (** realizability predicate **)
 
     val (ind_info, thy3') = thy2 |>
-      Inductive.add_inductive_global (serial_string ())
+      Inductive.add_inductive_global (serial ())
         {quiet_mode = false, verbose = false, kind = Thm.generatedK, alt_name = Binding.empty,
           coind = false, no_elim = false, no_ind = false, skip_mono = false, fork_mono = false}
         rlzpreds rlzparams (map (fn (rintr, intr) =>
--- a/src/HOL/Tools/primrec.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/primrec.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -275,7 +275,7 @@
         [Simplifier.simp_add, Nitpick_Simps.add, Quickcheck_RecFun_Simps.add]);
   in
     lthy
-    |> set_group ? LocalTheory.set_group (serial_string ())
+    |> set_group ? LocalTheory.set_group (serial ())
     |> add_primrec_simple fixes (map snd spec)
     |-> (fn (prefix, simps) => fold_map (LocalTheory.note Thm.generatedK)
           (attr_bindings prefix ~~ simps)
--- a/src/HOL/Tools/quickcheck_generators.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/quickcheck_generators.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -317,7 +317,7 @@
       |> fold (fn (v, sort) => Vartab.update ((v, 0), sort)) raw_vs
       |> fold meet_random insts;
   in SOME (fn (v, _) => (v, (the o Vartab.lookup vtab) (v, 0)))
-  end handle CLASS_ERROR => NONE;
+  end handle Sorts.CLASS_ERROR _ => NONE;
 
 fun ensure_random_datatype config raw_tycos thy =
   let
--- a/src/HOL/Tools/record.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/record.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -1810,7 +1810,7 @@
 
 fun record_definition (args, bname) parent (parents: parent_info list) raw_fields thy =
   let
-    val external_names = NameSpace.external_names (Sign.naming_of thy);
+    val external_names = Name_Space.external_names (Sign.naming_of thy);
 
     val alphas = map fst args;
     val name = Sign.full_bname thy bname;
--- a/src/HOL/Tools/res_atp.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/res_atp.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -355,7 +355,7 @@
     if run_blacklist_filter andalso is_package_def name then I
     else
       let val xname = Facts.extern facts name in
-        if NameSpace.is_hidden xname then I
+        if Name_Space.is_hidden xname then I
         else cons (xname, filter_out ResAxioms.bad_for_atp ths)
       end) facts [];
 
--- a/src/HOL/Tools/res_axioms.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/Tools/res_axioms.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -85,7 +85,7 @@
             val rhs = list_abs_free (map dest_Free args, HOLogic.choice_const T $ xtp)
                     (*Forms a lambda-abstraction over the formal parameters*)
             val (c, thy') =
-              Sign.declare_const [Markup.property_internal] ((Binding.name cname, cT), NoSyn) thy
+              Sign.declare_const ((Binding.conceal (Binding.name cname), cT), NoSyn) thy
             val cdef = cname ^ "_def"
             val thy'' = Theory.add_defs_i true false [(Binding.name cdef, Logic.mk_equals (c, rhs))] thy'
             val ax = Thm.axiom thy'' (Sign.full_bname thy'' cdef)
--- a/src/HOL/ex/Predicate_Compile.thy	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/ex/Predicate_Compile.thy	Mon Oct 26 09:03:57 2009 +0100
@@ -12,6 +12,6 @@
 begin
 
 setup {* Predicate_Compile.setup *}
-setup {* Quickcheck.add_generator ("pred_compile", Pred_Compile_Quickcheck.quickcheck) *}
+setup {* Quickcheck.add_generator ("pred_compile", Predicate_Compile_Quickcheck.quickcheck) *}
 
 end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Predicate_Compile_Alternative_Defs.thy	Mon Oct 26 09:03:57 2009 +0100
@@ -0,0 +1,76 @@
+theory Predicate_Compile_Alternative_Defs
+imports Predicate_Compile
+begin
+
+section {* Set operations *}
+(*
+definition Empty where "Empty == {}"
+declare empty_def[symmetric, code_pred_inline]
+*)
+declare eq_reflection[OF empty_def, code_pred_inline] 
+(*
+definition Union where "Union A B == A Un B"
+
+lemma [code_pred_intros]: "A x ==> Union A B x"
+and  [code_pred_intros] : "B x ==> Union A B x"
+unfolding Union_def Un_def Collect_def mem_def by auto
+
+code_pred Union
+unfolding Union_def Un_def Collect_def mem_def by auto
+
+declare Union_def[symmetric, code_pred_inline]
+*)
+declare eq_reflection[OF Un_def, code_pred_inline]
+
+section {* Alternative list definitions *}
+ 
+subsection {* Alternative rules for set *}
+
+lemma set_ConsI1 [code_pred_intros]:
+  "set (x # xs) x"
+unfolding mem_def[symmetric, of _ x]
+by auto
+
+lemma set_ConsI2 [code_pred_intros]:
+  "set xs x ==> set (x' # xs) x" 
+unfolding mem_def[symmetric, of _ x]
+by auto
+
+code_pred set
+proof -
+  case set
+  from this show thesis
+    apply (case_tac a1)
+    apply auto
+    unfolding mem_def[symmetric, of _ a2]
+    apply auto
+    unfolding mem_def
+    apply auto
+    done
+qed
+
+
+subsection {* Alternative rules for list_all2 *}
+
+lemma list_all2_NilI [code_pred_intros]: "list_all2 P [] []"
+by auto
+
+lemma list_all2_ConsI [code_pred_intros]: "list_all2 P xs ys ==> P x y ==> list_all2 P (x#xs) (y#ys)"
+by auto
+
+code_pred list_all2
+proof -
+  case list_all2
+  from this show thesis
+    apply -
+    apply (case_tac a1)
+    apply (case_tac a2)
+    apply auto
+    apply (case_tac a2)
+    apply auto
+    done
+qed
+
+
+
+end
\ No newline at end of file
--- a/src/HOL/ex/Predicate_Compile_ex.thy	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/ex/Predicate_Compile_ex.thy	Mon Oct 26 09:03:57 2009 +0100
@@ -1,5 +1,5 @@
 theory Predicate_Compile_ex
-imports Main Predicate_Compile
+imports Main Predicate_Compile_Alternative_Defs
 begin
 
 inductive even :: "nat \<Rightarrow> bool" and odd :: "nat \<Rightarrow> bool" where
@@ -7,66 +7,216 @@
   | "even n \<Longrightarrow> odd (Suc n)"
   | "odd n \<Longrightarrow> even (Suc n)"
 
-code_pred even .
+code_pred (mode: [], [1]) even .
+code_pred [depth_limited] even .
+code_pred [rpred] even .
 
 thm odd.equation
 thm even.equation
+thm odd.depth_limited_equation
+thm even.depth_limited_equation
+thm even.rpred_equation
+thm odd.rpred_equation
 
 values "{x. even 2}"
 values "{x. odd 2}"
 values 10 "{n. even n}"
 values 10 "{n. odd n}"
+values [depth_limit = 2] "{x. even 6}"
+values [depth_limit = 7] "{x. even 6}"
+values [depth_limit = 2] "{x. odd 7}"
+values [depth_limit = 8] "{x. odd 7}"
+
+values [depth_limit = 7] 10 "{n. even n}"
+
+definition odd' where "odd' x == \<not> even x"
+
+code_pred [inductify] odd' .
+code_pred [inductify, depth_limited] odd' .
+code_pred [inductify, rpred] odd' .
+
+thm odd'.depth_limited_equation
+values [depth_limit = 2] "{x. odd' 7}"
+values [depth_limit = 9] "{x. odd' 7}"
 
 inductive append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
     "append [] xs xs"
   | "append xs ys zs \<Longrightarrow> append (x # xs) ys (x # zs)"
 
-code_pred append .
-code_pred (inductify_all) (rpred) append .
+code_pred (mode: [1, 2], [3], [2, 3], [1, 3], [1, 2, 3]) append .
+code_pred [depth_limited] append .
+code_pred [rpred] append .
 
 thm append.equation
+thm append.depth_limited_equation
 thm append.rpred_equation
 
 values "{(ys, xs). append xs ys [0, Suc 0, 2]}"
 values "{zs. append [0, Suc 0, 2] [17, 8] zs}"
-values "{ys. append [0, Suc 0, 2] ys [0, Suc 0, 2, 17, 0,5]}"
+values "{ys. append [0, Suc 0, 2] ys [0, Suc 0, 2, 17, 0, 5]}"
+values [depth_limit = 3] "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
+values [random] 15 "{(ys, zs). append [1::nat, 2] ys zs}"
+
+value [code] "Predicate.the (append_1_2 [0::int, 1, 2] [3, 4, 5])"
+value [code] "Predicate.the (append_3 ([]::int list))"
+
+subsection {* Tricky case with alternative rules *}
+
+inductive append2
+where
+  "append2 [] xs xs"
+| "append2 xs ys zs \<Longrightarrow> append2 (x # xs) ys (x # zs)"
+
+lemma append2_Nil: "append2 [] (xs::'b list) xs"
+  by (simp add: append2.intros(1))
+
+lemmas [code_pred_intros] = append2_Nil append2.intros(2)
+
+code_pred append2
+proof -
+  case append2
+  from append2.cases[OF append2(1)] append2(2-3) show thesis by blast
+qed
+
+subsection {* Tricky cases with tuples *}
+
+inductive zerozero :: "nat * nat => bool"
+where
+  "zerozero (0, 0)"
+
+code_pred zerozero .
+code_pred [rpred] zerozero .
+
+inductive tupled_append :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
+where
+  "tupled_append ([], xs, xs)"
+| "tupled_append (xs, ys, zs) \<Longrightarrow> tupled_append (x # xs, ys, x # zs)"
+
+code_pred tupled_append .
+code_pred [rpred] tupled_append .
+thm tupled_append.equation
+(*
+TODO: values with tupled modes
+values "{xs. tupled_append ([1,2,3], [4,5], xs)}"
+*)
+
+inductive tupled_append'
+where
+"tupled_append' ([], xs, xs)"
+| "[| ys = fst (xa, y); x # zs = snd (xa, y);
+ tupled_append' (xs, ys, zs) |] ==> tupled_append' (x # xs, xa, y)"
+
+code_pred tupled_append' .
+thm tupled_append'.equation
+
+inductive tupled_append'' :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
+where
+  "tupled_append'' ([], xs, xs)"
+| "ys = fst yszs ==> x # zs = snd yszs ==> tupled_append'' (xs, ys, zs) \<Longrightarrow> tupled_append'' (x # xs, yszs)"
+
+thm tupled_append''.cases
+
+code_pred [inductify] tupled_append'' .
+thm tupled_append''.equation
+
+inductive tupled_append''' :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
+where
+  "tupled_append''' ([], xs, xs)"
+| "yszs = (ys, zs) ==> tupled_append''' (xs, yszs) \<Longrightarrow> tupled_append''' (x # xs, ys, x # zs)"
+
+code_pred [inductify] tupled_append''' .
+thm tupled_append'''.equation
+
+inductive map_ofP :: "('a \<times> 'b) list \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
+where
+  "map_ofP ((a, b)#xs) a b"
+| "map_ofP xs a b \<Longrightarrow> map_ofP (x#xs) a b"
+
+code_pred (mode: [1], [1, 2], [1, 2, 3], [1, 3]) map_ofP .
+thm map_ofP.equation
+
+inductive filter1
+for P
+where
+  "filter1 P [] []"
+| "P x ==> filter1 P xs ys ==> filter1 P (x#xs) (x#ys)"
+| "\<not> P x ==> filter1 P xs ys ==> filter1 P (x#xs) ys"
+
+code_pred (mode: [1], [1, 2]) filter1 .
+code_pred [depth_limited] filter1 .
+code_pred [rpred] filter1 .
+
+thm filter1.equation
+
+inductive filter2
+where
+  "filter2 P [] []"
+| "P x ==> filter2 P xs ys ==> filter2 P (x#xs) (x#ys)"
+| "\<not> P x ==> filter2 P xs ys ==> filter2 P (x#xs) ys"
+
+code_pred (mode: [1, 2, 3], [1, 2]) filter2 .
+code_pred [depth_limited] filter2 .
+code_pred [rpred] filter2 .
+thm filter2.equation
+thm filter2.rpred_equation
+
+inductive filter3
+for P
+where
+  "List.filter P xs = ys ==> filter3 P xs ys"
+
+code_pred filter3 .
+code_pred [depth_limited] filter3 .
+thm filter3.depth_limited_equation
+(*code_pred [rpred] filter3 .*)
+inductive filter4
+where
+  "List.filter P xs = ys ==> filter4 P xs ys"
+
+code_pred filter4 .
+code_pred [depth_limited] filter4 .
+code_pred [rpred] filter4 .
+
+section {* reverse *}
 
 inductive rev where
     "rev [] []"
   | "rev xs xs' ==> append xs' [x] ys ==> rev (x#xs) ys"
 
-code_pred rev .
+code_pred (mode: [1], [2], [1, 2]) rev .
 
 thm rev.equation
 
 values "{xs. rev [0, 1, 2, 3::nat] xs}"
 
+inductive tupled_rev where
+  "tupled_rev ([], [])"
+| "tupled_rev (xs, xs') \<Longrightarrow> tupled_append (xs', [x], ys) \<Longrightarrow> tupled_rev (x#xs, ys)"
+
+code_pred tupled_rev .
+thm tupled_rev.equation
+
 inductive partition :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
   for f where
     "partition f [] [] []"
   | "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)"
 
-code_pred partition .
+code_pred (mode: [1], [2, 3], [1, 2], [1, 3], [1, 2, 3]) partition .
+code_pred [depth_limited] partition .
+code_pred [rpred] partition .
 
-thm partition.equation
+inductive tupled_partition :: "('a \<Rightarrow> bool) \<Rightarrow> ('a list \<times> 'a list \<times> 'a list) \<Rightarrow> bool"
+  for f where
+   "tupled_partition f ([], [], [])"
+  | "f x \<Longrightarrow> tupled_partition f (xs, ys, zs) \<Longrightarrow> tupled_partition f (x # xs, x # ys, zs)"
+  | "\<not> f x \<Longrightarrow> tupled_partition f (xs, ys, zs) \<Longrightarrow> tupled_partition f (x # xs, ys, x # zs)"
+
+code_pred tupled_partition .
+
+thm tupled_partition.equation
 
 
-inductive member
-for xs
-where "x \<in> set xs ==> member xs x"
-
-lemma [code_pred_intros]:
-  "member (x#xs') x"
-by (auto intro: member.intros)
-
-lemma [code_pred_intros]:
-"member xs x ==> member (x'#xs) x"
-by (auto intro: member.intros elim!: member.cases)
-(* strange bug must be repaired! *)
-(*
-code_pred member sorry
-*)
 inductive is_even :: "nat \<Rightarrow> bool"
 where
   "n mod 2 = 0 \<Longrightarrow> is_even n"
@@ -88,18 +238,20 @@
   case tranclp
   from this converse_tranclpE[OF this(1)] show thesis by metis
 qed
-(*
-code_pred (inductify_all) (rpred) tranclp .
+
+code_pred [depth_limited] tranclp .
+code_pred [rpred] tranclp .
 thm tranclp.equation
 thm tranclp.rpred_equation
-*)
+
 inductive succ :: "nat \<Rightarrow> nat \<Rightarrow> bool" where
     "succ 0 1"
   | "succ m n \<Longrightarrow> succ (Suc m) (Suc n)"
 
 code_pred succ .
-
+code_pred [rpred] succ .
 thm succ.equation
+thm succ.rpred_equation
 
 values 10 "{(m, n). succ n m}"
 values "{m. succ 0 m}"
@@ -141,6 +293,16 @@
  (While (%s. s!0 > 0) (Seq (Ass 0 (%s. s!0 - 1)) (Ass 1 (%s. s!1 + 1))))
  [3,5] t}"
 
+inductive tupled_exec :: "(com \<times> state \<times> state) \<Rightarrow> bool" where
+"tupled_exec (Skip, s, s)" |
+"tupled_exec (Ass x e, s, s[x := e(s)])" |
+"tupled_exec (c1, s1, s2) ==> tupled_exec (c2, s2, s3) ==> tupled_exec (Seq c1 c2, s1, s3)" |
+"b s ==> tupled_exec (c1, s, t) ==> tupled_exec (IF b c1 c2, s, t)" |
+"~b s ==> tupled_exec (c2, s, t) ==> tupled_exec (IF b c1 c2, s, t)" |
+"~b s ==> tupled_exec (While b c, s, s)" |
+"b s1 ==> tupled_exec (c, s1, s2) ==> tupled_exec (While b c, s2, s3) ==> tupled_exec (While b c, s1, s3)"
+
+code_pred tupled_exec .
 
 subsection{* CCS *}
 
@@ -171,6 +333,17 @@
 values 3 "{(a,q). step (par nil nil) a q}"
 *)
 
+inductive tupled_step :: "(proc \<times> nat \<times> proc) \<Rightarrow> bool"
+where
+"tupled_step (pre n p, n, p)" |
+"tupled_step (p1, a, q) \<Longrightarrow> tupled_step (or p1 p2, a, q)" |
+"tupled_step (p2, a, q) \<Longrightarrow> tupled_step (or p1 p2, a, q)" |
+"tupled_step (p1, a, q) \<Longrightarrow> tupled_step (par p1 p2, a, par q p2)" |
+"tupled_step (p2, a, q) \<Longrightarrow> tupled_step (par p1 p2, a, par p1 q)"
+
+code_pred tupled_step .
+thm tupled_step.equation
+
 subsection {* divmod *}
 
 inductive divmod_rel :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
@@ -179,52 +352,75 @@
 
 code_pred divmod_rel ..
 
-value [code] "Predicate.singleton (divmod_rel_1_2 1705 42)"
+value [code] "Predicate.the (divmod_rel_1_2 1705 42)"
 
 section {* Executing definitions *}
 
 definition Min
 where "Min s r x \<equiv> s x \<and> (\<forall>y. r x y \<longrightarrow> x = y)"
 
-code_pred (inductify_all) Min .
+code_pred [inductify] Min .
 
 subsection {* Examples with lists *}
 
-inductive filterP for Pa where
-"(filterP::('a => bool) => 'a list => 'a list => bool) (Pa::'a => bool) [] []"
-| "[| (res::'a list) = (y::'a) # (resa::'a list); (filterP::('a => bool) => 'a list => 'a list => bool) (Pa::'a => bool) (xt::'a list) resa; Pa y |]
-==> filterP Pa (y # xt) res"
-| "[| (filterP::('a => bool) => 'a list => 'a list => bool) (Pa::'a => bool) (xt::'a list) (res::'a list); ~ Pa (y::'a) |] ==> filterP Pa (y # xt) res"
+subsubsection {* Lexicographic order *}
+
+thm lexord_def
+code_pred [inductify] lexord .
+code_pred [inductify, rpred] lexord .
+thm lexord.equation
+thm lexord.rpred_equation
+
+inductive less_than_nat :: "nat * nat => bool"
+where
+  "less_than_nat (0, x)"
+| "less_than_nat (x, y) ==> less_than_nat (Suc x, Suc y)"
+ 
+code_pred less_than_nat .
+
+code_pred [depth_limited] less_than_nat .
+code_pred [rpred] less_than_nat .
 
+inductive test_lexord :: "nat list * nat list => bool"
+where
+  "lexord less_than_nat (xs, ys) ==> test_lexord (xs, ys)"
+
+code_pred [rpred] test_lexord .
+code_pred [depth_limited] test_lexord .
+thm test_lexord.depth_limited_equation
+thm test_lexord.rpred_equation
+
+values "{x. test_lexord ([1, 2, 3], [1, 2, 5])}"
+values [depth_limit = 5] "{x. test_lexord ([1, 2, 3], [1, 2, 5])}"
+(*values [random] "{xys. test_lexord xys}"*)
+(*values [depth_limit = 5 random] "{xy. lexord less_than_nat xy}"*)
 (*
-code_pred (inductify_all) (rpred) filterP .
-thm filterP.rpred_equation
+lemma "(u, v) : lexord less_than_nat ==> (x @ u, y @ v) : lexord less_than_nat"
+quickcheck[generator=pred_compile]
+oops
 *)
-
-code_pred (inductify_all) lexord .
-
-thm lexord.equation
-
-lemma "(u, v) : lexord r ==> (x @ u, y @ v) : lexord r"
-(*quickcheck[generator=pred_compile]*)
-oops
-
 lemmas [code_pred_def] = lexn_conv lex_conv lenlex_conv
 
-code_pred (inductify_all) lexn .
+code_pred [inductify] lexn .
 thm lexn.equation
 
-code_pred (inductify_all) lenlex .
+code_pred [rpred] lexn .
+
+thm lexn.rpred_equation
+
+code_pred [inductify, show_steps] lenlex .
 thm lenlex.equation
-(*
-code_pred (inductify_all) (rpred) lenlex .
+
+code_pred [inductify, rpred] lenlex .
 thm lenlex.rpred_equation
-*)
+
 thm lists.intros
-code_pred (inductify_all) lists .
+code_pred [inductify] lists .
 
 thm lists.equation
 
+section {* AVL Tree Example *}
+
 datatype 'a tree = ET | MKT 'a "'a tree" "'a tree" nat
 fun height :: "'a tree => nat" where
 "height ET = 0"
@@ -236,40 +432,107 @@
   "avl (MKT x l r h) = ((height l = height r \<or> height l = 1 + height r \<or> height r = 1+height l) \<and> 
   h = max (height l) (height r) + 1 \<and> avl l \<and> avl r)"
 
-code_pred (inductify_all) avl .
+code_pred [inductify] avl .
 thm avl.equation
 
-lemma [code_pred_inline]: "bot_fun_inst.bot_fun == (\<lambda>(y::'a::type). False)"
-unfolding bot_fun_inst.bot_fun[symmetric] bot_bool_eq[symmetric] bot_fun_eq by simp
+code_pred [rpred] avl .
+thm avl.rpred_equation
+(*values [random] 10 "{t. avl (t::int tree)}"*)
 
 fun set_of
 where
 "set_of ET = {}"
 | "set_of (MKT n l r h) = insert n (set_of l \<union> set_of r)"
 
-fun is_ord
+fun is_ord :: "nat tree => bool"
 where
 "is_ord ET = True"
 | "is_ord (MKT n l r h) =
  ((\<forall>n' \<in> set_of l. n' < n) \<and> (\<forall>n' \<in> set_of r. n < n') \<and> is_ord l \<and> is_ord r)"
 
-declare Un_def[code_pred_def]
-
-code_pred (inductify_all) set_of .
+code_pred (mode: [1], [1, 2]) [inductify] set_of .
 thm set_of.equation
-(* FIXME *)
-(*
-code_pred (inductify_all) is_ord .
+
+code_pred [inductify] is_ord .
 thm is_ord.equation
-*)
+code_pred [rpred] is_ord .
+thm is_ord.rpred_equation
+
 section {* Definitions about Relations *}
 
-code_pred (inductify_all) converse .
+code_pred [inductify] converse .
 thm converse.equation
+code_pred [inductify] rel_comp .
+thm rel_comp.equation
+code_pred [inductify] Image .
+thm Image.equation
+(*TODO: *)
+ML {* Toplevel.debug := true *}
+declare Id_on_def[unfolded UNION_def, code_pred_def]
+
+code_pred [inductify] Id_on .
+thm Id_on.equation
+code_pred [inductify] Domain .
+thm Domain.equation
+code_pred [inductify] Range .
+thm sym_def
+code_pred [inductify] Field .
+declare Sigma_def[unfolded UNION_def, code_pred_def]
+declare refl_on_def[unfolded UNION_def, code_pred_def]
+code_pred [inductify] refl_on .
+thm refl_on.equation
+code_pred [inductify] total_on .
+thm total_on.equation
+(*
+code_pred [inductify] sym .
+thm sym.equation
+*)
+code_pred [inductify] antisym .
+thm antisym.equation
+code_pred [inductify] trans .
+thm trans.equation
+code_pred [inductify] single_valued .
+thm single_valued.equation
+code_pred [inductify] inv_image .
+thm inv_image.equation
 
-code_pred (inductify_all) Domain .
-thm Domain.equation
+section {* List functions *}
+
+code_pred [inductify] length .
+thm size_listP.equation
+code_pred [inductify, rpred] length .
+thm size_listP.rpred_equation
+values [random] 20 "{xs. size_listP (xs::nat list) (5::nat)}"
 
+code_pred [inductify] concat .
+code_pred [inductify] hd .
+code_pred [inductify] tl .
+code_pred [inductify] last .
+code_pred [inductify] butlast .
+(*code_pred [inductify] listsum .*)
+code_pred [inductify] take .
+code_pred [inductify] drop .
+code_pred [inductify] zip .
+code_pred [inductify] upt .
+code_pred [inductify] remdups .
+code_pred [inductify] remove1 .
+code_pred [inductify] removeAll .
+code_pred [inductify] distinct .
+code_pred [inductify] replicate .
+code_pred [inductify] splice .
+code_pred [inductify] List.rev .
+code_pred [inductify] map .
+code_pred [inductify] foldr .
+code_pred [inductify] foldl .
+code_pred [inductify] filter .
+code_pred [inductify, rpred] filter .
+thm filterP.rpred_equation
+
+definition test where "test xs = filter (\<lambda>x. x = (1::nat)) xs"
+code_pred [inductify] test .
+thm testP.equation
+code_pred [inductify, rpred] test .
+thm testP.rpred_equation
 
 section {* Context Free Grammar *}
 
@@ -283,15 +546,86 @@
 | "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
 | "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1"
 
-code_pred (inductify_all) S\<^isub>1p .
+code_pred [inductify] S\<^isub>1p .
+code_pred [inductify, rpred] S\<^isub>1p .
+thm S\<^isub>1p.equation
+thm S\<^isub>1p.rpred_equation
 
-thm S\<^isub>1p.equation
+values [random] 5 "{x. S\<^isub>1p x}"
+
+inductive is_a where
+  "is_a a"
+
+inductive is_b where
+  "is_b b"
 
-theorem S\<^isub>1_sound:
-"w \<in> S\<^isub>1 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-quickcheck[generator=pred_compile]
+code_pred is_a .
+code_pred [depth_limited] is_a .
+code_pred [rpred] is_a .
+
+values [random] "{x. is_a x}"
+code_pred [depth_limited] is_b .
+code_pred [rpred] is_b .
+
+code_pred [inductify, depth_limited] filter .
+
+values [depth_limit=5] "{x. filterP is_a [a, b] x}"
+values [depth_limit=3] "{x. filterP is_b [a, b] x}"
+
+lemma "w \<in> S\<^isub>1 \<Longrightarrow> length (filter (\<lambda>x. x = a) w) = 1"
+(*quickcheck[generator=pred_compile, size=10]*)
 oops
 
+inductive test_lemma where
+  "S\<^isub>1p w ==> filterP is_a w r1 ==> size_listP r1 r2 ==> filterP is_b w r3 ==> size_listP r3 r4 ==> r2 \<noteq> r4 ==> test_lemma w"
+(*
+code_pred [rpred] test_lemma .
+*)
+(*
+definition test_lemma' where
+  "test_lemma' w == (w \<in> S\<^isub>1 \<and> (\<not> length [x <- w. x = a] = length [x <- w. x = b]))"
+
+code_pred [inductify, rpred] test_lemma' .
+thm test_lemma'.rpred_equation
+*)
+(*thm test_lemma'.rpred_equation*)
+(*
+values [depth_limit=3 random] "{x. S\<^isub>1 x}"
+*)
+code_pred [depth_limited] is_b .
+(*
+code_pred [inductify, depth_limited] filter .
+*)
+thm filterP.intros
+thm filterP.equation
+(*
+values [depth_limit=3] "{x. filterP is_b [a, b] x}"
+find_theorems "test_lemma'_hoaux"
+code_pred [depth_limited] test_lemma'_hoaux .
+thm test_lemma'_hoaux.depth_limited_equation
+values [depth_limit=2] "{x. test_lemma'_hoaux b}"
+inductive test1 where
+  "\<not> test_lemma'_hoaux x ==> test1 x"
+code_pred test1 .
+code_pred [depth_limited] test1 .
+thm test1.depth_limited_equation
+thm test_lemma'_hoaux.depth_limited_equation
+thm test1.intros
+
+values [depth_limit=2] "{x. test1 b}"
+
+thm filterP.intros
+thm filterP.depth_limited_equation
+values [depth_limit=3] "{x. filterP test_lemma'_hoaux [a, b] x}"
+values [depth_limit=4 random] "{w. test_lemma w}"
+values [depth_limit=4 random] "{w. test_lemma' w}"
+*)
+(*
+theorem S\<^isub>1_sound:
+"w \<in> S\<^isub>1p \<Longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+quickcheck[generator=pred_compile, size=15]
+oops
+*)
 inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where
   "[] \<in> S\<^isub>2"
 | "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2"
@@ -299,14 +633,18 @@
 | "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2"
 | "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2"
 | "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2"
-(*
-code_pred (inductify_all) (rpred) S\<^isub>2 .
-ML {* Predicate_Compile_Core.intros_of @{theory} @{const_name "B\<^isub>2"} *}
-*)
+
+code_pred [inductify, rpred] S\<^isub>2 .
+thm S\<^isub>2.rpred_equation
+thm A\<^isub>2.rpred_equation
+thm B\<^isub>2.rpred_equation
+
+values [random] 10 "{x. S\<^isub>2 x}"
+
 theorem S\<^isub>2_sound:
 "w \<in> S\<^isub>2 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
 (*quickcheck[generator=SML]*)
-quickcheck[generator=pred_compile, size=15, iterations=100]
+(*quickcheck[generator=pred_compile, size=15, iterations=1]*)
 oops
 
 inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
@@ -317,23 +655,35 @@
 | "w \<in> S\<^isub>3 \<Longrightarrow> b # w \<in> B\<^isub>3"
 | "\<lbrakk>v \<in> B\<^isub>3; w \<in> B\<^isub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>3"
 
+code_pred [inductify] S\<^isub>3 .
+thm S\<^isub>3.equation
+
+values 10 "{x. S\<^isub>3 x}"
 (*
-code_pred (inductify_all) S\<^isub>3 .
-*)
 theorem S\<^isub>3_sound:
 "w \<in> S\<^isub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
 quickcheck[generator=pred_compile, size=10, iterations=1]
 oops
-
+*)
 lemma "\<not> (length w > 2) \<or> \<not> (length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b])"
-quickcheck[size=10, generator = pred_compile]
+(*quickcheck[size=10, generator = pred_compile]*)
 oops
+(*
+inductive test
+where
+  "length [x \<leftarrow> w. a = x] = length [x \<leftarrow> w. x = b] ==> test w"
+ML {* @{term "[x \<leftarrow> w. x = a]"} *}
+code_pred (inductify_all) test .
 
+thm test.equation
+*)
+(*
 theorem S\<^isub>3_complete:
-"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>3"
+"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. b = x] \<longrightarrow> w \<in> S\<^isub>3"
 (*quickcheck[generator=SML]*)
 quickcheck[generator=pred_compile, size=10, iterations=100]
 oops
+*)
 
 inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where
   "[] \<in> S\<^isub>4"
@@ -343,15 +693,15 @@
 | "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4"
 | "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4"
 | "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4"
-
+(*
 theorem S\<^isub>4_sound:
 "w \<in> S\<^isub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
 quickcheck[generator = pred_compile, size=2, iterations=1]
 oops
-
+*)
 theorem S\<^isub>4_complete:
 "length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>4"
-quickcheck[generator = pred_compile, size=5, iterations=1]
+(*quickcheck[generator = pred_compile, size=5, iterations=1]*)
 oops
 
 theorem S\<^isub>4_A\<^isub>4_B\<^isub>4_sound_and_complete:
@@ -361,8 +711,8 @@
 (*quickcheck[generator = pred_compile, size=5, iterations=1]*)
 oops
 
+section {* Lambda *}
 
-section {* Lambda *}
 datatype type =
     Atom nat
   | Fun type type    (infixr "\<Rightarrow>" 200)
@@ -378,15 +728,15 @@
   "[]\<langle>i\<rangle> = None"
 | "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"
 
-(*
+
 inductive nth_el' :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> bool"
 where
   "nth_el' (x # xs) 0 x"
 | "nth_el' xs i y \<Longrightarrow> nth_el' (x # xs) (Suc i) y"
-*)
+
 inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool"  ("_ \<turnstile> _ : _" [50, 50, 50] 50)
   where
-    Var [intro!]: "nth_el env x = Some T \<Longrightarrow> env \<turnstile> Var x : T"
+    Var [intro!]: "nth_el' env x T \<Longrightarrow> env \<turnstile> Var x : T"
   | Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> U)"
 (*  | App [intro!]: "env \<turnstile> s : T \<Rightarrow> U \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U" *)
   | App [intro!]: "env \<turnstile> s : U \<Rightarrow> T \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U"
@@ -414,22 +764,8 @@
   | abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t"
 
 lemma "Gamma \<turnstile> t : T \<Longrightarrow> t \<rightarrow>\<^sub>\<beta> t' \<Longrightarrow> Gamma \<turnstile> t' : T"
-quickcheck[generator = pred_compile, size = 10, iterations = 1000]
+(*quickcheck[generator = pred_compile, size = 10, iterations = 1]*)
 oops
-(* FIXME *)
-(*
-inductive test for P where
-"[| filter P vs = res |]
-==> test P vs res"
-
-code_pred test .
-*)
-(*
-export_code test_for_1_yields_1_2 in SML file -
-code_pred (inductify_all) (rpred) test .
-
-thm test.equation
-*)
 
 lemma filter_eq_ConsD:
  "filter P ys = x#xs \<Longrightarrow>
--- a/src/HOL/ex/RPred.thy	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/HOL/ex/RPred.thy	Mon Oct 26 09:03:57 2009 +0100
@@ -2,7 +2,7 @@
 imports Quickcheck Random Predicate
 begin
 
-types 'a rpred = "Random.seed \<Rightarrow> ('a Predicate.pred \<times> Random.seed)"
+types 'a "rpred" = "Random.seed \<Rightarrow> ('a Predicate.pred \<times> Random.seed)"
 
 section {* The RandomPred Monad *}
 
@@ -33,9 +33,9 @@
 
 (* Missing a good definition for negation: not_rpred *)
 
-definition not_rpred :: "unit Predicate.pred \<Rightarrow> unit rpred"
+definition not_rpred :: "unit rpred \<Rightarrow> unit rpred"
 where
-  "not_rpred = Pair o Predicate.not_pred"
+  "not_rpred P = (\<lambda>s. let (P', s') = P s in if Predicate.eval P' () then (Orderings.bot, s') else (Predicate.single (), s'))"
 
 definition lift_pred :: "'a Predicate.pred \<Rightarrow> 'a rpred"
   where
@@ -44,9 +44,9 @@
 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 = bind P (return o f)"
+definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a rpred \<Rightarrow> 'b rpred)"
+  where "map f P = bind P (return o f)"
 
-hide (open) const return bind supp 
+hide (open) const return bind supp map
   
 end
\ No newline at end of file
--- a/src/Pure/General/binding.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/General/binding.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -10,7 +10,7 @@
 signature BINDING =
 sig
   type binding
-  val dest: binding -> (string * bool) list * bstring
+  val dest: binding -> bool * (string * bool) list * bstring
   val make: bstring * Position.T -> binding
   val pos_of: binding -> Position.T
   val name: bstring -> binding
@@ -27,6 +27,7 @@
   val prefix_of: binding -> (string * bool) list
   val map_prefix: ((string * bool) list -> (string * bool) list) -> binding -> binding
   val prefix: bool -> string -> binding -> binding
+  val conceal: binding -> binding
   val str_of: binding -> string
 end;
 
@@ -38,19 +39,21 @@
 (* datatype *)
 
 abstype binding = Binding of
- {prefix: (string * bool) list,     (*system prefix*)
+ {conceal: bool,                    (*internal -- for foundational purposes only*)
+  prefix: (string * bool) list,     (*system prefix*)
   qualifier: (string * bool) list,  (*user qualifier*)
   name: bstring,                    (*base name*)
   pos: Position.T}                  (*source position*)
 with
 
-fun make_binding (prefix, qualifier, name, pos) =
-  Binding {prefix = prefix, qualifier = qualifier, name = name, pos = pos};
+fun make_binding (conceal, prefix, qualifier, name, pos) =
+  Binding {conceal = conceal, prefix = prefix, qualifier = qualifier, name = name, pos = pos};
 
-fun map_binding f (Binding {prefix, qualifier, name, pos}) =
-  make_binding (f (prefix, qualifier, name, pos));
+fun map_binding f (Binding {conceal, prefix, qualifier, name, pos}) =
+  make_binding (f (conceal, prefix, qualifier, name, pos));
 
-fun dest (Binding {prefix, qualifier, name, ...}) = (prefix @ qualifier, name);
+fun dest (Binding {conceal, prefix, qualifier, name, ...}) =
+  (conceal, prefix @ qualifier, name);
 
 
 
@@ -58,7 +61,7 @@
 
 (* name and position *)
 
-fun make (name, pos) = make_binding ([], [], name, pos);
+fun make (name, pos) = make_binding (false, [], [], name, pos);
 fun name name = make (name, Position.none);
 
 fun pos_of (Binding {pos, ...}) = pos;
@@ -66,7 +69,10 @@
 
 fun eq_name (b, b') = name_of b = name_of b';
 
-fun map_name f = map_binding (fn (prefix, qualifier, name, pos) => (prefix, qualifier, f name, pos));
+fun map_name f =
+  map_binding (fn (conceal, prefix, qualifier, name, pos) =>
+    (conceal, prefix, qualifier, f name, pos));
+
 val prefix_name = map_name o prefix;
 val suffix_name = map_name o suffix;
 
@@ -77,13 +83,14 @@
 (* user qualifier *)
 
 fun qualify _ "" = I
-  | qualify mandatory qual = map_binding (fn (prefix, qualifier, name, pos) =>
-      (prefix, (qual, mandatory) :: qualifier, name, pos));
+  | qualify mandatory qual =
+      map_binding (fn (conceal, prefix, qualifier, name, pos) =>
+        (conceal, prefix, (qual, mandatory) :: qualifier, name, pos));
 
 fun qualified_name "" = empty
   | qualified_name s =
       let val (qualifier, name) = split_last (Long_Name.explode s)
-      in make_binding ([], map (rpair false) qualifier, name, Position.none) end;
+      in make_binding (false, [], map (rpair false) qualifier, name, Position.none) end;
 
 fun qualified_name_of (b as Binding {qualifier, name, ...}) =
   if is_empty b then ""
@@ -94,13 +101,21 @@
 
 fun prefix_of (Binding {prefix, ...}) = prefix;
 
-fun map_prefix f = map_binding (fn (prefix, qualifier, name, pos) =>
-  (f prefix, qualifier, name, pos));
+fun map_prefix f =
+  map_binding (fn (conceal, prefix, qualifier, name, pos) =>
+    (conceal, f prefix, qualifier, name, pos));
 
 fun prefix _ "" = I
   | prefix mandatory prfx = map_prefix (cons (prfx, mandatory));
 
 
+(* conceal *)
+
+val conceal =
+  map_binding (fn (_, prefix, qualifier, name, pos) =>
+    (true, prefix, qualifier, name, pos));
+
+
 (* str_of *)
 
 fun str_of binding =
--- a/src/Pure/General/markup.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/General/markup.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -13,11 +13,11 @@
   val nameN: string
   val name: string -> T -> T
   val bindingN: string val binding: string -> T
-  val groupN: string
-  val theory_nameN: string
   val kindN: string
   val internalK: string
-  val property_internal: Properties.property
+  val entityN: string val entity: string -> T
+  val defN: string
+  val refN: string
   val lineN: string
   val columnN: string
   val offsetN: string
@@ -149,16 +149,20 @@
 
 val (bindingN, binding) = markup_string "binding" nameN;
 
-val groupN = "group";
-val theory_nameN = "theory_name";
-
 
 (* kind *)
 
 val kindN = "kind";
 
 val internalK = "internal";
-val property_internal = (kindN, internalK);
+
+
+(* formal entities *)
+
+val (entityN, entity) = markup_string "entity" nameN;
+
+val defN = "def";
+val refN = "ref";
 
 
 (* position *)
--- a/src/Pure/General/markup.scala	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/General/markup.scala	Mon Oct 26 09:03:57 2009 +0100
@@ -15,6 +15,13 @@
   val KIND = "kind"
 
 
+  /* formal entities */
+
+  val ENTITY = "entity"
+  val DEF = "def"
+  val REF = "ref"
+
+
   /* position */
 
   val LINE = "line"
--- a/src/Pure/General/name_space.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/General/name_space.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -20,33 +20,40 @@
   val hidden: string -> string
   val is_hidden: string -> bool
   type T
-  val empty: T
+  val empty: string -> T
+  val kind_of: T -> string
+  val the_entry: T -> string ->
+    {concealed: bool, group: serial option, theory_name: string, pos: Position.T, id: serial}
+  val is_concealed: T -> string -> bool
   val intern: T -> xstring -> string
-  val extern: T -> string -> xstring
   val extern_flags: {long_names: bool, short_names: bool, unique_names: bool} ->
     T -> string -> xstring
+  val extern: T -> string -> xstring
   val hide: bool -> string -> T -> T
   val merge: T * T -> T
   type naming
   val default_naming: naming
-  val declare: naming -> binding -> T -> string * T
-  val full_name: naming -> binding -> string
-  val external_names: naming -> string -> string list
+  val conceal: naming -> naming
+  val set_group: serial -> naming -> naming
+  val set_theory_name: string -> naming -> naming
   val add_path: string -> naming -> naming
   val root_path: naming -> naming
   val parent_path: naming -> naming
   val mandatory_path: string -> naming -> naming
+  val full_name: naming -> binding -> string
+  val external_names: naming -> string -> string list
+  val declare: bool -> naming -> binding -> T -> string * T
   type 'a table = T * 'a Symtab.table
-  val define: naming -> binding * 'a -> 'a table -> string * 'a table (*exception Symtab.DUP*)
-  val empty_table: 'a table
-  val merge_tables: ('a * 'a -> bool) -> 'a table * 'a table -> 'a table (*exception Symtab.DUP*)
-  val join_tables: (string -> 'a * 'a -> 'a) (*exception Symtab.DUP/Symtab.SAME*) ->
-    'a table * 'a table -> 'a table (*exception Symtab.DUP*)
+  val define: bool -> naming -> binding * 'a -> 'a table -> string * 'a table
+  val empty_table: string -> 'a table
+  val merge_tables: 'a table * 'a table -> 'a table
+  val join_tables: (string -> 'a * 'a -> 'a) (*Symtab.SAME*) ->
+    'a table * 'a table -> 'a table
   val dest_table: 'a table -> (string * 'a) list
   val extern_table: 'a table -> (xstring * 'a) list
 end;
 
-structure NameSpace: NAME_SPACE =
+structure Name_Space: NAME_SPACE =
 struct
 
 
@@ -58,33 +65,76 @@
 val is_hidden = String.isPrefix "??.";
 
 
+(* datatype entry *)
+
+type entry =
+ {externals: xstring list,
+  concealed: bool,
+  group: serial option,
+  theory_name: string,
+  pos: Position.T,
+  id: serial};
+
+fun str_of_entry def (name, {pos, id, ...}: entry) =
+  let
+    val occurrence = (if def then Markup.defN else Markup.refN, string_of_int id);
+    val props = occurrence :: Position.properties_of pos;
+  in Markup.markup (Markup.properties props (Markup.entity name)) name end;
+
+fun err_dup kind entry1 entry2 =
+  error ("Duplicate " ^ kind ^ " declaration " ^
+    quote (str_of_entry true entry1) ^ " vs. " ^ quote (str_of_entry true entry2));
+
+
 (* datatype T *)
 
 datatype T =
-  NameSpace of
-    (string list * string list) Symtab.table *   (*internals, hidden internals*)
-    xstring list Symtab.table;                   (*externals*)
+  Name_Space of
+   {kind: string,
+    internals: (string list * string list) Symtab.table,  (*visible, hidden*)
+    entries: entry Symtab.table};
+
+fun make_name_space (kind, internals, entries) =
+  Name_Space {kind = kind, internals = internals, entries = entries};
+
+fun map_name_space f (Name_Space {kind = kind, internals = internals, entries = entries}) =
+  make_name_space (f (kind, internals, entries));
+
+fun map_internals f xname = map_name_space (fn (kind, internals, entries) =>
+  (kind, Symtab.map_default (xname, ([], [])) f internals, entries));
+
 
-val empty = NameSpace (Symtab.empty, Symtab.empty);
+fun empty kind = make_name_space (kind, Symtab.empty, Symtab.empty);
+
+fun kind_of (Name_Space {kind, ...}) = kind;
 
-fun lookup (NameSpace (tab, _)) xname =
-  (case Symtab.lookup tab xname of
+fun the_entry (Name_Space {kind, entries, ...}) name =
+  (case Symtab.lookup entries name of
+    NONE => error ("Unknown " ^ kind ^ " " ^ quote name)
+  | SOME {concealed, group, theory_name, pos, id, ...} =>
+      {concealed = concealed, group = group, theory_name = theory_name, pos = pos, id = id});
+
+fun is_concealed space name = #concealed (the_entry space name);
+
+
+(* name accesses *)
+
+fun lookup (Name_Space {internals, ...}) xname =
+  (case Symtab.lookup internals xname of
     NONE => (xname, true)
   | SOME ([], []) => (xname, true)
   | SOME ([name], _) => (name, true)
   | SOME (name :: _, _) => (name, false)
   | SOME ([], name' :: _) => (hidden name', true));
 
-fun get_accesses (NameSpace (_, xtab)) name =
-  (case Symtab.lookup xtab name of
+fun get_accesses (Name_Space {entries, ...}) name =
+  (case Symtab.lookup entries name of
     NONE => [name]
-  | SOME xnames => xnames);
+  | SOME {externals, ...} => externals);
 
-fun put_accesses name xnames (NameSpace (tab, xtab)) =
-  NameSpace (tab, Symtab.update (name, xnames) xtab);
-
-fun valid_accesses (NameSpace (tab, _)) name = Symtab.fold (fn (xname, (names, _)) =>
-  if not (null names) andalso hd names = name then cons xname else I) tab [];
+fun valid_accesses (Name_Space {internals, ...}) name =
+  Symtab.fold (fn (xname, (names, _)) =>
+    if not (null names) andalso hd names = name then cons xname else I) internals [];
 
 
 (* intern and extern *)
@@ -116,21 +166,13 @@
     unique_names = ! unique_names} space name;
 
 
-(* basic operations *)
-
-local
-
-fun map_space f xname (NameSpace (tab, xtab)) =
-  NameSpace (Symtab.map_default (xname, ([], [])) f tab, xtab);
+(* modify internals *)
 
-in
-
-val del_name = map_space o apfst o remove (op =);
-fun del_name_extra name = map_space (apfst (fn [] => [] | x :: xs => x :: remove (op =) name xs));
-val add_name = map_space o apfst o update (op =);
-val add_name' = map_space o apsnd o update (op =);
-
-end;
+val del_name = map_internals o apfst o remove (op =);
+fun del_name_extra name =
+  map_internals (apfst (fn [] => [] | x :: xs => x :: remove (op =) name xs));
+val add_name = map_internals o apfst o update (op =);
+val add_name' = map_internals o apsnd o update (op =);
 
 
 (* hide *)
@@ -152,17 +194,24 @@
 
 (* merge *)
 
-fun merge (NameSpace (tab1, xtab1), NameSpace (tab2, xtab2)) =
+fun merge
+  (Name_Space {kind = kind1, internals = internals1, entries = entries1},
+    Name_Space {kind = kind2, internals = internals2, entries = entries2}) =
   let
-    val tab' = (tab1, tab2) |> Symtab.join
+    val kind' =
+      if kind1 = kind2 then kind1
+      else error ("Attempt to merge different kinds of name spaces " ^
+        quote kind1 ^ " vs. " ^ quote kind2);
+    val internals' = (internals1, internals2) |> Symtab.join
       (K (fn ((names1, names1'), (names2, names2')) =>
-        if pointer_eq (names1, names2) andalso pointer_eq (names1', names2') then raise Symtab.SAME
+        if pointer_eq (names1, names2) andalso pointer_eq (names1', names2')
+        then raise Symtab.SAME
         else (Library.merge (op =) (names1, names2), Library.merge (op =) (names1', names2'))));
-    val xtab' = (xtab1, xtab2) |> Symtab.join
-      (K (fn xnames =>
-        if pointer_eq xnames then raise Symtab.SAME
-        else (Library.merge (op =) xnames)));
-  in NameSpace (tab', xtab') end;
+    val entries' = (entries1, entries2) |> Symtab.join
+      (fn name => fn (entry1, entry2) =>
+        if #id entry1 = #id entry2 then raise Symtab.SAME
+        else err_dup kind' (name, entry1) (name, entry2));
+  in make_name_space (kind', internals', entries') end;
 
 
 
@@ -170,36 +219,59 @@
 
 (* datatype naming *)
 
-datatype naming = Naming of (string * bool) list;
-fun map_naming f (Naming path) = Naming (f path);
+datatype naming = Naming of
+ {conceal: bool,
+  group: serial option,
+  theory_name: string,
+  path: (string * bool) list};
 
-val default_naming = Naming [];
+fun make_naming (conceal, group, theory_name, path) =
+  Naming {conceal = conceal, group = group, theory_name = theory_name, path = path};
+
+fun map_naming f (Naming {conceal, group, theory_name, path}) =
+  make_naming (f (conceal, group, theory_name, path));
+
+fun map_path f = map_naming (fn (conceal, group, theory_name, path) =>
+  (conceal, group, theory_name, f path));
+
 
-fun add_path elems = map_naming (fn path => path @ [(elems, false)]);
-val root_path = map_naming (fn _ => []);
-val parent_path = map_naming (perhaps (try (#1 o split_last)));
-fun mandatory_path elems = map_naming (fn path => path @ [(elems, true)]);
+val default_naming = make_naming (false, NONE, "", []);
+
+val conceal = map_naming (fn (_, group, theory_name, path) =>
+  (true, group, theory_name, path));
+
+fun set_group group = map_naming (fn (conceal, _, theory_name, path) =>
+  (conceal, SOME group, theory_name, path));
+
+fun set_theory_name theory_name = map_naming (fn (conceal, group, _, path) =>
+  (conceal, group, theory_name, path));
+
+fun add_path elems = map_path (fn path => path @ [(elems, false)]);
+val root_path = map_path (fn _ => []);
+val parent_path = map_path (perhaps (try (#1 o split_last)));
+fun mandatory_path elems = map_path (fn path => path @ [(elems, true)]);
 
 
 (* full name *)
 
 fun err_bad binding = error ("Bad name binding " ^ quote (Binding.str_of binding));
 
-fun name_spec (Naming path) binding =
+fun name_spec (Naming {conceal = conceal1, path, ...}) binding =
   let
-    val (prefix, name) = Binding.dest binding;
+    val (conceal2, prefix, name) = Binding.dest binding;
     val _ = Long_Name.is_qualified name andalso err_bad binding;
 
+    val concealed = conceal1 orelse conceal2;
     val spec1 = maps (fn (a, b) => map (rpair b) (Long_Name.explode a)) (path @ prefix);
     val spec2 = if name = "" then [] else [(name, true)];
     val spec = spec1 @ spec2;
     val _ =
       exists (fn (a, _) => a = "" orelse a = "??" orelse exists_string (fn s => s = "\"") a) spec
       andalso err_bad binding;
-  in if null spec2 then [] else spec end;
+  in (concealed, if null spec2 then [] else spec) end;
 
-fun full naming = name_spec naming #> map fst;
-fun full_name naming = full naming #> Long_Name.implode;
+fun full_name naming =
+  name_spec naming #> #2 #> map #1 #> Long_Name.implode;
 
 
 (* accesses *)
@@ -215,7 +287,7 @@
 
 fun accesses naming binding =
   let
-    val spec = name_spec naming binding;
+    val spec = #2 (name_spec naming binding);
     val sfxs = mandatory_suffixes spec;
     val pfxs = mandatory_prefixes spec;
   in pairself (map Long_Name.implode) (sfxs @ pfxs, sfxs) end;
@@ -225,13 +297,32 @@
 
 (* declaration *)
 
-fun declare naming binding space =
+fun new_entry strict entry =
+  map_name_space (fn (kind, internals, entries) =>
+    let
+      val entries' =
+        (if strict then Symtab.update_new else Symtab.update) entry entries
+          handle Symtab.DUP dup =>
+            err_dup kind (dup, the (Symtab.lookup entries dup)) entry;
+    in (kind, internals, entries') end);
+
+fun declare strict naming binding space =
   let
-    val names = full naming binding;
-    val name = Long_Name.implode names;
+    val Naming {group, theory_name, ...} = naming;
+    val (concealed, spec) = name_spec naming binding;
+    val (accs, accs') = accesses naming binding;
+
+    val name = Long_Name.implode (map fst spec);
     val _ = name = "" andalso err_bad binding;
-    val (accs, accs') = accesses naming binding;
-    val space' = space |> fold (add_name name) accs |> put_accesses name accs';
+
+    val entry =
+     {externals = accs',
+      concealed = concealed,
+      group = group,
+      theory_name = theory_name,
+      pos = Position.default (Binding.pos_of binding),
+      id = serial ()};
+    val space' = space |> fold (add_name name) accs |> new_entry strict (name, entry);
   in (name, space') end;
 
 
@@ -240,14 +331,14 @@
 
 type 'a table = T * 'a Symtab.table;
 
-fun define naming (binding, x) (space, tab) =
-  let val (name, space') = declare naming binding space
-  in (name, (space', Symtab.update_new (name, x) tab)) end;
+fun define strict naming (binding, x) (space, tab) =
+  let val (name, space') = declare strict naming binding space
+  in (name, (space', Symtab.update (name, x) tab)) end;
 
-val empty_table = (empty, Symtab.empty);
+fun empty_table kind = (empty kind, Symtab.empty);
 
-fun merge_tables eq ((space1, tab1), (space2, tab2)) =
-  (merge (space1, space2), Symtab.merge eq (tab1, tab2));
+fun merge_tables ((space1, tab1), (space2, tab2)) =
+  (merge (space1, space2), Symtab.merge (K true) (tab1, tab2));
 
 fun join_tables f ((space1, tab1), (space2, tab2)) =
   (merge (space1, space2), Symtab.join f (tab1, tab2));
@@ -261,6 +352,6 @@
 
 end;
 
-structure Basic_Name_Space: BASIC_NAME_SPACE = NameSpace;
+structure Basic_Name_Space: BASIC_NAME_SPACE = Name_Space;
 open Basic_Name_Space;
 
--- a/src/Pure/General/position.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/General/position.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -37,6 +37,7 @@
   val range: T -> T -> range
   val thread_data: unit -> T
   val setmp_thread_data: T -> ('a -> 'b) -> 'a -> 'b
+  val default: T -> T
 end;
 
 structure Position: POSITION =
@@ -178,4 +179,8 @@
 
 end;
 
+fun default pos =
+  if pos = none then thread_data ()
+  else pos;
+
 end;
--- a/src/Pure/General/table.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/General/table.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -54,8 +54,7 @@
   val update_list: ('a * 'a -> bool) -> key * 'a -> 'a list table -> 'a list table
   val make_list: (key * 'a) list -> 'a list table
   val dest_list: 'a list table -> (key * 'a) list
-  val merge_list: ('a * 'a -> bool) ->
-    'a list table * 'a list table -> 'a list table                     (*exception DUP*)
+  val merge_list: ('a * 'a -> bool) -> 'a list table * 'a list table -> 'a list table
 end;
 
 functor Table(Key: KEY): TABLE =
--- a/src/Pure/Isar/attrib.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/Isar/attrib.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -67,35 +67,33 @@
 
 structure Attributes = TheoryDataFun
 (
-  type T = (((src -> attribute) * string) * stamp) NameSpace.table;
-  val empty = NameSpace.empty_table;
+  type T = ((src -> attribute) * string) Name_Space.table;
+  val empty : T = Name_Space.empty_table "attribute";
   val copy = I;
   val extend = I;
-  fun merge _ tables : T = NameSpace.merge_tables (eq_snd (op =)) tables handle Symtab.DUP dup =>
-    error ("Attempt to merge different versions of attribute " ^ quote dup);
+  fun merge _ tables : T = Name_Space.merge_tables tables;
 );
 
 fun print_attributes thy =
   let
     val attribs = Attributes.get thy;
-    fun prt_attr (name, ((_, comment), _)) = Pretty.block
+    fun prt_attr (name, (_, comment)) = Pretty.block
       [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
   in
-    [Pretty.big_list "attributes:" (map prt_attr (NameSpace.extern_table attribs))]
+    [Pretty.big_list "attributes:" (map prt_attr (Name_Space.extern_table attribs))]
     |> Pretty.chunks |> Pretty.writeln
   end;
 
-fun add_attribute name att comment thy = thy |> Attributes.map (fn atts =>
-  #2 (NameSpace.define (Sign.naming_of thy) (name, ((att, comment), stamp ())) atts)
-    handle Symtab.DUP dup => error ("Duplicate declaration of attribute " ^ quote dup));
+fun add_attribute name att comment thy = thy
+  |> Attributes.map (#2 o Name_Space.define true (Sign.naming_of thy) (name, (att, comment)));
 
 
 (* name space *)
 
-val intern = NameSpace.intern o #1 o Attributes.get;
+val intern = Name_Space.intern o #1 o Attributes.get;
 val intern_src = Args.map_name o intern;
 
-val extern = NameSpace.extern o #1 o Attributes.get o ProofContext.theory_of;
+val extern = Name_Space.extern o #1 o Attributes.get o ProofContext.theory_of;
 
 
 (* pretty printing *)
@@ -117,7 +115,7 @@
       let val ((name, _), pos) = Args.dest_src src in
         (case Symtab.lookup attrs name of
           NONE => error ("Unknown attribute: " ^ quote name ^ Position.str_of pos)
-        | SOME ((att, _), _) => (Position.report (Markup.attribute name) pos; att src))
+        | SOME (att, _) => (Position.report (Markup.attribute name) pos; att src))
       end;
   in attr end;
 
@@ -340,7 +338,7 @@
         Pretty.block [Pretty.str (name ^ ": " ^ Config.print_type value ^ " ="), Pretty.brk 1,
           Pretty.str (Config.print_value value)]
       end;
-    val configs = NameSpace.extern_table (#1 (Attributes.get thy), Configs.get thy);
+    val configs = Name_Space.extern_table (#1 (Attributes.get thy), Configs.get thy);
   in Pretty.writeln (Pretty.big_list "configuration options" (map prt configs)) end;
 
 
--- a/src/Pure/Isar/class.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/Isar/class.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -237,7 +237,7 @@
         val syn = (the_default NoSyn o AList.lookup Binding.eq_name global_syntax) b;
       in
         thy
-        |> Sign.declare_const [] ((b, ty0), syn)
+        |> Sign.declare_const ((b, ty0), syn)
         |> snd
         |> pair ((Name.of_binding b, ty), (c, ty'))
       end;
--- a/src/Pure/Isar/class_target.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/Isar/class_target.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -21,10 +21,8 @@
 
   val begin: class list -> sort -> Proof.context -> Proof.context
   val init: class -> theory -> Proof.context
-  val declare: class -> Properties.T
-    -> (binding * mixfix) * term -> theory -> theory
-  val abbrev: class -> Syntax.mode -> Properties.T
-    -> (binding * mixfix) * term -> theory -> theory
+  val declare: class -> (binding * mixfix) * term -> theory -> theory
+  val abbrev: class -> Syntax.mode -> (binding * mixfix) * term -> theory -> theory
   val class_prefix: string -> string
   val refresh_syntax: class -> Proof.context -> Proof.context
   val redeclare_operations: theory -> sort -> Proof.context -> Proof.context
@@ -325,7 +323,7 @@
 
 val class_prefix = Logic.const_of_class o Long_Name.base_name;
 
-fun declare class pos ((c, mx), dict) thy =
+fun declare class ((c, mx), dict) thy =
   let
     val morph = morphism thy class;
     val b = Morphism.binding morph c;
@@ -337,7 +335,7 @@
       |> map_types Type.strip_sorts;
   in
     thy
-    |> Sign.declare_const pos ((b, Type.strip_sorts ty'), mx)
+    |> Sign.declare_const ((b, Type.strip_sorts ty'), mx)
     |> snd
     |> Thm.add_def false false (b_def, def_eq)
     |>> Thm.varifyT
@@ -347,7 +345,7 @@
     |> Sign.add_const_constraint (c', SOME ty')
   end;
 
-fun abbrev class prmode pos ((c, mx), rhs) thy =
+fun abbrev class prmode ((c, mx), rhs) thy =
   let
     val morph = morphism thy class;
     val unchecks = these_unchecks thy [class];
@@ -358,7 +356,7 @@
     val rhs'' = map_types Logic.varifyT rhs';
   in
     thy
-    |> Sign.add_abbrev (#1 prmode) pos (b, rhs'')
+    |> Sign.add_abbrev (#1 prmode) (b, rhs'')
     |> snd
     |> Sign.add_const_constraint (c', SOME ty')
     |> Sign.notation true prmode [(Const (c', ty'), mx)]
--- a/src/Pure/Isar/expression.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/Isar/expression.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -639,7 +639,7 @@
     val ([pred_def], defs_thy) =
       thy
       |> bodyT = propT ? Sign.add_advanced_trfuns ([], [], [aprop_tr' (length args) name], [])
-      |> Sign.declare_const [] ((bname, predT), NoSyn) |> snd
+      |> Sign.declare_const ((bname, predT), NoSyn) |> snd
       |> PureThy.add_defs false
         [((Binding.map_name Thm.def_name bname, Logic.mk_equals (head, body)), [Thm.kind_internal])];
     val defs_ctxt = ProofContext.init defs_thy |> Variable.declare_term head;
--- a/src/Pure/Isar/isar_cmd.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/Isar/isar_cmd.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -400,7 +400,7 @@
     val {classes = (space, algebra), ...} = Type.rep_tsig (Sign.tsig_of thy);
     val {classes, ...} = Sorts.rep_algebra algebra;
     fun entry (c, (i, (_, cs))) =
-      (i, {name = NameSpace.extern space c, ID = c, parents = cs,
+      (i, {name = Name_Space.extern space c, ID = c, parents = cs,
             dir = "", unfold = true, path = ""});
     val gr =
       Graph.fold (cons o entry) classes []
--- a/src/Pure/Isar/local_theory.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/Isar/local_theory.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -9,22 +9,21 @@
 signature LOCAL_THEORY =
 sig
   type operations
-  val group_of: local_theory -> string
-  val group_properties_of: local_theory -> Properties.T
-  val group_position_of: local_theory -> Properties.T
-  val set_group: string -> local_theory -> local_theory
+  val affirm: local_theory -> local_theory
+  val naming_of: local_theory -> Name_Space.naming
+  val full_name: local_theory -> binding -> string
+  val map_naming: (Name_Space.naming -> Name_Space.naming) -> local_theory -> local_theory
+  val conceal: local_theory -> local_theory
+  val set_group: serial -> local_theory -> local_theory
   val target_of: local_theory -> Proof.context
   val raw_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory
   val raw_theory: (theory -> theory) -> local_theory -> local_theory
   val checkpoint: local_theory -> local_theory
-  val full_naming: local_theory -> NameSpace.naming
-  val full_name: local_theory -> binding -> string
   val theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory
   val theory: (theory -> theory) -> local_theory -> local_theory
   val target_result: (Proof.context -> 'a * Proof.context) -> local_theory -> 'a * local_theory
   val target: (Proof.context -> Proof.context) -> local_theory -> local_theory
   val map_contexts: (Context.generic -> Context.generic) -> local_theory -> local_theory
-  val affirm: local_theory -> local_theory
   val pretty: local_theory -> Pretty.T list
   val abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory ->
     (term * term) * local_theory
@@ -71,13 +70,14 @@
   exit: local_theory -> Proof.context};
 
 datatype lthy = LThy of
- {group: string,
+ {naming: Name_Space.naming,
   theory_prefix: string,
   operations: operations,
   target: Proof.context};
 
-fun make_lthy (group, theory_prefix, operations, target) =
-  LThy {group = group, theory_prefix = theory_prefix, operations = operations, target = target};
+fun make_lthy (naming, theory_prefix, operations, target) =
+  LThy {naming = naming, theory_prefix = theory_prefix,
+    operations = operations, target = target};
 
 
 (* context data *)
@@ -94,33 +94,30 @@
   | SOME (LThy data) => data);
 
 fun map_lthy f lthy =
-  let val {group, theory_prefix, operations, target} = get_lthy lthy
-  in Data.put (SOME (make_lthy (f (group, theory_prefix, operations, target)))) lthy end;
+  let val {naming, theory_prefix, operations, target} = get_lthy lthy
+  in Data.put (SOME (make_lthy (f (naming, theory_prefix, operations, target)))) lthy end;
+
+val affirm = tap get_lthy;
 
 
-(* group *)
+(* naming *)
 
-val group_of = #group o get_lthy;
+val naming_of = #naming o get_lthy;
+val full_name = Name_Space.full_name o naming_of;
 
-fun group_properties_of lthy =
-  (case group_of lthy of
-    "" => []
-  | group => [(Markup.groupN, group)]);
+fun map_naming f = map_lthy (fn (naming, theory_prefix, operations, target) =>
+  (f naming, theory_prefix, operations, target));
 
-fun group_position_of lthy =
-  group_properties_of lthy @ Position.properties_of (Position.thread_data ());
-
-fun set_group group = map_lthy (fn (_, theory_prefix, operations, target) =>
-  (group, theory_prefix, operations, target));
+val conceal = map_naming Name_Space.conceal;
+val set_group = map_naming o Name_Space.set_group;
 
 
 (* target *)
 
 val target_of = #target o get_lthy;
-val affirm = tap target_of;
 
-fun map_target f = map_lthy (fn (group, theory_prefix, operations, target) =>
-  (group, theory_prefix, operations, f target));
+fun map_target f = map_lthy (fn (naming, theory_prefix, operations, target) =>
+  (naming, theory_prefix, operations, f target));
 
 
 (* substructure mappings *)
@@ -137,16 +134,12 @@
 
 val checkpoint = raw_theory Theory.checkpoint;
 
-fun full_naming lthy =
-  Sign.naming_of (ProofContext.theory_of lthy)
-  |> NameSpace.mandatory_path (#theory_prefix (get_lthy lthy));
-
-fun full_name naming = NameSpace.full_name (full_naming naming);
-
-fun theory_result f lthy = lthy |> raw_theory_result (fn thy => thy
-  |> Sign.mandatory_path (#theory_prefix (get_lthy lthy))
-  |> f
-  ||> Sign.restore_naming thy);
+fun theory_result f lthy =
+  lthy |> raw_theory_result (fn thy =>
+    thy
+    |> Sign.map_naming (K (naming_of lthy))
+    |> f
+    ||> Sign.restore_naming thy);
 
 fun theory f = #2 o theory_result (f #> pair ());
 
@@ -196,13 +189,19 @@
 
 (* init *)
 
-fun init theory_prefix operations target = target |> Data.map
-  (fn NONE => SOME (make_lthy ("", theory_prefix, operations, target))
-    | SOME _ => error "Local theory already initialized")
-  |> checkpoint;
+fun init theory_prefix operations target =
+  let val naming =
+    Sign.naming_of (ProofContext.theory_of target)
+    |> Name_Space.mandatory_path theory_prefix;
+  in
+    target |> Data.map
+      (fn NONE => SOME (make_lthy (naming, theory_prefix, operations, target))
+        | SOME _ => error "Local theory already initialized")
+    |> checkpoint
+  end;
 
 fun restore lthy =
-  let val {group = _, theory_prefix, operations, target} = get_lthy lthy
+  let val {theory_prefix, operations, target, ...} = get_lthy lthy
   in init theory_prefix operations target end;
 
 val reinit = checkpoint o operation #reinit;
--- a/src/Pure/Isar/locale.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/Isar/locale.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -124,15 +124,15 @@
 
 structure Locales = TheoryDataFun
 (
-  type T = locale NameSpace.table;
-  val empty = NameSpace.empty_table;
+  type T = locale Name_Space.table;
+  val empty : T = Name_Space.empty_table "locale";
   val copy = I;
   val extend = I;
-  fun merge _ = NameSpace.join_tables (K merge_locale);
+  fun merge _ = Name_Space.join_tables (K merge_locale);
 );
 
-val intern = NameSpace.intern o #1 o Locales.get;
-val extern = NameSpace.extern o #1 o Locales.get;
+val intern = Name_Space.intern o #1 o Locales.get;
+val extern = Name_Space.extern o #1 o Locales.get;
 
 val get_locale = Symtab.lookup o #2 o Locales.get;
 val defined = Symtab.defined o #2 o Locales.get;
@@ -143,7 +143,7 @@
   | NONE => error ("Unknown locale " ^ quote name));
 
 fun register_locale binding parameters spec intros axioms decls notes dependencies thy =
-  thy |> Locales.map (NameSpace.define (Sign.naming_of thy)
+  thy |> Locales.map (Name_Space.define true (Sign.naming_of thy)
     (binding,
       mk_locale ((parameters, spec, intros, axioms),
         ((pairself (map (fn decl => (decl, stamp ()))) decls, map (fn n => (n, stamp ())) notes),
@@ -153,7 +153,7 @@
   Locales.map o apsnd o Symtab.map_entry name o map_locale o apsnd;
 
 fun print_locales thy =
-  Pretty.strs ("locales:" :: map #1 (NameSpace.extern_table (Locales.get thy)))
+  Pretty.strs ("locales:" :: map #1 (Name_Space.extern_table (Locales.get thy)))
   |> Pretty.writeln;
 
 
--- a/src/Pure/Isar/method.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/Isar/method.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -322,32 +322,30 @@
 
 structure Methods = TheoryDataFun
 (
-  type T = (((src -> Proof.context -> method) * string) * stamp) NameSpace.table;
-  val empty = NameSpace.empty_table;
+  type T = ((src -> Proof.context -> method) * string) Name_Space.table;
+  val empty : T = Name_Space.empty_table "method";
   val copy = I;
   val extend = I;
-  fun merge _ tables : T = NameSpace.merge_tables (eq_snd (op =)) tables handle Symtab.DUP dup =>
-    error ("Attempt to merge different versions of method " ^ quote dup);
+  fun merge _ tables : T = Name_Space.merge_tables tables;
 );
 
 fun print_methods thy =
   let
     val meths = Methods.get thy;
-    fun prt_meth (name, ((_, comment), _)) = Pretty.block
+    fun prt_meth (name, (_, comment)) = Pretty.block
       [Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
   in
-    [Pretty.big_list "methods:" (map prt_meth (NameSpace.extern_table meths))]
+    [Pretty.big_list "methods:" (map prt_meth (Name_Space.extern_table meths))]
     |> Pretty.chunks |> Pretty.writeln
   end;
 
-fun add_method name meth comment thy = thy |> Methods.map (fn meths =>
-  #2 (NameSpace.define (Sign.naming_of thy) (name, ((meth, comment), stamp ())) meths)
-    handle Symtab.DUP dup => error ("Duplicate declaration of method " ^ quote dup));
+fun add_method name meth comment thy = thy
+  |> Methods.map (#2 o Name_Space.define true (Sign.naming_of thy) (name, (meth, comment)));
 
 
 (* get methods *)
 
-val intern = NameSpace.intern o #1 o Methods.get;
+val intern = Name_Space.intern o #1 o Methods.get;
 val defined = Symtab.defined o #2 o Methods.get;
 
 fun method_i thy =
@@ -357,11 +355,11 @@
       let val ((name, _), pos) = Args.dest_src src in
         (case Symtab.lookup meths name of
           NONE => error ("Unknown proof method: " ^ quote name ^ Position.str_of pos)
-        | SOME ((mth, _), _) => (Position.report (Markup.method name) pos; mth src))
+        | SOME (mth, _) => (Position.report (Markup.method name) pos; mth src))
       end;
   in meth end;
 
-fun method thy = method_i thy o Args.map_name (NameSpace.intern (#1 (Methods.get thy)));
+fun method thy = method_i thy o Args.map_name (Name_Space.intern (#1 (Methods.get thy)));
 
 
 (* method setup *)
--- a/src/Pure/Isar/object_logic.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/Isar/object_logic.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -90,7 +90,7 @@
     val base_sort = get_base_sort thy;
     val b = Binding.map_name (Syntax.type_name mx) a;
     val _ = has_duplicates (op =) vs andalso
-      error ("Duplicate parameters in type declaration: " ^ quote (Binding.str_of b));
+      error ("Duplicate parameters in type declaration " ^ quote (Binding.str_of b));
     val name = Sign.full_name thy b;
     val n = length vs;
     val T = Type (name, map (fn v => TFree (v, [])) vs);
--- a/src/Pure/Isar/proof.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/Isar/proof.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -162,7 +162,8 @@
   make_node (f (context, facts, mode, goal));
 
 val init_context =
-  ProofContext.set_stmt true #> ProofContext.reset_naming;
+  ProofContext.set_stmt true #>
+  ProofContext.map_naming (K ProofContext.local_naming);
 
 fun init ctxt =
   State (Stack.init (make_node (init_context ctxt, NONE, Forward, NONE)));
--- a/src/Pure/Isar/proof_context.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/Isar/proof_context.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -21,7 +21,10 @@
   val restore_mode: Proof.context -> Proof.context -> Proof.context
   val abbrev_mode: Proof.context -> bool
   val set_stmt: bool -> Proof.context -> Proof.context
-  val naming_of: Proof.context -> NameSpace.naming
+  val local_naming: Name_Space.naming
+  val map_naming: (Name_Space.naming -> Name_Space.naming) -> Proof.context -> Proof.context
+  val naming_of: Proof.context -> Name_Space.naming
+  val restore_naming: Proof.context -> Proof.context -> Proof.context
   val full_name: Proof.context -> binding -> string
   val consts_of: Proof.context -> Consts.T
   val const_syntax_name: Proof.context -> string -> string
@@ -92,10 +95,6 @@
   val get_fact_single: Proof.context -> Facts.ref -> thm
   val get_thms: Proof.context -> xstring -> thm list
   val get_thm: Proof.context -> xstring -> thm
-  val add_path: string -> Proof.context -> Proof.context
-  val mandatory_path: string -> Proof.context -> Proof.context
-  val restore_naming: Proof.context -> Proof.context -> Proof.context
-  val reset_naming: Proof.context -> Proof.context
   val note_thmss: string -> (Thm.binding * (thm list * attribute list) list) list ->
     Proof.context -> (string * thm list) list * Proof.context
   val put_thms: bool -> string * thm list option -> Proof.context -> Proof.context
@@ -120,8 +119,7 @@
   val target_notation: bool -> Syntax.mode -> (term * mixfix) list -> morphism ->
     Context.generic -> Context.generic
   val add_const_constraint: string * typ option -> Proof.context -> Proof.context
-  val add_abbrev: string -> Properties.T ->
-    binding * term -> Proof.context -> (term * term) * Proof.context
+  val add_abbrev: string -> binding * term -> Proof.context -> (term * term) * Proof.context
   val revert_abbrev: string -> string -> Proof.context -> Proof.context
   val verbose: bool Unsynchronized.ref
   val setmp_verbose: ('a -> 'b) -> 'a -> 'b
@@ -134,9 +132,6 @@
   val prems_limit: int Unsynchronized.ref
   val pretty_ctxt: Proof.context -> Pretty.T list
   val pretty_context: Proof.context -> Pretty.T list
-  val query_type: Proof.context -> string -> Properties.T
-  val query_const: Proof.context -> string -> Properties.T
-  val query_class: Proof.context -> string -> Properties.T
 end;
 
 structure ProofContext: PROOF_CONTEXT =
@@ -170,7 +165,7 @@
 datatype ctxt =
   Ctxt of
    {mode: mode,                                       (*inner syntax mode*)
-    naming: NameSpace.naming,                         (*local naming conventions*)
+    naming: Name_Space.naming,                        (*local naming conventions*)
     syntax: LocalSyntax.T,                            (*local syntax*)
     consts: Consts.T * Consts.T,                      (*local/global consts*)
     facts: Facts.T,                                   (*local facts*)
@@ -180,7 +175,7 @@
   Ctxt {mode = mode, naming = naming, syntax = syntax,
     consts = consts, facts = facts, cases = cases};
 
-val local_naming = NameSpace.default_naming |> NameSpace.add_path "local";
+val local_naming = Name_Space.default_naming |> Name_Space.add_path "local";
 
 structure ContextData = ProofDataFun
 (
@@ -231,7 +226,8 @@
   map_mode (fn (_, pattern, schematic, abbrev) => (stmt, pattern, schematic, abbrev));
 
 val naming_of = #naming o rep_context;
-val full_name = NameSpace.full_name o naming_of;
+val restore_naming = map_naming o K o naming_of
+val full_name = Name_Space.full_name o naming_of;
 
 val syntax_of = #syntax o rep_context;
 val syn_of = LocalSyntax.syn_of o syntax_of;
@@ -922,14 +918,6 @@
 end;
 
 
-(* name space operations *)
-
-val add_path        = map_naming o NameSpace.add_path;
-val mandatory_path  = map_naming o NameSpace.mandatory_path;
-val restore_naming  = map_naming o K o naming_of;
-val reset_naming    = map_naming (K local_naming);
-
-
 (* facts *)
 
 local
@@ -1059,13 +1047,13 @@
       in cert_term ctxt (Const (c, T)); T end;
   in ctxt |> (map_consts o apfst) (Consts.constrain (c, Option.map prepT opt_T)) end;
 
-fun add_abbrev mode tags (b, raw_t) ctxt =
+fun add_abbrev mode (b, raw_t) ctxt =
   let
     val t0 = cert_term (ctxt |> set_mode mode_abbrev) raw_t
       handle ERROR msg => cat_error msg ("in constant abbreviation " ^ quote (Binding.str_of b));
     val [t] = Variable.exportT_terms (Variable.declare_term t0 ctxt) ctxt [t0];
     val ((lhs, rhs), consts') = consts_of ctxt
-      |> Consts.abbreviate (Syntax.pp ctxt) (tsig_of ctxt) (naming_of ctxt) mode tags (b, t);
+      |> Consts.abbreviate (Syntax.pp ctxt) (tsig_of ctxt) (naming_of ctxt) mode (b, t);
   in
     ctxt
     |> (map_consts o apfst) (K consts')
@@ -1230,7 +1218,7 @@
       | add_abbr (c, (T, SOME t)) =
           if not show_globals andalso Symtab.defined globals c then I
           else cons (c, Logic.mk_equals (Const (c, T), t));
-    val abbrevs = NameSpace.extern_table (space, Symtab.make (Symtab.fold add_abbr consts []));
+    val abbrevs = Name_Space.extern_table (space, Symtab.make (Symtab.fold add_abbr consts []));
   in
     if null abbrevs andalso not (! verbose) then []
     else [Pretty.big_list "abbreviations:" (map (pretty_term_abbrev ctxt o #2) abbrevs)]
@@ -1391,14 +1379,4 @@
     verb single (fn () => Pretty.big_list "default sorts:" (map prt_defS (Vartab.dest sorts)))
   end;
 
-
-(* query meta data *)
-
-val query_type = Type.the_tags o tsig_of;
-
-fun query_const ctxt name =
-  Consts.the_tags (consts_of ctxt) name handle TYPE (msg, _, _) => error msg;
-
-fun query_class ctxt name = query_const ctxt (Logic.const_of_class name);
-
 end;
--- a/src/Pure/Isar/specification.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/Isar/specification.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -161,7 +161,7 @@
     val xs = map (fn ((b, T), _) => (Name.of_binding b, T)) vars;
 
     (*consts*)
-    val (consts, consts_thy) = thy |> fold_map (Theory.specify_const []) vars;
+    val (consts, consts_thy) = thy |> fold_map Theory.specify_const vars;
     val subst = Term.subst_atomic (map Free xs ~~ consts);
 
     (*axioms*)
--- a/src/Pure/Isar/theory_target.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/Isar/theory_target.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -158,8 +158,8 @@
     val global_facts = PureThy.map_facts #2 facts'
       |> Attrib.map_facts (if is_locale then K I else Attrib.attribute_i thy);
   in
-    lthy |> LocalTheory.theory
-      (PureThy.note_thmss_grouped kind (LocalTheory.group_of lthy) global_facts #> snd)
+    lthy
+    |> LocalTheory.theory (PureThy.note_thmss kind global_facts #> snd)
     |> not is_locale ? LocalTheory.target (ProofContext.note_thmss kind global_facts #> snd)
     |> is_locale ? LocalTheory.target (Locale.add_thmss target kind target_facts)
     |> ProofContext.note_thmss kind local_facts
@@ -173,7 +173,7 @@
   else if not is_class then (NoSyn, mx, NoSyn)
   else (mx, NoSyn, NoSyn);
 
-fun locale_const (Target {target, is_class, ...}) (prmode as (mode, _)) tags ((b, mx), rhs) phi =
+fun locale_const (Target {target, is_class, ...}) (prmode as (mode, _)) ((b, mx), rhs) phi =
   let
     val b' = Morphism.binding phi b;
     val rhs' = Morphism.term phi rhs;
@@ -187,8 +187,8 @@
   in
     not (is_class andalso (similar_body orelse class_global)) ?
       (Context.mapping_result
-        (Sign.add_abbrev PrintMode.internal tags arg)
-        (ProofContext.add_abbrev PrintMode.internal tags arg)
+        (Sign.add_abbrev PrintMode.internal arg)
+        (ProofContext.add_abbrev PrintMode.internal arg)
       #-> (fn (lhs' as Const (d, _), _) =>
           similar_body ?
             (Context.mapping (Sign.revert_abbrev mode d) (ProofContext.revert_abbrev mode d) #>
@@ -199,7 +199,6 @@
 
 fun declare_const (ta as Target {target, is_locale, is_class, ...}) depends ((b, T), mx) lthy =
   let
-    val tags = LocalTheory.group_position_of lthy;
     val xs = filter depends (#1 (ProofContext.inferred_fixes (LocalTheory.target_of lthy)));
     val U = map #2 xs ---> T;
     val (mx1, mx2, mx3) = fork_mixfix ta mx;
@@ -215,13 +214,13 @@
                 if mx3 <> NoSyn then syntax_error c'
                 else LocalTheory.theory_result (Overloading.declare (c', U))
                   ##> Overloading.confirm b
-            | NONE => LocalTheory.theory_result (Sign.declare_const tags ((b, U), mx3))));
+            | NONE => LocalTheory.theory_result (Sign.declare_const ((b, U), mx3))));
     val (const, lthy') = lthy |> declare_const;
     val t = Term.list_comb (const, map Free xs);
   in
     lthy'
-    |> is_locale ? term_syntax ta (locale_const ta Syntax.mode_default tags ((b, mx2), t))
-    |> is_class ? class_target ta (Class_Target.declare target tags ((b, mx1), t))
+    |> is_locale ? term_syntax ta (locale_const ta Syntax.mode_default ((b, mx2), t))
+    |> is_class ? class_target ta (Class_Target.declare target ((b, mx1), t))
     |> LocalDefs.add_def ((b, NoSyn), t)
   end;
 
@@ -230,7 +229,6 @@
 
 fun abbrev (ta as Target {target, is_locale, is_class, ...}) prmode ((b, mx), t) lthy =
   let
-    val tags = LocalTheory.group_position_of lthy;
     val thy_ctxt = ProofContext.init (ProofContext.theory_of lthy);
     val target_ctxt = LocalTheory.target_of lthy;
 
@@ -243,17 +241,17 @@
   in
     lthy |>
      (if is_locale then
-        LocalTheory.theory_result (Sign.add_abbrev PrintMode.internal tags (b, global_rhs))
+        LocalTheory.theory_result (Sign.add_abbrev PrintMode.internal (b, global_rhs))
         #-> (fn (lhs, _) =>
           let val lhs' = Term.list_comb (Logic.unvarify lhs, xs) in
-            term_syntax ta (locale_const ta prmode tags ((b, mx2), lhs')) #>
-            is_class ? class_target ta (Class_Target.abbrev target prmode tags ((b, mx1), t'))
+            term_syntax ta (locale_const ta prmode ((b, mx2), lhs')) #>
+            is_class ? class_target ta (Class_Target.abbrev target prmode ((b, mx1), t'))
           end)
       else
         LocalTheory.theory
-          (Sign.add_abbrev (#1 prmode) tags (b, global_rhs) #-> (fn (lhs, _) =>
+          (Sign.add_abbrev (#1 prmode) (b, global_rhs) #-> (fn (lhs, _) =>
            Sign.notation true prmode [(lhs, mx3)])))
-    |> ProofContext.add_abbrev PrintMode.internal tags (b, t) |> snd
+    |> ProofContext.add_abbrev PrintMode.internal (b, t) |> snd
     |> LocalDefs.fixed_abbrev ((b, NoSyn), t)
   end;
 
--- a/src/Pure/Thy/thm_deps.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/Thy/thm_deps.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -48,37 +48,49 @@
 
 fun unused_thms (base_thys, thys) =
   let
-    fun add_fact (name, ths) =
+    fun add_fact space (name, ths) =
       if exists (fn thy => PureThy.defined_fact thy name) base_thys then I
-      else fold_rev (fn th => (case Thm.get_name th of "" => I | a => cons (a, th))) ths;
+      else
+        let val {concealed, group, ...} = Name_Space.the_entry space name in
+          fold_rev (fn th =>
+            (case Thm.get_name th of
+              "" => I
+            | a => cons (a, (th, concealed, group)))) ths
+        end;
+    fun add_facts facts = Facts.fold_static (add_fact (Facts.space_of facts)) facts;
+
     val thms =
-      fold (Facts.fold_static add_fact o PureThy.facts_of) thys []
+      fold (add_facts o PureThy.facts_of) thys []
       |> sort_distinct (string_ord o pairself #1);
 
     val tab =
       Proofterm.fold_body_thms
-        (fn (name, prop, _) => name <> "" ? Symtab.insert_list (op =) (name, prop))
-        (map (Proofterm.strip_thm o Thm.proof_body_of o snd) thms) Symtab.empty;
-    fun is_unused (name, th) =
-      not (member (op aconv) (Symtab.lookup_list tab name) (Thm.prop_of th));
+        (fn (a, prop, _) => a <> "" ? Symtab.insert_list (op =) (a, prop))
+        (map (Proofterm.strip_thm o Thm.proof_body_of o #1 o #2) thms) Symtab.empty;
+
+    fun is_unused (a, th) =
+      not (member (op aconv) (Symtab.lookup_list tab a) (Thm.prop_of th));
 
     (* groups containing at least one used theorem *)
-    val grps = fold (fn p as (_, thm) => if is_unused p then I else
-      case Thm.get_group thm of
-        NONE => I | SOME grp => Symtab.update (grp, ())) thms Symtab.empty;
-    val (thms', _) = fold (fn p as (s, thm) => fn q as (thms, grps') =>
-      if member (op =) [Thm.theoremK, Thm.generatedK, Thm.lemmaK, Thm.corollaryK, Thm.axiomK] (Thm.get_kind thm)
-        andalso is_unused p
+    val used_groups = fold (fn (a, (th, _, group)) =>
+      if is_unused (a, th) then I
+      else
+        (case group of
+          NONE => I
+        | SOME grp => Inttab.update (grp, ()))) thms Inttab.empty;
+    val (thms', _) = fold (fn (a, (th, concealed, group)) => fn q as (thms, grps') =>
+      if member (op =) [Thm.theoremK, Thm.generatedK, Thm.lemmaK, Thm.corollaryK, Thm.axiomK]
+        (Thm.get_kind th) andalso not concealed andalso is_unused (a, th)
       then
-        (case Thm.get_group thm of
-           NONE => (p :: thms, grps')
+        (case group of
+           NONE => ((a, th) :: thms, grps')
          | SOME grp =>
              (* do not output theorem if another theorem in group has been used *)
-             if Symtab.defined grps grp then q
+             if Inttab.defined used_groups grp then q
              (* output at most one unused theorem per group *)
-             else if Symtab.defined grps' grp then q
-             else (p :: thms, Symtab.update (grp, ()) grps'))
-      else q) thms ([], Symtab.empty)
+             else if Inttab.defined grps' grp then q
+             else ((a, th) :: thms, Inttab.update (grp, ()) grps'))
+      else q) thms ([], Inttab.empty)
   in rev thms' end;
 
 end;
--- a/src/Pure/Thy/thy_output.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/Thy/thy_output.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -420,9 +420,9 @@
   ("show_sorts", setmp_CRITICAL Syntax.show_sorts o boolean),
   ("show_structs", setmp_CRITICAL show_structs o boolean),
   ("show_question_marks", setmp_CRITICAL show_question_marks o boolean),
-  ("long_names", setmp_CRITICAL NameSpace.long_names o boolean),
-  ("short_names", setmp_CRITICAL NameSpace.short_names o boolean),
-  ("unique_names", setmp_CRITICAL NameSpace.unique_names o boolean),
+  ("long_names", setmp_CRITICAL Name_Space.long_names o boolean),
+  ("short_names", setmp_CRITICAL Name_Space.short_names o boolean),
+  ("unique_names", setmp_CRITICAL Name_Space.unique_names o boolean),
   ("eta_contract", setmp_CRITICAL Syntax.eta_contract o boolean),
   ("display", setmp_CRITICAL display o boolean),
   ("break", setmp_CRITICAL break o boolean),
--- a/src/Pure/Tools/find_consts.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/Tools/find_consts.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -87,9 +87,8 @@
     val thy = ProofContext.theory_of ctxt;
     val low_ranking = 10000;
 
-    fun not_internal consts (nm, _) =
-      if member (op =) (Consts.the_tags consts nm) Markup.property_internal
-      then NONE else SOME low_ranking;
+    fun user_visible consts (nm, _) =
+      if Consts.is_concealed consts nm then NONE else SOME low_ranking;
 
     fun make_pattern crit =
       let
@@ -119,7 +118,7 @@
     val consts = Sign.consts_of thy;
     val (_, consts_tab) = #constants (Consts.dest consts);
     fun eval_entry c =
-      fold filter_const (not_internal consts :: criteria)
+      fold filter_const (user_visible consts :: criteria)
         (SOME (c, low_ranking));
 
     val matches =
--- a/src/Pure/Tools/find_theorems.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/Tools/find_theorems.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -326,7 +326,7 @@
 local
 
 val index_ord = option_ord (K EQUAL);
-val hidden_ord = bool_ord o pairself NameSpace.is_hidden;
+val hidden_ord = bool_ord o pairself Name_Space.is_hidden;
 val qual_ord = int_ord o pairself (length o Long_Name.explode);
 val txt_ord = int_ord o pairself size;
 
@@ -355,7 +355,8 @@
     val space = Facts.space_of (PureThy.facts_of (ProofContext.theory_of ctxt));
 
     val shorten =
-      NameSpace.extern_flags {long_names = false, short_names = false, unique_names = false} space;
+      Name_Space.extern_flags
+        {long_names = false, short_names = false, unique_names = false} space;
 
     fun nicer (Facts.Named ((x, _), i)) (Facts.Named ((y, _), j)) =
           nicer_name (shorten x, i) (shorten y, j)
--- a/src/Pure/axclass.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/axclass.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -36,7 +36,7 @@
   val lookup_inst_param: Consts.T -> ((string * string) * 'a) list -> string * typ -> 'a option
   val param_of_inst: theory -> string * string -> string
   val inst_of_param: theory -> string -> (string * string) option
-  val arity_property: theory -> class * string -> string -> string list
+  val thynames_of_arity: theory -> class * string -> string list
   type cache
   val of_sort: theory -> typ * sort -> cache -> thm list * cache  (*exception Sorts.CLASS_ERROR*)
   val cache: cache
@@ -92,8 +92,8 @@
 val arity_prefix = "arity_";
 
 type instances =
-  ((class * class) * thm) list *
-  ((class * sort list) * thm) list Symtab.table;
+  ((class * class) * thm) list *  (*classrel theorems*)
+  ((class * sort list) * (thm * string)) list Symtab.table;  (*arity theorems with theory name*)
 
 fun merge_instances ((classrel1, arities1): instances, (classrel2, arities2)) =
  (merge (eq_fst op =) (classrel1, classrel2),
@@ -175,35 +175,32 @@
 
 fun the_arity thy a (c, Ss) =
   (case AList.lookup (op =) (Symtab.lookup_list (#2 (get_instances thy)) a) (c, Ss) of
-    SOME th => Thm.transfer thy th
+    SOME (th, _) => Thm.transfer thy th
   | NONE => error ("Unproven type arity " ^
       Syntax.string_of_arity (ProofContext.init thy) (a, Ss, [c])));
 
-fun arity_property thy (c, a) x =
-  Symtab.lookup_list (snd (get_instances thy)) a
-  |> map_filter (fn ((c', _), th) => if c = c'
-      then AList.lookup (op =) (Thm.get_tags th) x else NONE)
+fun thynames_of_arity thy (c, a) =
+  Symtab.lookup_list (#2 (get_instances thy)) a
+  |> map_filter (fn ((c', _), (_, name)) => if c = c' then SOME name else NONE)
   |> rev;
 
-fun insert_arity_completions thy (t, ((c, Ss), th)) arities =
+fun insert_arity_completions thy (t, ((c, Ss), (th, thy_name))) arities =
   let
     val algebra = Sign.classes_of thy;
-    val super_class_completions = Sign.super_classes thy c
+    val super_class_completions =
+      Sign.super_classes thy c
       |> filter_out (fn c1 => exists (fn ((c2, Ss2), _) => c1 = c2
-          andalso Sorts.sorts_le algebra (Ss2, Ss)) (Symtab.lookup_list arities t))
-    val tags = Thm.get_tags th;
+          andalso Sorts.sorts_le algebra (Ss2, Ss)) (Symtab.lookup_list arities t));
     val completions = map (fn c1 => (Sorts.weaken algebra
       (fn (th, c2) => fn c3 => th RS the_classrel thy (c2, c3)) (th, c) c1
-        |> Thm.map_tags (K tags) |> Thm.close_derivation, c1)) super_class_completions;
-    val arities' = fold (fn (th1, c1) => Symtab.cons_list (t, ((c1, Ss), th1)))
+        |> Thm.close_derivation, c1)) super_class_completions;
+    val arities' = fold (fn (th1, c1) => Symtab.cons_list (t, ((c1, Ss), (th1, thy_name))))
       completions arities;
-  in (completions, arities') end;
+  in (null completions, arities') end;
 
 fun put_arity ((t, Ss, c), th) thy =
   let
-    val th' = (Thm.map_tags o AList.default (op =))
-      (Markup.theory_nameN, Context.theory_name thy) th;
-    val arity' = (t, ((c, Ss), th'));
+    val arity' = (t, ((c, Ss), (th, Context.theory_name thy)));
   in
     thy
     |> map_instances (fn (classrel, arities) => (classrel,
@@ -216,11 +213,10 @@
 fun complete_arities thy =
   let
     val arities = snd (get_instances thy);
-    val (completions, arities') = arities
-      |> fold_map (insert_arity_completions thy) (Symtab.dest_list arities)
-      |>> flat;
-  in if null completions
-    then NONE
+    val (finished, arities') = arities
+      |> fold_map (insert_arity_completions thy) (Symtab.dest_list arities);
+  in
+    if forall I finished then NONE
     else SOME (thy |> map_instances (fn (classrel, _) => (classrel, arities')))
   end;
 
@@ -309,7 +305,7 @@
   in
     thy
     |> Sign.mandatory_path name_inst
-    |> Sign.declare_const [] ((Binding.name c', T'), NoSyn)
+    |> Sign.declare_const ((Binding.name c', T'), NoSyn)
     |-> (fn const' as Const (c'', _) =>
       Thm.add_def false true
         (Binding.name (Thm.def_name c'), Logic.mk_equals (Const (c, T'), const'))
--- a/src/Pure/codegen.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/codegen.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -337,15 +337,16 @@
     val tc = Sign.intern_type thy s;
   in
     case Symtab.lookup (snd (#types (Type.rep_tsig (Sign.tsig_of thy)))) tc of
-      SOME ((Type.LogicalType i, _), _) =>
+      SOME (Type.LogicalType i) =>
         if num_args_of (fst syn) > i then
           error ("More arguments than corresponding type constructor " ^ s)
-        else (case AList.lookup (op =) types tc of
-          NONE => CodegenData.put {codegens = codegens,
-            tycodegens = tycodegens, consts = consts,
-            types = (tc, syn) :: types,
-            preprocs = preprocs, modules = modules} thy
-        | SOME _ => error ("Type " ^ tc ^ " already associated with code"))
+        else
+          (case AList.lookup (op =) types tc of
+            NONE => CodegenData.put {codegens = codegens,
+              tycodegens = tycodegens, consts = consts,
+              types = (tc, syn) :: types,
+              preprocs = preprocs, modules = modules} thy
+          | SOME _ => error ("Type " ^ tc ^ " already associated with code"))
     | _ => error ("Not a type constructor: " ^ s)
   end;
 
@@ -445,13 +446,8 @@
 fun map_node k f (gr, x) = (Graph.map_node k f gr, x);
 fun new_node p (gr, x) = (Graph.new_node p gr, x);
 
-fun thyname_of thy f x = the (AList.lookup (op =) (f x) Markup.theory_nameN);
-
-fun thyname_of_type thy =
-  thyname_of thy (Type.the_tags (Sign.tsig_of thy));
-
-fun thyname_of_const thy =
-  thyname_of thy (Consts.the_tags (Sign.consts_of thy));
+fun thyname_of_type thy = #theory_name o Name_Space.the_entry (Sign.type_space thy);
+fun thyname_of_const thy = #theory_name o Name_Space.the_entry (Sign.const_space thy);
 
 fun rename_terms ts =
   let
--- a/src/Pure/consts.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/consts.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -11,15 +11,15 @@
   val eq_consts: T * T -> bool
   val retrieve_abbrevs: T -> string list -> term -> (term * term) list
   val dest: T ->
-   {constants: (typ * term option) NameSpace.table,
-    constraints: typ NameSpace.table}
+   {constants: (typ * term option) Name_Space.table,
+    constraints: typ Name_Space.table}
   val the_type: T -> string -> typ                             (*exception TYPE*)
   val the_abbreviation: T -> string -> typ * term              (*exception TYPE*)
   val type_scheme: T -> string -> typ                          (*exception TYPE*)
-  val the_tags: T -> string -> Properties.T                    (*exception TYPE*)
   val is_monomorphic: T -> string -> bool                      (*exception TYPE*)
   val the_constraint: T -> string -> typ                       (*exception TYPE*)
-  val space_of: T -> NameSpace.T
+  val space_of: T -> Name_Space.T
+  val is_concealed: T -> string -> bool
   val intern: T -> xstring -> string
   val extern: T -> string -> xstring
   val extern_early: T -> string -> xstring
@@ -29,9 +29,9 @@
   val certify: Pretty.pp -> Type.tsig -> bool -> T -> term -> term  (*exception TYPE*)
   val typargs: T -> string * typ -> typ list
   val instance: T -> string * typ list -> typ
-  val declare: bool -> NameSpace.naming -> Properties.T -> (binding * typ) -> T -> T
+  val declare: bool -> Name_Space.naming -> binding * typ -> T -> T
   val constrain: string * typ option -> T -> T
-  val abbreviate: Pretty.pp -> Type.tsig -> NameSpace.naming -> string -> Properties.T ->
+  val abbreviate: Pretty.pp -> Type.tsig -> Name_Space.naming -> string ->
     binding * term -> T -> (term * term) * T
   val revert_abbrev: string -> string -> T -> T
   val hide: bool -> string -> T -> T
@@ -46,11 +46,11 @@
 
 (* datatype T *)
 
-type decl = {T: typ, typargs: int list list, tags: Properties.T, authentic: bool};
+type decl = {T: typ, typargs: int list list, authentic: bool};
 type abbrev = {rhs: term, normal_rhs: term, force_expand: bool};
 
 datatype T = Consts of
- {decls: ((decl * abbrev option) * serial) NameSpace.table,
+ {decls: (decl * abbrev option) Name_Space.table,
   constraints: typ Symtab.table,
   rev_abbrevs: (term * term) Item_Net.T Symtab.table};
 
@@ -70,7 +70,8 @@
 
 (* reverted abbrevs *)
 
-val empty_abbrevs = Item_Net.init (fn ((t, u), (t', u')) => t aconv t' andalso u aconv u') #1;
+val empty_abbrevs =
+  Item_Net.init (fn ((t, u), (t', u')) => t aconv t' andalso u aconv u') #1;
 
 fun insert_abbrevs mode abbrs =
   Symtab.map_default (mode, empty_abbrevs) (Item_Net.insert abbrs);
@@ -84,7 +85,7 @@
 
 fun dest (Consts {decls = (space, decls), constraints, ...}) =
  {constants = (space,
-    Symtab.fold (fn (c, (({T, ...}, abbr), _)) =>
+    Symtab.fold (fn (c, ({T, ...}, abbr)) =>
       Symtab.update (c, (T, Option.map #rhs abbr))) decls Symtab.empty),
   constraints = (space, constraints)};
 
@@ -93,7 +94,7 @@
 
 fun the_const (Consts {decls = (_, tab), ...}) c =
   (case Symtab.lookup tab c of
-    SOME (decl, _) => decl
+    SOME decl => decl
   | NONE => raise TYPE ("Unknown constant: " ^ quote c, [], []));
 
 fun the_type consts c =
@@ -109,7 +110,6 @@
 val the_decl = #1 oo the_const;
 val type_scheme = #T oo the_decl;
 val type_arguments = #typargs oo the_decl;
-val the_tags = #tags oo the_decl;
 
 val is_monomorphic = null oo type_arguments;
 
@@ -123,8 +123,10 @@
 
 fun space_of (Consts {decls = (space, _), ...}) = space;
 
-val intern = NameSpace.intern o space_of;
-val extern = NameSpace.extern o space_of;
+val is_concealed = Name_Space.is_concealed o space_of;
+
+val intern = Name_Space.intern o space_of;
+val extern = Name_Space.extern o space_of;
 
 fun extern_early consts c =
   (case try (the_const consts) c of
@@ -221,27 +223,20 @@
 
 (** build consts **)
 
-fun err_dup_const c =
-  error ("Duplicate declaration of constant " ^ quote c);
-
-fun extend_decls naming decl tab = NameSpace.define naming decl tab
-  handle Symtab.DUP c => err_dup_const c;
-
-
 (* name space *)
 
 fun hide fully c = map_consts (fn (decls, constraints, rev_abbrevs) =>
-  (apfst (NameSpace.hide fully c) decls, constraints, rev_abbrevs));
+  (apfst (Name_Space.hide fully c) decls, constraints, rev_abbrevs));
 
 
 (* declarations *)
 
-fun declare authentic naming tags (b, declT) = map_consts (fn (decls, constraints, rev_abbrevs) =>
-  let
-    val tags' = tags |> Position.default_properties (Position.thread_data ());
-    val decl = {T = declT, typargs = typargs_of declT, tags = tags', authentic = authentic};
-    val (_, decls') = decls |> extend_decls naming (b, ((decl, NONE), serial ()));
-  in (decls', constraints, rev_abbrevs) end);
+fun declare authentic naming (b, declT) =
+  map_consts (fn (decls, constraints, rev_abbrevs) =>
+    let
+      val decl = {T = declT, typargs = typargs_of declT, authentic = authentic};
+      val (_, decls') = decls |> Name_Space.define true naming (b, (decl, NONE));
+    in (decls', constraints, rev_abbrevs) end);
 
 
 (* constraints *)
@@ -271,14 +266,14 @@
 
 in
 
-fun abbreviate pp tsig naming mode tags (b, raw_rhs) consts =
+fun abbreviate pp tsig naming mode (b, raw_rhs) consts =
   let
     val cert_term = certify pp tsig false consts;
     val expand_term = certify pp tsig true consts;
     val force_expand = mode = PrintMode.internal;
 
     val _ = Term.exists_subterm Term.is_Var raw_rhs andalso
-      error ("Illegal schematic variables on rhs of abbreviation: " ^ Binding.str_of b);
+      error ("Illegal schematic variables on rhs of abbreviation " ^ quote (Binding.str_of b));
 
     val rhs = raw_rhs
       |> Term.map_types (Type.cert_typ tsig)
@@ -286,15 +281,14 @@
       |> Term.close_schematic_term;
     val normal_rhs = expand_term rhs;
     val T = Term.fastype_of rhs;
-    val lhs = Const (NameSpace.full_name naming b, T);
+    val lhs = Const (Name_Space.full_name naming b, T);
   in
     consts |> map_consts (fn (decls, constraints, rev_abbrevs) =>
       let
-        val tags' = tags |> Position.default_properties (Position.thread_data ());
-        val decl = {T = T, typargs = typargs_of T, tags = tags', authentic = true};
+        val decl = {T = T, typargs = typargs_of T, authentic = true};
         val abbr = {rhs = rhs, normal_rhs = normal_rhs, force_expand = force_expand};
         val (_, decls') = decls
-          |> extend_decls naming (b, ((decl, SOME abbr), serial ()));
+          |> Name_Space.define true naming (b, (decl, SOME abbr));
         val rev_abbrevs' = rev_abbrevs
           |> insert_abbrevs mode (rev_abbrev lhs rhs);
       in (decls', constraints, rev_abbrevs') end)
@@ -313,14 +307,13 @@
 
 (* empty and merge *)
 
-val empty = make_consts (NameSpace.empty_table, Symtab.empty, Symtab.empty);
+val empty = make_consts (Name_Space.empty_table "constant", Symtab.empty, Symtab.empty);
 
 fun merge
    (Consts {decls = decls1, constraints = constraints1, rev_abbrevs = rev_abbrevs1},
     Consts {decls = decls2, constraints = constraints2, rev_abbrevs = rev_abbrevs2}) =
   let
-    val decls' = NameSpace.merge_tables (eq_snd (op =)) (decls1, decls2)
-      handle Symtab.DUP c => err_dup_const c;
+    val decls' = Name_Space.merge_tables (decls1, decls2);
     val constraints' = Symtab.join (K fst) (constraints1, constraints2);
     val rev_abbrevs' = Symtab.join (K Item_Net.merge) (rev_abbrevs1, rev_abbrevs2);
   in make_consts (decls', constraints', rev_abbrevs') end;
--- a/src/Pure/display.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/display.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -146,14 +146,14 @@
       [Pretty.str "default sort:", Pretty.brk 1, prt_sort S];
 
     val tfrees = map (fn v => TFree (v, []));
-    fun pretty_type syn (t, ((Type.LogicalType n, _), _)) =
+    fun pretty_type syn (t, (Type.LogicalType n)) =
           if syn then NONE
           else SOME (prt_typ (Type (t, tfrees (Name.invents Name.context Name.aT n))))
-      | pretty_type syn (t, ((Type.Abbreviation (vs, U, syn'), _), _)) =
+      | pretty_type syn (t, (Type.Abbreviation (vs, U, syn'))) =
           if syn <> syn' then NONE
           else SOME (Pretty.block
             [prt_typ (Type (t, tfrees vs)), Pretty.str " =", Pretty.brk 1, prt_typ U])
-      | pretty_type syn (t, ((Type.Nonterminal, _), _)) =
+      | pretty_type syn (t, Type.Nonterminal) =
           if not syn then NONE
           else SOME (prt_typ (Type (t, [])));
 
@@ -179,25 +179,24 @@
     val {restricts, reducts} = Defs.dest defs;
     val {naming = _, syn = _, tsig, consts} = Sign.rep_sg thy;
     val {constants, constraints} = Consts.dest consts;
-    val extern_const = NameSpace.extern (#1 constants);
+    val extern_const = Name_Space.extern (#1 constants);
     val {classes, default, types, ...} = Type.rep_tsig tsig;
     val (class_space, class_algebra) = classes;
     val {classes, arities} = Sorts.rep_algebra class_algebra;
 
-    val clsses = NameSpace.dest_table (class_space, Symtab.make (Graph.dest classes));
-    val tdecls = NameSpace.dest_table types;
-    val arties = NameSpace.dest_table (Sign.type_space thy, arities);
+    val clsses = Name_Space.dest_table (class_space, Symtab.make (Graph.dest classes));
+    val tdecls = Name_Space.dest_table types;
+    val arties = Name_Space.dest_table (Sign.type_space thy, arities);
 
-    fun prune_const c = not verbose andalso
-      member (op =) (Consts.the_tags consts c) Markup.property_internal;
-    val cnsts = NameSpace.extern_table (#1 constants,
+    fun prune_const c = not verbose andalso Consts.is_concealed consts c;
+    val cnsts = Name_Space.extern_table (#1 constants,
       Symtab.make (filter_out (prune_const o fst) (Symtab.dest (#2 constants))));
 
     val log_cnsts = map_filter (fn (c, (ty, NONE)) => SOME (c, ty) | _ => NONE) cnsts;
     val abbrevs = map_filter (fn (c, (ty, SOME t)) => SOME (c, (ty, t)) | _ => NONE) cnsts;
-    val cnstrs = NameSpace.extern_table constraints;
+    val cnstrs = Name_Space.extern_table constraints;
 
-    val axms = NameSpace.extern_table axioms;
+    val axms = Name_Space.extern_table axioms;
 
     val (reds0, (reds1, reds2)) = filter_out (prune_const o fst o fst) reducts
       |> map (fn (lhs, rhs) =>
--- a/src/Pure/drule.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/drule.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -452,7 +452,7 @@
 
 val read_prop = certify o SimpleSyntax.read_prop;
 
-fun get_def thy = Thm.axiom thy o NameSpace.intern (Theory.axiom_space thy) o Thm.def_name;
+fun get_def thy = Thm.axiom thy o Name_Space.intern (Theory.axiom_space thy) o Thm.def_name;
 
 fun store_thm name th =
   Context.>>> (Context.map_theory_result (PureThy.store_thm (Binding.name name, th)));
--- a/src/Pure/facts.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/facts.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -20,7 +20,8 @@
   val selections: string * thm list -> (ref * thm) list
   type T
   val empty: T
-  val space_of: T -> NameSpace.T
+  val space_of: T -> Name_Space.T
+  val is_concealed: T -> string -> bool
   val intern: T -> xstring -> string
   val extern: T -> string -> xstring
   val lookup: Context.generic -> T -> string -> (bool * thm list) option
@@ -31,9 +32,9 @@
   val props: T -> thm list
   val could_unify: T -> term -> thm list
   val merge: T * T -> T
-  val add_global: NameSpace.naming -> binding * thm list -> T -> string * T
-  val add_local: bool -> NameSpace.naming -> binding * thm list -> T -> string * T
-  val add_dynamic: NameSpace.naming -> binding * (Context.generic -> thm list) -> T -> string * T
+  val add_global: Name_Space.naming -> binding * thm list -> T -> string * T
+  val add_local: bool -> Name_Space.naming -> binding * thm list -> T -> string * T
+  val add_dynamic: Name_Space.naming -> binding * (Context.generic -> thm list) -> T -> string * T
   val del: string -> T -> T
   val hide: bool -> string -> T -> T
 end;
@@ -44,7 +45,7 @@
 (** fact references **)
 
 fun the_single _ [th] : thm = th
-  | the_single name _ = error ("Single theorem expected " ^ quote name);
+  | the_single name _ = error ("Expected singleton fact " ^ quote name);
 
 
 (* datatype interval *)
@@ -122,11 +123,12 @@
 datatype fact = Static of thm list | Dynamic of Context.generic -> thm list;
 
 datatype T = Facts of
- {facts: (fact * serial) NameSpace.table,
+ {facts: fact Name_Space.table,
   props: thm Net.net};
 
 fun make_facts facts props = Facts {facts = facts, props = props};
-val empty = make_facts NameSpace.empty_table Net.empty;
+
+val empty = make_facts (Name_Space.empty_table "fact") Net.empty;
 
 
 (* named facts *)
@@ -136,18 +138,21 @@
 val space_of = #1 o facts_of;
 val table_of = #2 o facts_of;
 
-val intern = NameSpace.intern o space_of;
-val extern = NameSpace.extern o space_of;
+val is_concealed = Name_Space.is_concealed o space_of;
+
+val intern = Name_Space.intern o space_of;
+val extern = Name_Space.extern o space_of;
 
 val defined = Symtab.defined o table_of;
 
 fun lookup context facts name =
   (case Symtab.lookup (table_of facts) name of
     NONE => NONE
-  | SOME (Static ths, _) => SOME (true, ths)
-  | SOME (Dynamic f, _) => SOME (false, f context));
+  | SOME (Static ths) => SOME (true, ths)
+  | SOME (Dynamic f) => SOME (false, f context));
 
-fun fold_static f = Symtab.fold (fn (name, (Static ths, _)) => f (name, ths) | _ => I) o table_of;
+fun fold_static f =
+  Symtab.fold (fn (name, Static ths) => f (name, ths) | _ => I) o table_of;
 
 
 (* content difference *)
@@ -174,61 +179,52 @@
 
 (* merge facts *)
 
-fun err_dup_fact name = error ("Duplicate fact " ^ quote name);
-
-(* FIXME stamp identity! *)
-fun eq_entry ((Static ths1, _), (Static ths2, _)) = is_equal (list_ord Thm.thm_ord (ths1, ths2))
-  | eq_entry ((_, id1: serial), (_, id2)) = id1 = id2;
-
 fun merge (Facts {facts = facts1, props = props1}, Facts {facts = facts2, props = props2}) =
   let
-    val facts' = NameSpace.merge_tables eq_entry (facts1, facts2)
-      handle Symtab.DUP dup => err_dup_fact dup;
+    val facts' = Name_Space.merge_tables (facts1, facts2);
     val props' = Net.merge (is_equal o prop_ord) (props1, props2);
   in make_facts facts' props' end;
 
 
 (* add static entries *)
 
-fun add permissive do_props naming (b, ths) (Facts {facts, props}) =
+local
+
+fun add strict do_props naming (b, ths) (Facts {facts, props}) =
   let
-    val (name, facts') = if Binding.is_empty b then ("", facts)
-      else let
-        val (space, tab) = facts;
-        val (name, space') = NameSpace.declare naming b space;
-        val entry = (name, (Static ths, serial ()));
-        val tab' = (if permissive then Symtab.update else Symtab.update_new) entry tab
-          handle Symtab.DUP dup => err_dup_fact dup;
-      in (name, (space', tab')) end;
+    val (name, facts') =
+      if Binding.is_empty b then ("", facts)
+      else Name_Space.define strict naming (b, Static ths) facts;
     val props' =
-      if do_props then fold (fn th => Net.insert_term (K false) (Thm.full_prop_of th, th)) ths props
+      if do_props
+      then fold (fn th => Net.insert_term (K false) (Thm.full_prop_of th, th)) ths props
       else props;
   in (name, make_facts facts' props') end;
 
-val add_global = add false false;
-val add_local = add true;
+in
+
+val add_global = add true false;
+val add_local = add false;
+
+end;
 
 
 (* add dynamic entries *)
 
-fun add_dynamic naming (b, f) (Facts {facts = (space, tab), props}) =
-  let
-    val (name, space') = NameSpace.declare naming b space;
-    val entry = (name, (Dynamic f, serial ()));
-    val tab' = Symtab.update_new entry tab
-      handle Symtab.DUP dup => err_dup_fact dup;
-  in (name, make_facts (space', tab') props) end;
+fun add_dynamic naming (b, f) (Facts {facts, props}) =
+  let val (name, facts') = Name_Space.define true naming (b, Dynamic f) facts;
+  in (name, make_facts facts' props) end;
 
 
 (* remove entries *)
 
 fun del name (Facts {facts = (space, tab), props}) =
   let
-    val space' = NameSpace.hide true name space handle ERROR _ => space;
+    val space' = Name_Space.hide true name space handle ERROR _ => space;
     val tab' = Symtab.delete_safe name tab;
   in make_facts (space', tab') props end;
 
 fun hide fully name (Facts {facts = (space, tab), props}) =
-  make_facts (NameSpace.hide fully name space, tab) props;
+  make_facts (Name_Space.hide fully name space, tab) props;
 
 end;
--- a/src/Pure/more_thm.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/more_thm.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -83,9 +83,6 @@
   val has_name_hint: thm -> bool
   val get_name_hint: thm -> string
   val put_name_hint: string -> thm -> thm
-  val get_group: thm -> string option
-  val put_group: string -> thm -> thm
-  val group: string -> attribute
   val axiomK: string
   val assumptionK: string
   val definitionK: string
@@ -417,15 +414,6 @@
 fun put_name_hint name = untag_rule Markup.nameN #> tag_rule (Markup.nameN, name);
 
 
-(* theorem groups *)
-
-fun get_group thm = Properties.get (Thm.get_tags thm) Markup.groupN;
-
-fun put_group name = if name = "" then I else Thm.map_tags (Properties.put (Markup.groupN, name));
-
-fun group name = rule_attribute (K (put_group name));
-
-
 (* theorem kinds *)
 
 val axiomK = "axiom";
--- a/src/Pure/name.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/name.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -45,7 +45,7 @@
 
 (* checked binding *)
 
-val of_binding = Long_Name.base_name o NameSpace.full_name NameSpace.default_naming;
+val of_binding = Long_Name.base_name o Name_Space.full_name Name_Space.default_naming;
 
 
 (* encoded bounds *)
--- a/src/Pure/pure_thy.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/pure_thy.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -32,8 +32,6 @@
   val add_thms_dynamic: binding * (Context.generic -> thm list) -> theory -> theory
   val note_thmss: string -> (Thm.binding * (thm list * attribute list) list) list
     -> theory -> (string * thm list) list * theory
-  val note_thmss_grouped: string -> string -> (Thm.binding * (thm list * attribute list) list) list
-    -> theory -> (string * thm list) list * theory
   val add_axioms: ((binding * term) * attribute list) list -> theory -> thm list * theory
   val add_axioms_cmd: ((binding * string) * attribute list) list -> theory -> thm list * theory
   val add_defs: bool -> ((binding * term) * attribute list) list ->
@@ -146,7 +144,7 @@
   else
     let
       val naming = Sign.naming_of thy;
-      val name = NameSpace.full_name naming b;
+      val name = Name_Space.full_name naming b;
       val (thy', thms') =
         register_proofs (apsnd (post_name name) (app_att (thy, pre_name name thms)));
       val thms'' = map (Thm.transfer thy') thms';
@@ -192,9 +190,7 @@
 
 (* note_thmss *)
 
-local
-
-fun gen_note_thmss tag = fold_map (fn ((b, more_atts), ths_atts) => fn thy =>
+fun note_thmss kind = fold_map (fn ((b, more_atts), ths_atts) => fn thy =>
   let
     val pos = Binding.pos_of b;
     val name = Sign.full_name thy b;
@@ -203,16 +199,9 @@
     fun app (x, (ths, atts)) = Library.foldl_map (Thm.theory_attributes atts) (x, ths);
     val (thms, thy') = thy |> enter_thms
       (name_thmss true pos) (name_thms false true pos) (apsnd flat o Library.foldl_map app)
-      (b, map (fn (ths, atts) => (ths, surround tag (atts @ more_atts))) ths_atts);
+      (b, map (fn (ths, atts) => (ths, surround (Thm.kind kind) (atts @ more_atts))) ths_atts);
   in ((name, thms), thy') end);
 
-in
-
-fun note_thmss k = gen_note_thmss (Thm.kind k);
-fun note_thmss_grouped k g = gen_note_thmss (Thm.kind k #> Thm.group g);
-
-end;
-
 
 (* store axioms as theorems *)
 
--- a/src/Pure/sign.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/sign.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -8,11 +8,12 @@
 signature SIGN =
 sig
   val rep_sg: theory ->
-   {naming: NameSpace.naming,
+   {naming: Name_Space.naming,
     syn: Syntax.syntax,
     tsig: Type.tsig,
     consts: Consts.T}
-  val naming_of: theory -> NameSpace.naming
+  val map_naming: (Name_Space.naming -> Name_Space.naming) -> theory -> theory
+  val naming_of: theory -> Name_Space.naming
   val full_name: theory -> binding -> string
   val full_name_path: theory -> string -> binding -> string
   val full_bname: theory -> bstring -> string
@@ -44,9 +45,9 @@
   val const_typargs: theory -> string * typ -> typ list
   val const_instance: theory -> string * typ list -> typ
   val mk_const: theory -> string * typ list -> term
-  val class_space: theory -> NameSpace.T
-  val type_space: theory -> NameSpace.T
-  val const_space: theory -> NameSpace.T
+  val class_space: theory -> Name_Space.T
+  val type_space: theory -> Name_Space.T
+  val const_space: theory -> Name_Space.T
   val intern_class: theory -> xstring -> string
   val extern_class: theory -> string -> xstring
   val intern_type: theory -> xstring -> string
@@ -89,10 +90,10 @@
   val del_modesyntax: Syntax.mode -> (string * string * mixfix) list -> theory -> theory
   val del_modesyntax_i: Syntax.mode -> (string * typ * mixfix) list -> theory -> theory
   val notation: bool -> Syntax.mode -> (term * mixfix) list -> theory -> theory
-  val declare_const: Properties.T -> (binding * typ) * mixfix -> theory -> term * theory
+  val declare_const: (binding * typ) * mixfix -> theory -> term * theory
   val add_consts: (binding * string * mixfix) list -> theory -> theory
   val add_consts_i: (binding * typ * mixfix) list -> theory -> theory
-  val add_abbrev: string -> Properties.T -> binding * term -> theory -> (term * term) * theory
+  val add_abbrev: string -> binding * term -> theory -> (term * term) * theory
   val revert_abbrev: string -> string -> theory -> theory
   val add_const_constraint: string * typ option -> theory -> theory
   val primitive_class: binding * class list -> theory -> theory
@@ -137,7 +138,7 @@
 (** datatype sign **)
 
 datatype sign = Sign of
- {naming: NameSpace.naming,     (*common naming conventions*)
+ {naming: Name_Space.naming,    (*common naming conventions*)
   syn: Syntax.syntax,           (*concrete syntax for terms, types, sorts*)
   tsig: Type.tsig,              (*order-sorted signature of types*)
   consts: Consts.T};            (*polymorphic constants*)
@@ -150,17 +151,17 @@
   type T = sign;
   val copy = I;
   fun extend (Sign {syn, tsig, consts, ...}) =
-    make_sign (NameSpace.default_naming, syn, tsig, consts);
+    make_sign (Name_Space.default_naming, syn, tsig, consts);
 
   val empty =
-    make_sign (NameSpace.default_naming, Syntax.basic_syn, Type.empty_tsig, Consts.empty);
+    make_sign (Name_Space.default_naming, Syntax.basic_syn, Type.empty_tsig, Consts.empty);
 
   fun merge pp (sign1, sign2) =
     let
       val Sign {naming = _, syn = syn1, tsig = tsig1, consts = consts1} = sign1;
       val Sign {naming = _, syn = syn2, tsig = tsig2, consts = consts2} = sign2;
 
-      val naming = NameSpace.default_naming;
+      val naming = Name_Space.default_naming;
       val syn = Syntax.merge_syntaxes syn1 syn2;
       val tsig = Type.merge_tsigs pp (tsig1, tsig2);
       val consts = Consts.merge (consts1, consts2);
@@ -182,10 +183,10 @@
 
 val naming_of = #naming o rep_sg;
 
-val full_name = NameSpace.full_name o naming_of;
-fun full_name_path thy path = NameSpace.full_name (NameSpace.add_path path (naming_of thy));
+val full_name = Name_Space.full_name o naming_of;
+fun full_name_path thy path = Name_Space.full_name (Name_Space.add_path path (naming_of thy));
 
-fun full_bname thy = NameSpace.full_name (naming_of thy) o Binding.name;
+fun full_bname thy = Name_Space.full_name (naming_of thy) o Binding.name;
 fun full_bname_path thy path = full_name_path thy path o Binding.name;
 
 
@@ -240,12 +241,12 @@
 val type_space = #1 o #types o Type.rep_tsig o tsig_of;
 val const_space = Consts.space_of o consts_of;
 
-val intern_class = NameSpace.intern o class_space;
-val extern_class = NameSpace.extern o class_space;
-val intern_type = NameSpace.intern o type_space;
-val extern_type = NameSpace.extern o type_space;
-val intern_const = NameSpace.intern o const_space;
-val extern_const = NameSpace.extern o const_space;
+val intern_class = Name_Space.intern o class_space;
+val extern_class = Name_Space.extern o class_space;
+val intern_type = Name_Space.intern o type_space;
+val extern_type = Name_Space.extern o type_space;
+val intern_const = Name_Space.intern o const_space;
+val extern_const = Name_Space.extern o const_space;
 
 val intern_sort = map o intern_class;
 val extern_sort = map o extern_class;
@@ -433,8 +434,7 @@
   let
     val syn' = Syntax.update_type_gram (map (fn (a, n, mx) => (Name.of_binding a, n, mx)) types) syn;
     val decls = map (fn (a, n, mx) => (Binding.map_name (Syntax.type_name mx) a, n)) types;
-    val tags = [(Markup.theory_nameN, Context.theory_name thy)];
-    val tsig' = fold (Type.add_type naming tags) decls tsig;
+    val tsig' = fold (Type.add_type naming) decls tsig;
   in (naming, syn', tsig', consts) end);
 
 
@@ -443,7 +443,7 @@
 fun add_nonterminals ns thy = thy |> map_sign (fn (naming, syn, tsig, consts) =>
   let
     val syn' = Syntax.update_consts (map Name.of_binding ns) syn;
-    val tsig' = fold (Type.add_nonterminal naming []) ns tsig;
+    val tsig' = fold (Type.add_nonterminal naming) ns tsig;
   in (naming, syn', tsig', consts) end);
 
 
@@ -457,7 +457,7 @@
       val b = Binding.map_name (Syntax.type_name mx) a;
       val abbr = (b, vs, certify_typ_mode Type.mode_syntax thy (parse_typ ctxt rhs))
         handle ERROR msg => cat_error msg ("in type abbreviation " ^ quote (Binding.str_of b));
-      val tsig' = Type.add_abbrev naming [] abbr tsig;
+      val tsig' = Type.add_abbrev naming abbr tsig;
     in (naming, syn', tsig', consts) end);
 
 val add_tyabbrs = fold (gen_add_tyabbr Syntax.parse_typ);
@@ -495,7 +495,7 @@
 
 local
 
-fun gen_add_consts parse_typ authentic tags raw_args thy =
+fun gen_add_consts parse_typ authentic raw_args thy =
   let
     val ctxt = ProofContext.init thy;
     val prepT = Type.no_tvars o Term.no_dummyT o certify_typ thy o parse_typ ctxt;
@@ -510,24 +510,22 @@
         val T' = Logic.varifyT T;
       in ((b, T'), (c_syn, T', mx), Const (c, T)) end;
     val args = map prep raw_args;
-    val tags' = tags |> Properties.put (Markup.theory_nameN, Context.theory_name thy);
   in
     thy
-    |> map_consts (fold (Consts.declare authentic (naming_of thy) tags' o #1) args)
+    |> map_consts (fold (Consts.declare authentic (naming_of thy) o #1) args)
     |> add_syntax_i (map #2 args)
     |> pair (map #3 args)
   end;
 
 in
 
-fun add_consts args = snd o gen_add_consts Syntax.parse_typ false [] args;
-fun add_consts_i args = snd o gen_add_consts (K I) false [] args;
+fun add_consts args = snd o gen_add_consts Syntax.parse_typ false args;
+fun add_consts_i args = snd o gen_add_consts (K I) false args;
 
-fun declare_const tags ((b, T), mx) thy =
+fun declare_const ((b, T), mx) thy =
   let
     val pos = Binding.pos_of b;
-    val tags' = Position.default_properties pos tags;
-    val ([const as Const (c, _)], thy') = gen_add_consts (K I) true tags' [(b, T, mx)] thy;
+    val ([const as Const (c, _)], thy') = gen_add_consts (K I) true [(b, T, mx)] thy;
     val _ = Position.report (Markup.const_decl c) pos;
   in (const, thy') end;
 
@@ -536,14 +534,14 @@
 
 (* abbreviations *)
 
-fun add_abbrev mode tags (b, raw_t) thy =
+fun add_abbrev mode (b, raw_t) thy =
   let
     val pp = Syntax.pp_global thy;
     val prep_tm = no_frees pp o Term.no_dummy_patterns o cert_term_abbrev thy;
     val t = (prep_tm raw_t handle TYPE (msg, _, _) => error msg | TERM (msg, _) => error msg)
       handle ERROR msg => cat_error msg ("in constant abbreviation " ^ quote (Binding.str_of b));
     val (res, consts') = consts_of thy
-      |> Consts.abbreviate pp (tsig_of thy) (naming_of thy) mode tags (b, t);
+      |> Consts.abbreviate pp (tsig_of thy) (naming_of thy) mode (b, t);
   in (res, thy |> map_consts (K consts')) end;
 
 fun revert_abbrev mode c = map_consts (Consts.revert_abbrev mode c);
@@ -612,10 +610,10 @@
 
 (* naming *)
 
-val add_path = map_naming o NameSpace.add_path;
-val root_path = map_naming NameSpace.root_path;
-val parent_path = map_naming NameSpace.parent_path;
-val mandatory_path = map_naming o NameSpace.mandatory_path;
+val add_path = map_naming o Name_Space.add_path;
+val root_path = map_naming Name_Space.root_path;
+val parent_path = map_naming Name_Space.parent_path;
+val mandatory_path = map_naming o Name_Space.mandatory_path;
 
 fun local_path thy = thy |> root_path |> add_path (Context.theory_name thy);
 
--- a/src/Pure/simplifier.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/simplifier.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -143,18 +143,14 @@
 
 (** named simprocs **)
 
-fun err_dup_simproc name = error ("Duplicate simproc: " ^ quote name);
-
-
 (* data *)
 
 structure Simprocs = GenericDataFun
 (
-  type T = simproc NameSpace.table;
-  val empty = NameSpace.empty_table;
+  type T = simproc Name_Space.table;
+  val empty : T = Name_Space.empty_table "simproc";
   val extend = I;
-  fun merge _ simprocs = NameSpace.merge_tables eq_simproc simprocs
-    handle Symtab.DUP dup => err_dup_simproc dup;
+  fun merge _ simprocs = Name_Space.merge_tables simprocs;
 );
 
 
@@ -163,7 +159,7 @@
 fun get_simproc context xname =
   let
     val (space, tab) = Simprocs.get context;
-    val name = NameSpace.intern space xname;
+    val name = Name_Space.intern space xname;
   in
     (case Symtab.lookup tab name of
       SOME proc => proc
@@ -181,9 +177,9 @@
 fun gen_simproc prep {name, lhss, proc, identifier} lthy =
   let
     val b = Binding.name name;
-    val naming = LocalTheory.full_naming lthy;
+    val naming = LocalTheory.naming_of lthy;
     val simproc = make_simproc
-      {name = LocalTheory.full_name lthy b,
+      {name = Name_Space.full_name naming b,
        lhss =
         let
           val lhss' = prep lthy lhss;
@@ -201,9 +197,7 @@
         val b' = Morphism.binding phi b;
         val simproc' = morph_simproc phi simproc;
       in
-        Simprocs.map (fn simprocs =>
-          NameSpace.define naming (b', simproc') simprocs |> snd
-            handle Symtab.DUP dup => err_dup_simproc dup)
+        Simprocs.map (#2 o Name_Space.define true naming (b', simproc'))
         #> map_ss (fn ss => ss addsimprocs [simproc'])
       end)
   end;
--- a/src/Pure/theory.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/theory.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -19,7 +19,7 @@
   val checkpoint: theory -> theory
   val copy: theory -> theory
   val requires: theory -> string -> string -> unit
-  val axiom_space: theory -> NameSpace.T
+  val axiom_space: theory -> Name_Space.T
   val axiom_table: theory -> term Symtab.table
   val axioms_of: theory -> (string * term) list
   val all_axioms_of: theory -> (string * term) list
@@ -35,7 +35,7 @@
   val add_defs: bool -> bool -> (binding * string) list -> theory -> theory
   val add_finals_i: bool -> term list -> theory -> theory
   val add_finals: bool -> string list -> theory -> theory
-  val specify_const: Properties.T -> (binding * typ) * mixfix -> theory -> term * theory
+  val specify_const: (binding * typ) * mixfix -> theory -> term * theory
 end
 
 structure Theory: THEORY =
@@ -80,29 +80,27 @@
   perhaps (perhaps_loop (perhaps_apply (map fst wrappers)));
 
 datatype thy = Thy of
- {axioms: term NameSpace.table,
+ {axioms: term Name_Space.table,
   defs: Defs.T,
   wrappers: wrapper list * wrapper list};
 
 fun make_thy (axioms, defs, wrappers) = Thy {axioms = axioms, defs = defs, wrappers = wrappers};
 
-fun err_dup_axm dup = error ("Duplicate axiom: " ^ quote dup);
-
 structure ThyData = TheoryDataFun
 (
   type T = thy;
-  val empty = make_thy (NameSpace.empty_table, Defs.empty, ([], []));
+  val empty_axioms = Name_Space.empty_table "axiom" : term Name_Space.table;
+  val empty = make_thy (empty_axioms, Defs.empty, ([], []));
   val copy = I;
 
-  fun extend (Thy {axioms = _, defs, wrappers}) =
-    make_thy (NameSpace.empty_table, defs, wrappers);
+  fun extend (Thy {axioms = _, defs, wrappers}) = make_thy (empty_axioms, defs, wrappers);
 
   fun merge pp (thy1, thy2) =
     let
       val Thy {axioms = _, defs = defs1, wrappers = (bgs1, ens1)} = thy1;
       val Thy {axioms = _, defs = defs2, wrappers = (bgs2, ens2)} = thy2;
 
-      val axioms' = NameSpace.empty_table;
+      val axioms' = empty_axioms;
       val defs' = Defs.merge pp (defs1, defs2);
       val bgs' = Library.merge (eq_snd op =) (bgs1, bgs2);
       val ens' = Library.merge (eq_snd op =) (ens1, ens2);
@@ -143,7 +141,12 @@
   let
     val thy = Context.begin_thy Syntax.pp_global name imports;
     val wrappers = begin_wrappers thy;
-  in thy |> Sign.local_path |> apply_wrappers wrappers end;
+  in
+    thy
+    |> Sign.local_path
+    |> Sign.map_naming (Name_Space.set_theory_name name)
+    |> apply_wrappers wrappers
+  end;
 
 fun end_theory thy =
   thy |> apply_wrappers (end_wrappers thy) |> Context.finish_thy;
@@ -166,7 +169,7 @@
 
 fun read_axm thy (b, str) =
   cert_axm thy (b, Syntax.read_prop_global thy str) handle ERROR msg =>
-    cat_error msg ("The error(s) above occurred in axiom: " ^ quote (Binding.str_of b));
+    cat_error msg ("The error(s) above occurred in axiom " ^ quote (Binding.str_of b));
 
 
 (* add_axioms(_i) *)
@@ -176,8 +179,7 @@
 fun gen_add_axioms prep_axm raw_axms thy = thy |> map_axioms (fn axioms =>
   let
     val axms = map (apsnd Logic.varify o prep_axm thy) raw_axms;
-    val axioms' = fold (snd oo NameSpace.define (Sign.naming_of thy)) axms axioms
-      handle Symtab.DUP dup => err_dup_axm dup;
+    val axioms' = fold (#2 oo Name_Space.define true (Sign.naming_of thy)) axms axioms;
   in axioms' end);
 
 in
@@ -217,8 +219,8 @@
     val name = if a = "" then (#1 lhs ^ " axiom") else a;
   in thy |> map_defs (dependencies thy false false name lhs rhs) end;
 
-fun specify_const tags decl thy =
-  let val (t as Const const, thy') = Sign.declare_const tags decl thy
+fun specify_const decl thy =
+  let val (t as Const const, thy') = Sign.declare_const decl thy
   in (t, add_deps "" const [] thy') end;
 
 
--- a/src/Pure/thm.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/thm.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -1724,25 +1724,21 @@
 
 (* authentic derivation names *)
 
-fun err_dup_ora dup = error ("Duplicate oracle: " ^ quote dup);
-
 structure Oracles = TheoryDataFun
 (
-  type T = serial NameSpace.table;
-  val empty = NameSpace.empty_table;
+  type T = unit Name_Space.table;
+  val empty : T = Name_Space.empty_table "oracle";
   val copy = I;
   val extend = I;
-  fun merge _ oracles : T = NameSpace.merge_tables (op =) oracles
-    handle Symtab.DUP dup => err_dup_ora dup;
+  fun merge _ oracles : T = Name_Space.merge_tables oracles;
 );
 
-val extern_oracles = map #1 o NameSpace.extern_table o Oracles.get;
+val extern_oracles = map #1 o Name_Space.extern_table o Oracles.get;
 
 fun add_oracle (b, oracle) thy =
   let
     val naming = Sign.naming_of thy;
-    val (name, tab') = NameSpace.define naming (b, serial ()) (Oracles.get thy)
-      handle Symtab.DUP _ => err_dup_ora (Binding.str_of b);
+    val (name, tab') = Name_Space.define true naming (b, ()) (Oracles.get thy);
     val thy' = Oracles.put tab' thy;
   in ((name, invoke_oracle (Theory.check_thy thy') name oracle), thy') end;
 
--- a/src/Pure/type.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Pure/type.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -14,9 +14,9 @@
     Nonterminal
   type tsig
   val rep_tsig: tsig ->
-   {classes: NameSpace.T * Sorts.algebra,
+   {classes: Name_Space.T * Sorts.algebra,
     default: sort,
-    types: ((decl * Properties.T) * serial) NameSpace.table,
+    types: decl Name_Space.table,
     log_types: string list}
   val empty_tsig: tsig
   val defaultS: tsig -> sort
@@ -39,7 +39,6 @@
   val cert_typ: tsig -> typ -> typ
   val arity_number: tsig -> string -> int
   val arity_sorts: Pretty.pp -> tsig -> string -> sort -> sort list
-  val the_tags: tsig -> string -> Properties.T
 
   (*special treatment of type vars*)
   val strip_sorts: typ -> typ
@@ -70,12 +69,12 @@
   val eq_type: tyenv -> typ * typ -> bool
 
   (*extend and merge type signatures*)
-  val add_class: Pretty.pp -> NameSpace.naming -> binding * class list -> tsig -> tsig
+  val add_class: Pretty.pp -> Name_Space.naming -> binding * class list -> tsig -> tsig
   val hide_class: bool -> string -> tsig -> tsig
   val set_defsort: sort -> tsig -> tsig
-  val add_type: NameSpace.naming -> Properties.T -> binding * int -> tsig -> tsig
-  val add_abbrev: NameSpace.naming -> Properties.T -> binding * string list * typ -> tsig -> tsig
-  val add_nonterminal: NameSpace.naming -> Properties.T -> binding -> tsig -> tsig
+  val add_type: Name_Space.naming -> binding * int -> tsig -> tsig
+  val add_abbrev: Name_Space.naming -> binding * string list * typ -> tsig -> tsig
+  val add_nonterminal: Name_Space.naming -> binding -> tsig -> tsig
   val hide_type: bool -> string -> tsig -> tsig
   val add_arity: Pretty.pp -> arity -> tsig -> tsig
   val add_classrel: Pretty.pp -> class * class -> tsig -> tsig
@@ -94,18 +93,14 @@
   Abbreviation of string list * typ * bool |
   Nonterminal;
 
-fun str_of_decl (LogicalType _) = "logical type constructor"
-  | str_of_decl (Abbreviation _) = "type abbreviation"
-  | str_of_decl Nonterminal = "syntactic type";
-
 
 (* type tsig *)
 
 datatype tsig =
   TSig of {
-    classes: NameSpace.T * Sorts.algebra,   (*order-sorted algebra of type classes*)
+    classes: Name_Space.T * Sorts.algebra,  (*order-sorted algebra of type classes*)
     default: sort,                          (*default sort on input*)
-    types: ((decl * Properties.T) * serial) NameSpace.table, (*declared types*)
+    types: decl Name_Space.table,           (*declared types*)
     log_types: string list};                (*logical types sorted by number of arguments*)
 
 fun rep_tsig (TSig comps) = comps;
@@ -113,18 +108,18 @@
 fun make_tsig (classes, default, types, log_types) =
   TSig {classes = classes, default = default, types = types, log_types = log_types};
 
-fun build_tsig ((space, classes), default, types) =
+fun build_tsig (classes, default, types) =
   let
     val log_types =
-      Symtab.fold (fn (c, ((LogicalType n, _), _)) => cons (c, n) | _ => I) (snd types) []
-      |> Library.sort (Library.int_ord o pairself snd) |> map fst;
-  in make_tsig ((space, classes), default, types, log_types) end;
+      Symtab.fold (fn (c, LogicalType n) => cons (c, n) | _ => I) (snd types) []
+      |> Library.sort (int_ord o pairself snd) |> map fst;
+  in make_tsig (classes, default, types, log_types) end;
 
 fun map_tsig f (TSig {classes, default, types, log_types = _}) =
   build_tsig (f (classes, default, types));
 
 val empty_tsig =
-  build_tsig ((NameSpace.empty, Sorts.empty_algebra), [], NameSpace.empty_table);
+  build_tsig ((Name_Space.empty "class", Sorts.empty_algebra), [], Name_Space.empty_table "type");
 
 
 (* classes and sorts *)
@@ -167,12 +162,7 @@
 
 fun undecl_type c = "Undeclared type constructor: " ^ quote c;
 
-fun lookup_type (TSig {types, ...}) = Option.map fst o Symtab.lookup (snd types);
-
-fun the_tags tsig c =
-  (case lookup_type tsig c of
-    SOME (_, tags) => tags
-  | NONE => error (undecl_type c));
+fun lookup_type (TSig {types = (_, types), ...}) = Symtab.lookup types;
 
 
 (* certified types *)
@@ -201,13 +191,13 @@
             fun nargs n = if length Ts <> n then err (bad_nargs c) else ();
           in
             (case lookup_type tsig c of
-              SOME (LogicalType n, _) => (nargs n; Type (c, Ts'))
-            | SOME (Abbreviation (vs, U, syn), _) =>
+              SOME (LogicalType n) => (nargs n; Type (c, Ts'))
+            | SOME (Abbreviation (vs, U, syn)) =>
                (nargs (length vs);
                 if syn then check_logical c else ();
                 if normalize then inst_typ (vs ~~ Ts') U
                 else Type (c, Ts'))
-            | SOME (Nonterminal, _) => (nargs 0; check_logical c; T)
+            | SOME Nonterminal => (nargs 0; check_logical c; T)
             | NONE => err (undecl_type c))
           end
       | cert (TFree (x, S)) = TFree (x, cert_sort tsig S)
@@ -228,7 +218,7 @@
 
 fun arity_number tsig a =
   (case lookup_type tsig a of
-    SOME (LogicalType n, _) => n
+    SOME (LogicalType n) => n
   | _ => error (undecl_type a));
 
 fun arity_sorts _ tsig a [] = replicate (arity_number tsig a) []
@@ -515,12 +505,12 @@
     let
       val cs' = map (cert_class tsig) cs
         handle TYPE (msg, _, _) => error msg;
-      val (c', space') = space |> NameSpace.declare naming c;
+      val (c', space') = space |> Name_Space.declare true naming c;
       val classes' = classes |> Sorts.add_class pp (c', cs');
     in ((space', classes'), default, types) end);
 
 fun hide_class fully c = map_tsig (fn ((space, classes), default, types) =>
-  ((NameSpace.hide fully c space, classes), default, types));
+  ((Name_Space.hide fully c space, classes), default, types));
 
 
 (* arities *)
@@ -529,8 +519,8 @@
   let
     val _ =
       (case lookup_type tsig t of
-        SOME (LogicalType n, _) => if length Ss <> n then error (bad_nargs t) else ()
-      | SOME (decl, _) => error ("Illegal " ^ str_of_decl decl ^ ": " ^ quote t)
+        SOME (LogicalType n) => if length Ss <> n then error (bad_nargs t) else ()
+      | SOME _ => error ("Logical type constructor expected: " ^ quote t)
       | NONE => error (undecl_type t));
     val (Ss', S') = (map (cert_sort tsig) Ss, cert_sort tsig S)
       handle TYPE (msg, _, _) => error msg;
@@ -559,68 +549,50 @@
 
 local
 
-fun err_in_decls c decl decl' =
-  let val s = str_of_decl decl and s' = str_of_decl decl' in
-    if s = s' then error ("Duplicate declaration of " ^ s ^ ": " ^ quote c)
-    else error ("Conflict of " ^ s ^ " with " ^ s' ^ ": " ^ quote c)
-  end;
-
-fun new_decl naming tags (c, decl) (space, types) =
-  let
-    val tags' = tags |> Position.default_properties (Position.thread_data ());
-    val (c', space') = NameSpace.declare naming c space;
-    val types' =
-      (case Symtab.lookup types c' of
-        SOME ((decl', _), _) => err_in_decls c' decl decl'
-      | NONE => Symtab.update (c', ((decl, tags'), serial ())) types);
-  in (space', types') end;
-
-fun the_decl (_, types) = fst o fst o the o Symtab.lookup types;
+fun new_decl naming (c, decl) types =
+  #2 (Name_Space.define true naming (c, decl) types);
 
 fun map_types f = map_tsig (fn (classes, default, types) =>
   let
     val (space', tab') = f types;
-    val _ = NameSpace.intern space' "dummy" = "dummy" orelse
+    val _ = Name_Space.intern space' "dummy" = "dummy" orelse
       error "Illegal declaration of dummy type";
   in (classes, default, (space', tab')) end);
 
 fun syntactic types (Type (c, Ts)) =
-      (case Symtab.lookup types c of SOME ((Nonterminal, _), _) => true | _ => false)
+      (case Symtab.lookup types c of SOME Nonterminal => true | _ => false)
         orelse exists (syntactic types) Ts
   | syntactic _ _ = false;
 
 in
 
-fun add_type naming tags (c, n) =
-  if n < 0 then error ("Bad type constructor declaration: " ^ quote (Binding.str_of c))
-  else map_types (new_decl naming tags (c, LogicalType n));
+fun add_type naming (c, n) =
+  if n < 0 then error ("Bad type constructor declaration " ^ quote (Binding.str_of c))
+  else map_types (new_decl naming (c, LogicalType n));
 
-fun add_abbrev naming tags (a, vs, rhs) tsig = tsig |> map_types (fn types =>
+fun add_abbrev naming (a, vs, rhs) tsig = tsig |> map_types (fn types =>
   let
     fun err msg =
-      cat_error msg ("The error(s) above occurred in type abbreviation: " ^ quote (Binding.str_of a));
+      cat_error msg ("The error(s) above occurred in type abbreviation " ^
+        quote (Binding.str_of a));
     val rhs' = strip_sorts (no_tvars (cert_typ_mode mode_syntax tsig rhs))
       handle TYPE (msg, _, _) => err msg;
-  in
-    (case duplicates (op =) vs of
-      [] => []
-    | dups => err ("Duplicate variables on lhs: " ^ commas_quote dups));
-    (case subtract (op =) vs (map #1 (Term.add_tfreesT rhs' [])) of
-      [] => []
-    | extras => err ("Extra variables on rhs: " ^ commas_quote extras));
-    types |> new_decl naming tags (a, Abbreviation (vs, rhs', syntactic (#2 types) rhs'))
-  end);
+    val _ =
+      (case duplicates (op =) vs of
+        [] => []
+      | dups => err ("Duplicate variables on lhs: " ^ commas_quote dups));
+    val _ =
+      (case subtract (op =) vs (map #1 (Term.add_tfreesT rhs' [])) of
+        [] => []
+      | extras => err ("Extra variables on rhs: " ^ commas_quote extras));
+  in types |> new_decl naming (a, Abbreviation (vs, rhs', syntactic (#2 types) rhs')) end);
 
-fun add_nonterminal naming tags = map_types o new_decl naming tags o rpair Nonterminal;
-
-fun merge_types (types1, types2) =
-  NameSpace.merge_tables (Library.eq_snd (op = : serial * serial -> bool)) (types1, types2)
-    handle Symtab.DUP d => err_in_decls d (the_decl types1 d) (the_decl types2 d);
+fun add_nonterminal naming = map_types o new_decl naming o rpair Nonterminal;
 
 end;
 
 fun hide_type fully c = map_tsig (fn (classes, default, (space, types)) =>
-  (classes, default, (NameSpace.hide fully c space, types)));
+  (classes, default, (Name_Space.hide fully c space, types)));
 
 
 (* merge type signatures *)
@@ -632,10 +604,10 @@
     val (TSig {classes = (space2, classes2), default = default2, types = types2,
       log_types = _}) = tsig2;
 
-    val space' = NameSpace.merge (space1, space2);
+    val space' = Name_Space.merge (space1, space2);
     val classes' = Sorts.merge_algebra pp (classes1, classes2);
     val default' = Sorts.inter_sort classes' (default1, default2);
-    val types' = merge_types (types1, types2);
+    val types' = Name_Space.merge_tables (types1, types2);
   in build_tsig ((space', classes'), default', types') end;
 
 end;
--- a/src/Tools/Code/code_thingol.ML	Fri Oct 23 13:23:18 2009 +0200
+++ b/src/Tools/Code/code_thingol.ML	Mon Oct 26 09:03:57 2009 +0100
@@ -252,19 +252,15 @@
 (* policies *)
 
 local
-  fun thyname_of thy f x = the (AList.lookup (op =) (f x) Markup.theory_nameN);
-  fun thyname_of_class thy =
-    thyname_of thy (ProofContext.query_class (ProofContext.init thy));
-  fun thyname_of_tyco thy =
-    thyname_of thy (Type.the_tags (Sign.tsig_of thy));
-  fun thyname_of_instance thy inst = case AxClass.arity_property thy inst Markup.theory_nameN
-   of [] => error ("no such instance: " ^ quote (snd inst ^ " :: " ^ fst inst))
+  fun thyname_of_class thy = #theory_name o Name_Space.the_entry (Sign.class_space thy);
+  fun thyname_of_instance thy inst = case AxClass.thynames_of_arity thy inst
+   of [] => error ("No such instance: " ^ quote (snd inst ^ " :: " ^ fst inst))
     | thyname :: _ => thyname;
   fun thyname_of_const thy c = case AxClass.class_of_param thy c
    of SOME class => thyname_of_class thy class
     | NONE => (case Code.get_datatype_of_constr thy c
-       of SOME dtco => thyname_of_tyco thy dtco
-        | NONE => thyname_of thy (Consts.the_tags (Sign.consts_of thy)) c);
+       of SOME dtco => Codegen.thyname_of_type thy dtco
+        | NONE => Codegen.thyname_of_const thy c);
   fun purify_base "op &" = "and"
     | purify_base "op |" = "or"
     | purify_base "op -->" = "implies"
@@ -282,10 +278,11 @@
 
 fun namify_class thy = namify thy Long_Name.base_name thyname_of_class;
 fun namify_classrel thy = namify thy (fn (class1, class2) => 
-  Long_Name.base_name class2 ^ "_" ^ Long_Name.base_name class1) (fn thy => thyname_of_class thy o fst);
+    Long_Name.base_name class2 ^ "_" ^ Long_Name.base_name class1)
+  (fn thy => thyname_of_class thy o fst);
   (*order fits nicely with composed projections*)
 fun namify_tyco thy "fun" = "Pure.fun"
-  | namify_tyco thy tyco = namify thy Long_Name.base_name thyname_of_tyco tyco;
+  | namify_tyco thy tyco = namify thy Long_Name.base_name Codegen.thyname_of_type tyco;
 fun namify_instance thy = namify thy (fn (class, tyco) => 
   Long_Name.base_name class ^ "_" ^ Long_Name.base_name tyco) thyname_of_instance;
 fun namify_const thy = namify thy Long_Name.base_name thyname_of_const;