merged
authornipkow
Thu, 24 Sep 2009 17:26:05 +0200
changeset 32677 8854e892cf3e
parent 32675 5fe601aff9be (diff)
parent 32676 b1c85a117dec (current diff)
child 32682 304a47739407
merged
--- a/Admin/E/eproof	Thu Sep 24 17:25:42 2009 +0200
+++ b/Admin/E/eproof	Thu Sep 24 17:26:05 2009 +0200
@@ -11,6 +11,7 @@
 
 use File::Basename qw/ dirname /;
 use File::Temp qw/ tempfile /;
+use English;
 
 
 # E executables
@@ -44,7 +45,7 @@
 # run E, redirecting output into a temporary file
 
 my ($fh, $filename) = tempfile(UNLINK => 1);
-my $r = system "$eprover_cmd > $filename";
+my $r = system "$eprover_cmd > '$filename'";
 exit ($r >> 8) if $r != 0;
 
 
@@ -55,7 +56,7 @@
   # Note: Like the original eproof, we only look at the last 60 lines.
 
 if ($content =~ m/Total time\s*:\s*([0-9]+\.[0-9]+)/) {
-  $timelimit = $timelimit - $1 - 1;
+  $timelimit = int($timelimit - $1 - 1);
 
   if ($content =~ m/No proof found!/) {
     print "# Problem is satisfiable (or invalid), " .
@@ -85,7 +86,7 @@
   print if (m/# SZS status/ or m/"# Failure"/);
 }
 $r = system ("exec bash -c \"ulimit -S -t $timelimit; " .
-  "'$epclextract' $format -f --competition-framing $filename\"");
+  "'$epclextract' $format -f --competition-framing '$filename'\"");
   # Note: Setting the user time is not supported on Cygwin, i.e., ulimit fails
   # and prints and error message. How could we then limit the execution time?
 exit ($r >> 8);
--- a/Admin/isatest/isatest-makedist	Thu Sep 24 17:25:42 2009 +0200
+++ b/Admin/isatest/isatest-makedist	Thu Sep 24 17:26:05 2009 +0200
@@ -102,9 +102,9 @@
 #sleep 15
 $SSH macbroy21 "$MAKEALL $HOME/settings/at64-poly-5.1-para"
 sleep 15
-$SSH macbroy23 -l HOL images "$MAKEALL $HOME/settings/at-sml-dev-e"
-sleep 15
-$SSH atbroy101 "$MAKEALL $HOME/settings/at64-poly"
+#$SSH macbroy23 -l HOL images "$MAKEALL $HOME/settings/at-sml-dev-e"
+#sleep 15
+$SSH atbroy99 "$MAKEALL $HOME/settings/at64-poly"
 sleep 15
 $SSH macbroy2 "$MAKEALL $HOME/settings/mac-poly-M4; $MAKEALL $HOME/settings/mac-poly-M8; $MAKEALL $HOME/settings/mac-poly64-M4; $MAKEALL $HOME/settings/mac-poly64-M8"
 sleep 15
--- a/NEWS	Thu Sep 24 17:25:42 2009 +0200
+++ b/NEWS	Thu Sep 24 17:26:05 2009 +0200
@@ -102,6 +102,10 @@
 
   INCOMPATIBILITY.
 
+* Rules inf_absorb1, inf_absorb2, sup_absorb1, sup_absorb2 are no
+simp rules by default any longer.  The same applies to
+min_max.inf_absorb1 etc.!  INCOMPATIBILITY.
+
 * Power operations on relations and functions are now one dedicate
 constant "compow" with infix syntax "^^".  Power operations on
 multiplicative monoids retains syntax "^" and is now defined generic
--- a/lib/Tools/usedir	Thu Sep 24 17:25:42 2009 +0200
+++ b/lib/Tools/usedir	Thu Sep 24 17:26:05 2009 +0200
@@ -262,7 +262,7 @@
 else
   { echo "$ITEM FAILED";
     echo "(see also $LOG)";
-    echo; tail "$LOG"; echo; } >&2
+    echo; tail -n 20 "$LOG"; echo; } >&2
 fi
 
 exit "$RC"
--- a/src/HOL/Code_Eval.thy	Thu Sep 24 17:25:42 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,271 +0,0 @@
-(*  Title:      HOL/Code_Eval.thy
-    Author:     Florian Haftmann, TU Muenchen
-*)
-
-header {* Term evaluation using the generic code generator *}
-
-theory Code_Eval
-imports Plain Typerep Code_Numeral
-begin
-
-subsection {* Term representation *}
-
-subsubsection {* Terms and class @{text term_of} *}
-
-datatype "term" = dummy_term
-
-definition Const :: "String.literal \<Rightarrow> typerep \<Rightarrow> term" where
-  "Const _ _ = dummy_term"
-
-definition App :: "term \<Rightarrow> term \<Rightarrow> term" where
-  "App _ _ = dummy_term"
-
-code_datatype Const App
-
-class term_of = typerep +
-  fixes term_of :: "'a \<Rightarrow> term"
-
-lemma term_of_anything: "term_of x \<equiv> t"
-  by (rule eq_reflection) (cases "term_of x", cases t, simp)
-
-definition valapp :: "('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)
-  \<Rightarrow> 'a \<times> (unit \<Rightarrow> term) \<Rightarrow> 'b \<times> (unit \<Rightarrow> term)" where
-  "valapp f x = (fst f (fst x), \<lambda>u. App (snd f ()) (snd x ()))"
-
-lemma valapp_code [code, code_unfold]:
-  "valapp (f, tf) (x, tx) = (f x, \<lambda>u. App (tf ()) (tx ()))"
-  by (simp only: valapp_def fst_conv snd_conv)
-
-
-subsubsection {* @{text term_of} instances *}
-
-instantiation "fun" :: (typerep, typerep) term_of
-begin
-
-definition
-  "term_of (f \<Colon> 'a \<Rightarrow> 'b) = Const (STR ''dummy_pattern'') (Typerep.Typerep (STR ''fun'')
-     [Typerep.typerep TYPE('a), Typerep.typerep TYPE('b)])"
-
-instance ..
-
-end
-
-setup {*
-let
-  fun add_term_of tyco raw_vs thy =
-    let
-      val vs = map (fn (v, _) => (v, @{sort typerep})) raw_vs;
-      val ty = Type (tyco, map TFree vs);
-      val lhs = Const (@{const_name term_of}, ty --> @{typ term})
-        $ Free ("x", ty);
-      val rhs = @{term "undefined \<Colon> term"};
-      val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
-      fun triv_name_of t = (fst o dest_Free o fst o strip_comb o fst
-        o HOLogic.dest_eq o HOLogic.dest_Trueprop) t ^ "_triv";
-    in
-      thy
-      |> TheoryTarget.instantiation ([tyco], vs, @{sort term_of})
-      |> `(fn lthy => Syntax.check_term lthy eq)
-      |-> (fn eq => Specification.definition (NONE, ((Binding.name (triv_name_of eq), []), eq)))
-      |> snd
-      |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
-    end;
-  fun ensure_term_of (tyco, (raw_vs, _)) thy =
-    let
-      val need_inst = not (can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort term_of})
-        andalso can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort typerep};
-    in if need_inst then add_term_of tyco raw_vs thy else thy end;
-in
-  Code.type_interpretation ensure_term_of
-end
-*}
-
-setup {*
-let
-  fun mk_term_of_eq thy ty vs tyco (c, tys) =
-    let
-      val t = list_comb (Const (c, tys ---> ty),
-        map Free (Name.names Name.context "a" tys));
-      val (arg, rhs) = pairself (Thm.cterm_of thy o map_types Logic.unvarifyT o Logic.varify)
-        (t, (map_aterms (fn t as Free (v, ty) => HOLogic.mk_term_of ty t | t => t) o HOLogic.reflect_term) t)
-      val cty = Thm.ctyp_of thy ty;
-    in
-      @{thm term_of_anything}
-      |> Drule.instantiate' [SOME cty] [SOME arg, SOME rhs]
-      |> Thm.varifyT
-    end;
-  fun add_term_of_code tyco raw_vs raw_cs thy =
-    let
-      val algebra = Sign.classes_of thy;
-      val vs = map (fn (v, sort) =>
-        (v, curry (Sorts.inter_sort algebra) @{sort typerep} sort)) raw_vs;
-      val ty = Type (tyco, map TFree vs);
-      val cs = (map o apsnd o map o map_atyps)
-        (fn TFree (v, _) => TFree (v, (the o AList.lookup (op =) vs) v)) raw_cs;
-      val const = AxClass.param_of_inst thy (@{const_name term_of}, tyco);
-      val eqs = map (mk_term_of_eq thy ty vs tyco) cs;
-   in
-      thy
-      |> Code.del_eqns const
-      |> fold Code.add_eqn eqs
-    end;
-  fun ensure_term_of_code (tyco, (raw_vs, cs)) thy =
-    let
-      val has_inst = can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort term_of};
-    in if has_inst then add_term_of_code tyco raw_vs cs thy else thy end;
-in
-  Code.type_interpretation ensure_term_of_code
-end
-*}
-
-
-subsubsection {* Code generator setup *}
-
-lemmas [code del] = term.recs term.cases term.size
-lemma [code, code del]: "eq_class.eq (t1\<Colon>term) t2 \<longleftrightarrow> eq_class.eq t1 t2" ..
-
-lemma [code, code del]: "(term_of \<Colon> typerep \<Rightarrow> term) = term_of" ..
-lemma [code, code del]: "(term_of \<Colon> term \<Rightarrow> term) = term_of" ..
-lemma [code, code del]: "(term_of \<Colon> String.literal \<Rightarrow> term) = term_of" ..
-lemma [code, code del]:
-  "(Code_Eval.term_of \<Colon> 'a::{type, term_of} Predicate.pred \<Rightarrow> Code_Eval.term) = Code_Eval.term_of" ..
-lemma [code, code del]:
-  "(Code_Eval.term_of \<Colon> 'a::{type, term_of} Predicate.seq \<Rightarrow> Code_Eval.term) = Code_Eval.term_of" ..
-
-lemma term_of_char [unfolded typerep_fun_def typerep_char_def typerep_nibble_def, code]: "Code_Eval.term_of c =
-    (let (n, m) = nibble_pair_of_char c
-  in Code_Eval.App (Code_Eval.App (Code_Eval.Const (STR ''String.char.Char'') (TYPEREP(nibble \<Rightarrow> nibble \<Rightarrow> char)))
-    (Code_Eval.term_of n)) (Code_Eval.term_of m))"
-  by (subst term_of_anything) rule 
-
-code_type "term"
-  (Eval "Term.term")
-
-code_const Const and App
-  (Eval "Term.Const/ ((_), (_))" and "Term.$/ ((_), (_))")
-
-code_const "term_of \<Colon> String.literal \<Rightarrow> term"
-  (Eval "HOLogic.mk'_message'_string")
-
-code_reserved Eval HOLogic
-
-
-subsubsection {* Syntax *}
-
-definition termify :: "'a \<Rightarrow> term" where
-  [code del]: "termify x = dummy_term"
-
-abbreviation valtermify :: "'a \<Rightarrow> 'a \<times> (unit \<Rightarrow> term)" where
-  "valtermify x \<equiv> (x, \<lambda>u. termify x)"
-
-setup {*
-let
-  fun map_default f xs =
-    let val ys = map f xs
-    in if exists is_some ys
-      then SOME (map2 the_default xs ys)
-      else NONE
-    end;
-  fun subst_termify_app (Const (@{const_name termify}, T), [t]) =
-        if not (Term.has_abs t)
-        then if fold_aterms (fn Const _ => I | _ => K false) t true
-          then SOME (HOLogic.reflect_term t)
-          else error "Cannot termify expression containing variables"
-        else error "Cannot termify expression containing abstraction"
-    | subst_termify_app (t, ts) = case map_default subst_termify ts
-       of SOME ts' => SOME (list_comb (t, ts'))
-        | NONE => NONE
-  and subst_termify (Abs (v, T, t)) = (case subst_termify t
-       of SOME t' => SOME (Abs (v, T, t'))
-        | NONE => NONE)
-    | subst_termify t = subst_termify_app (strip_comb t) 
-  fun check_termify ts ctxt = map_default subst_termify ts
-    |> Option.map (rpair ctxt)
-in
-  Context.theory_map (Syntax.add_term_check 0 "termify" check_termify)
-end;
-*}
-
-locale term_syntax
-begin
-
-notation App (infixl "<\<cdot>>" 70)
-  and valapp (infixl "{\<cdot>}" 70)
-
-end
-
-interpretation term_syntax .
-
-no_notation App (infixl "<\<cdot>>" 70)
-  and valapp (infixl "{\<cdot>}" 70)
-
-
-subsection {* Numeric types *}
-
-definition term_of_num :: "'a\<Colon>{semiring_div} \<Rightarrow> 'a\<Colon>{semiring_div} \<Rightarrow> term" where
-  "term_of_num two = (\<lambda>_. dummy_term)"
-
-lemma (in term_syntax) term_of_num_code [code]:
-  "term_of_num two k = (if k = 0 then termify Int.Pls
-    else (if k mod two = 0
-      then termify Int.Bit0 <\<cdot>> term_of_num two (k div two)
-      else termify Int.Bit1 <\<cdot>> term_of_num two (k div two)))"
-  by (auto simp add: term_of_anything Const_def App_def term_of_num_def Let_def)
-
-lemma (in term_syntax) term_of_nat_code [code]:
-  "term_of (n::nat) = termify (number_of :: int \<Rightarrow> nat) <\<cdot>> term_of_num (2::nat) n"
-  by (simp only: term_of_anything)
-
-lemma (in term_syntax) term_of_int_code [code]:
-  "term_of (k::int) = (if k = 0 then termify (0 :: int)
-    else if k > 0 then termify (number_of :: int \<Rightarrow> int) <\<cdot>> term_of_num (2::int) k
-      else termify (uminus :: int \<Rightarrow> int) <\<cdot>> (termify (number_of :: int \<Rightarrow> int) <\<cdot>> term_of_num (2::int) (- k)))"
-  by (simp only: term_of_anything)
-
-lemma (in term_syntax) term_of_code_numeral_code [code]:
-  "term_of (k::code_numeral) = termify (number_of :: int \<Rightarrow> code_numeral) <\<cdot>> term_of_num (2::code_numeral) k"
-  by (simp only: term_of_anything)
-
-subsection {* Obfuscate *}
-
-print_translation {*
-let
-  val term = Const ("<TERM>", dummyT);
-  fun tr1' [_, _] = term;
-  fun tr2' [] = term;
-in
-  [(@{const_syntax Const}, tr1'),
-    (@{const_syntax App}, tr1'),
-    (@{const_syntax dummy_term}, tr2')]
-end
-*}
-
-hide const dummy_term App valapp
-hide (open) const Const termify valtermify term_of term_of_num
-
-
-subsection {* Evaluation setup *}
-
-ML {*
-signature EVAL =
-sig
-  val eval_ref: (unit -> term) option ref
-  val eval_term: theory -> term -> term
-end;
-
-structure Eval : EVAL =
-struct
-
-val eval_ref = ref (NONE : (unit -> term) option);
-
-fun eval_term thy t =
-  Code_ML.eval NONE ("Eval.eval_ref", eval_ref) I thy (HOLogic.mk_term_of (fastype_of t) t) [];
-
-end;
-*}
-
-setup {*
-  Value.add_evaluator ("code", Eval.eval_term o ProofContext.theory_of)
-*}
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Code_Evaluation.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -0,0 +1,271 @@
+(*  Title:      HOL/Code_Evaluation.thy
+    Author:     Florian Haftmann, TU Muenchen
+*)
+
+header {* Term evaluation using the generic code generator *}
+
+theory Code_Evaluation
+imports Plain Typerep Code_Numeral
+begin
+
+subsection {* Term representation *}
+
+subsubsection {* Terms and class @{text term_of} *}
+
+datatype "term" = dummy_term
+
+definition Const :: "String.literal \<Rightarrow> typerep \<Rightarrow> term" where
+  "Const _ _ = dummy_term"
+
+definition App :: "term \<Rightarrow> term \<Rightarrow> term" where
+  "App _ _ = dummy_term"
+
+code_datatype Const App
+
+class term_of = typerep +
+  fixes term_of :: "'a \<Rightarrow> term"
+
+lemma term_of_anything: "term_of x \<equiv> t"
+  by (rule eq_reflection) (cases "term_of x", cases t, simp)
+
+definition valapp :: "('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)
+  \<Rightarrow> 'a \<times> (unit \<Rightarrow> term) \<Rightarrow> 'b \<times> (unit \<Rightarrow> term)" where
+  "valapp f x = (fst f (fst x), \<lambda>u. App (snd f ()) (snd x ()))"
+
+lemma valapp_code [code, code_unfold]:
+  "valapp (f, tf) (x, tx) = (f x, \<lambda>u. App (tf ()) (tx ()))"
+  by (simp only: valapp_def fst_conv snd_conv)
+
+
+subsubsection {* @{text term_of} instances *}
+
+instantiation "fun" :: (typerep, typerep) term_of
+begin
+
+definition
+  "term_of (f \<Colon> 'a \<Rightarrow> 'b) = Const (STR ''dummy_pattern'') (Typerep.Typerep (STR ''fun'')
+     [Typerep.typerep TYPE('a), Typerep.typerep TYPE('b)])"
+
+instance ..
+
+end
+
+setup {*
+let
+  fun add_term_of tyco raw_vs thy =
+    let
+      val vs = map (fn (v, _) => (v, @{sort typerep})) raw_vs;
+      val ty = Type (tyco, map TFree vs);
+      val lhs = Const (@{const_name term_of}, ty --> @{typ term})
+        $ Free ("x", ty);
+      val rhs = @{term "undefined \<Colon> term"};
+      val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
+      fun triv_name_of t = (fst o dest_Free o fst o strip_comb o fst
+        o HOLogic.dest_eq o HOLogic.dest_Trueprop) t ^ "_triv";
+    in
+      thy
+      |> TheoryTarget.instantiation ([tyco], vs, @{sort term_of})
+      |> `(fn lthy => Syntax.check_term lthy eq)
+      |-> (fn eq => Specification.definition (NONE, ((Binding.name (triv_name_of eq), []), eq)))
+      |> snd
+      |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
+    end;
+  fun ensure_term_of (tyco, (raw_vs, _)) thy =
+    let
+      val need_inst = not (can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort term_of})
+        andalso can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort typerep};
+    in if need_inst then add_term_of tyco raw_vs thy else thy end;
+in
+  Code.type_interpretation ensure_term_of
+end
+*}
+
+setup {*
+let
+  fun mk_term_of_eq thy ty vs tyco (c, tys) =
+    let
+      val t = list_comb (Const (c, tys ---> ty),
+        map Free (Name.names Name.context "a" tys));
+      val (arg, rhs) = pairself (Thm.cterm_of thy o map_types Logic.unvarifyT o Logic.varify)
+        (t, (map_aterms (fn t as Free (v, ty) => HOLogic.mk_term_of ty t | t => t) o HOLogic.reflect_term) t)
+      val cty = Thm.ctyp_of thy ty;
+    in
+      @{thm term_of_anything}
+      |> Drule.instantiate' [SOME cty] [SOME arg, SOME rhs]
+      |> Thm.varifyT
+    end;
+  fun add_term_of_code tyco raw_vs raw_cs thy =
+    let
+      val algebra = Sign.classes_of thy;
+      val vs = map (fn (v, sort) =>
+        (v, curry (Sorts.inter_sort algebra) @{sort typerep} sort)) raw_vs;
+      val ty = Type (tyco, map TFree vs);
+      val cs = (map o apsnd o map o map_atyps)
+        (fn TFree (v, _) => TFree (v, (the o AList.lookup (op =) vs) v)) raw_cs;
+      val const = AxClass.param_of_inst thy (@{const_name term_of}, tyco);
+      val eqs = map (mk_term_of_eq thy ty vs tyco) cs;
+   in
+      thy
+      |> Code.del_eqns const
+      |> fold Code.add_eqn eqs
+    end;
+  fun ensure_term_of_code (tyco, (raw_vs, cs)) thy =
+    let
+      val has_inst = can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort term_of};
+    in if has_inst then add_term_of_code tyco raw_vs cs thy else thy end;
+in
+  Code.type_interpretation ensure_term_of_code
+end
+*}
+
+
+subsubsection {* Code generator setup *}
+
+lemmas [code del] = term.recs term.cases term.size
+lemma [code, code del]: "eq_class.eq (t1\<Colon>term) t2 \<longleftrightarrow> eq_class.eq t1 t2" ..
+
+lemma [code, code del]: "(term_of \<Colon> typerep \<Rightarrow> term) = term_of" ..
+lemma [code, code del]: "(term_of \<Colon> term \<Rightarrow> term) = term_of" ..
+lemma [code, code del]: "(term_of \<Colon> String.literal \<Rightarrow> term) = term_of" ..
+lemma [code, code del]:
+  "(Code_Evaluation.term_of \<Colon> 'a::{type, term_of} Predicate.pred \<Rightarrow> Code_Evaluation.term) = Code_Evaluation.term_of" ..
+lemma [code, code del]:
+  "(Code_Evaluation.term_of \<Colon> 'a::{type, term_of} Predicate.seq \<Rightarrow> Code_Evaluation.term) = Code_Evaluation.term_of" ..
+
+lemma term_of_char [unfolded typerep_fun_def typerep_char_def typerep_nibble_def, code]: "Code_Evaluation.term_of c =
+    (let (n, m) = nibble_pair_of_char c
+  in Code_Evaluation.App (Code_Evaluation.App (Code_Evaluation.Const (STR ''String.char.Char'') (TYPEREP(nibble \<Rightarrow> nibble \<Rightarrow> char)))
+    (Code_Evaluation.term_of n)) (Code_Evaluation.term_of m))"
+  by (subst term_of_anything) rule 
+
+code_type "term"
+  (Eval "Term.term")
+
+code_const Const and App
+  (Eval "Term.Const/ ((_), (_))" and "Term.$/ ((_), (_))")
+
+code_const "term_of \<Colon> String.literal \<Rightarrow> term"
+  (Eval "HOLogic.mk'_message'_string")
+
+code_reserved Eval HOLogic
+
+
+subsubsection {* Syntax *}
+
+definition termify :: "'a \<Rightarrow> term" where
+  [code del]: "termify x = dummy_term"
+
+abbreviation valtermify :: "'a \<Rightarrow> 'a \<times> (unit \<Rightarrow> term)" where
+  "valtermify x \<equiv> (x, \<lambda>u. termify x)"
+
+setup {*
+let
+  fun map_default f xs =
+    let val ys = map f xs
+    in if exists is_some ys
+      then SOME (map2 the_default xs ys)
+      else NONE
+    end;
+  fun subst_termify_app (Const (@{const_name termify}, T), [t]) =
+        if not (Term.has_abs t)
+        then if fold_aterms (fn Const _ => I | _ => K false) t true
+          then SOME (HOLogic.reflect_term t)
+          else error "Cannot termify expression containing variables"
+        else error "Cannot termify expression containing abstraction"
+    | subst_termify_app (t, ts) = case map_default subst_termify ts
+       of SOME ts' => SOME (list_comb (t, ts'))
+        | NONE => NONE
+  and subst_termify (Abs (v, T, t)) = (case subst_termify t
+       of SOME t' => SOME (Abs (v, T, t'))
+        | NONE => NONE)
+    | subst_termify t = subst_termify_app (strip_comb t) 
+  fun check_termify ts ctxt = map_default subst_termify ts
+    |> Option.map (rpair ctxt)
+in
+  Context.theory_map (Syntax.add_term_check 0 "termify" check_termify)
+end;
+*}
+
+locale term_syntax
+begin
+
+notation App (infixl "<\<cdot>>" 70)
+  and valapp (infixl "{\<cdot>}" 70)
+
+end
+
+interpretation term_syntax .
+
+no_notation App (infixl "<\<cdot>>" 70)
+  and valapp (infixl "{\<cdot>}" 70)
+
+
+subsection {* Numeric types *}
+
+definition term_of_num :: "'a\<Colon>{semiring_div} \<Rightarrow> 'a\<Colon>{semiring_div} \<Rightarrow> term" where
+  "term_of_num two = (\<lambda>_. dummy_term)"
+
+lemma (in term_syntax) term_of_num_code [code]:
+  "term_of_num two k = (if k = 0 then termify Int.Pls
+    else (if k mod two = 0
+      then termify Int.Bit0 <\<cdot>> term_of_num two (k div two)
+      else termify Int.Bit1 <\<cdot>> term_of_num two (k div two)))"
+  by (auto simp add: term_of_anything Const_def App_def term_of_num_def Let_def)
+
+lemma (in term_syntax) term_of_nat_code [code]:
+  "term_of (n::nat) = termify (number_of :: int \<Rightarrow> nat) <\<cdot>> term_of_num (2::nat) n"
+  by (simp only: term_of_anything)
+
+lemma (in term_syntax) term_of_int_code [code]:
+  "term_of (k::int) = (if k = 0 then termify (0 :: int)
+    else if k > 0 then termify (number_of :: int \<Rightarrow> int) <\<cdot>> term_of_num (2::int) k
+      else termify (uminus :: int \<Rightarrow> int) <\<cdot>> (termify (number_of :: int \<Rightarrow> int) <\<cdot>> term_of_num (2::int) (- k)))"
+  by (simp only: term_of_anything)
+
+lemma (in term_syntax) term_of_code_numeral_code [code]:
+  "term_of (k::code_numeral) = termify (number_of :: int \<Rightarrow> code_numeral) <\<cdot>> term_of_num (2::code_numeral) k"
+  by (simp only: term_of_anything)
+
+subsection {* Obfuscate *}
+
+print_translation {*
+let
+  val term = Const ("<TERM>", dummyT);
+  fun tr1' [_, _] = term;
+  fun tr2' [] = term;
+in
+  [(@{const_syntax Const}, tr1'),
+    (@{const_syntax App}, tr1'),
+    (@{const_syntax dummy_term}, tr2')]
+end
+*}
+
+hide const dummy_term App valapp
+hide (open) const Const termify valtermify term_of term_of_num
+
+
+subsection {* Evaluation setup *}
+
+ML {*
+signature EVAL =
+sig
+  val eval_ref: (unit -> term) option ref
+  val eval_term: theory -> term -> term
+end;
+
+structure Eval : EVAL =
+struct
+
+val eval_ref = ref (NONE : (unit -> term) option);
+
+fun eval_term thy t =
+  Code_ML.eval NONE ("Eval.eval_ref", eval_ref) I thy (HOLogic.mk_term_of (fastype_of t) t) [];
+
+end;
+*}
+
+setup {*
+  Value.add_evaluator ("code", Eval.eval_term o ProofContext.theory_of)
+*}
+
+end
--- a/src/HOL/Complete_Lattice.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Complete_Lattice.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -76,11 +76,11 @@
 
 lemma sup_bot [simp]:
   "x \<squnion> bot = x"
-  using bot_least [of x] by (simp add: sup_commute)
+  using bot_least [of x] by (simp add: sup_commute sup_absorb2)
 
 lemma inf_top [simp]:
   "x \<sqinter> top = x"
-  using top_greatest [of x] by (simp add: inf_commute)
+  using top_greatest [of x] by (simp add: inf_commute inf_absorb2)
 
 definition SUPR :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
   "SUPR A f = \<Squnion> (f ` A)"
--- a/src/HOL/Decision_Procs/Approximation.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Decision_Procs/Approximation.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -1904,7 +1904,7 @@
 	show "0 < real x * 2/3" using * by auto
 	show "real ?max + 1 \<le> real x * 2/3" using * up
 	  by (cases "0 < real x * real (lapprox_posrat prec 2 3) - 1",
-	      auto simp add: real_of_float_max)
+	      auto simp add: real_of_float_max min_max.sup_absorb1)
       qed
       finally have "real (?lb_horner (Float 1 -1)) + real (?lb_horner ?max)
 	\<le> ln (real x)"
@@ -3246,12 +3246,13 @@
         = map (` (variable_of_bound o prop_of)) prems
 
       fun add_deps (name, bnds)
-        = Graph.add_deps_acyclic
-            (name, remove (op =) name (Term.add_free_names (prop_of bnds) []))
+        = Graph.add_deps_acyclic (name,
+            remove (op =) name (Term.add_free_names (prop_of bnds) []))
+
       val order = Graph.empty
                   |> fold Graph.new_node variable_bounds
                   |> fold add_deps variable_bounds
-                  |> Graph.topological_order |> rev
+                  |> Graph.strong_conn |> map the_single |> rev
                   |> map_filter (AList.lookup (op =) variable_bounds)
 
       fun prepend_prem th tac
@@ -3338,7 +3339,7 @@
                       etac @{thm meta_eqE},
                       rtac @{thm impI}] i)
       THEN Subgoal.FOCUS (fn {prems, ...} => reorder_bounds_tac prems i) @{context} i
-      THEN TRY (filter_prems_tac (K false) i)
+      THEN DETERM (TRY (filter_prems_tac (K false) i))
       THEN DETERM (Reflection.genreify_tac ctxt form_equations NONE i)
       THEN rewrite_interpret_form_tac ctxt prec splitting taylor i
       THEN gen_eval_tac eval_oracle ctxt i))
@@ -3350,7 +3351,7 @@
 
   fun mk_approx' prec t = (@{const "approx'"}
                          $ HOLogic.mk_number @{typ nat} prec
-                         $ t $ @{term "[] :: (float * float) list"})
+                         $ t $ @{term "[] :: (float * float) option list"})
 
   fun dest_result (Const (@{const_name "Some"}, _) $
                    ((Const (@{const_name "Pair"}, _)) $
--- a/src/HOL/Decision_Procs/Ferrack.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Decision_Procs/Ferrack.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -512,7 +512,7 @@
   assumes g0: "numgcd t = 0"
   shows "Inum bs t = 0"
   using g0[simplified numgcd_def] 
-  by (induct t rule: numgcdh.induct, auto simp add: natabs0 maxcoeff_pos)
+  by (induct t rule: numgcdh.induct, auto simp add: natabs0 maxcoeff_pos min_max.sup_absorb2)
 
 lemma numgcdh_pos: assumes gp: "g \<ge> 0" shows "numgcdh t g \<ge> 0"
   using gp
--- a/src/HOL/Decision_Procs/ex/Approximation_Ex.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Decision_Procs/ex/Approximation_Ex.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -72,7 +72,9 @@
   shows "g / v * tan (35 * d) \<in> { 3 * d .. 3.1 * d }"
   using assms by (approximation 80)
 
-lemma "\<phi> \<in> { 0 .. 1 :: real } \<longrightarrow> \<phi> ^ 2 \<le> \<phi>"
-  by (approximation 30 splitting: \<phi>=1 taylor: \<phi> = 3)
+lemma "x \<in> { 0 .. 1 :: real } \<longrightarrow> x ^ 2 \<le> x"
+  by (approximation 30 splitting: x=1 taylor: x = 3)
+
+value [approximate] "10"
 
 end
--- a/src/HOL/Finite_Set.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Finite_Set.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -2966,11 +2966,11 @@
 
 lemma dual_max:
   "ord.max (op \<ge>) = min"
-  by (auto simp add: ord.max_def_raw expand_fun_eq)
+  by (auto simp add: ord.max_def_raw min_def expand_fun_eq)
 
 lemma dual_min:
   "ord.min (op \<ge>) = max"
-  by (auto simp add: ord.min_def_raw expand_fun_eq)
+  by (auto simp add: ord.min_def_raw max_def expand_fun_eq)
 
 lemma strict_below_fold1_iff:
   assumes "finite A" and "A \<noteq> {}"
--- a/src/HOL/HOL.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/HOL.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -2021,6 +2021,29 @@
 
 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 *}
 
--- a/src/HOL/Hoare_Parallel/Graph.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Hoare_Parallel/Graph.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -187,6 +187,8 @@
 
 subsubsection{* Graph 3 *}
 
+declare min_max.inf_absorb1 [simp] min_max.inf_absorb2 [simp]
+
 lemma Graph3: 
   "\<lbrakk> T\<in>Reach E; R<length E \<rbrakk> \<Longrightarrow> Reach(E[R:=(fst(E!R),T)]) \<subseteq> Reach E"
 apply (unfold Reach_def)
@@ -307,6 +309,8 @@
 apply force
 done
 
+declare min_max.inf_absorb1 [simp del] min_max.inf_absorb2 [simp del]
+
 subsubsection {* Graph 5 *}
 
 lemma Graph5: 
--- a/src/HOL/Induct/LList.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Induct/LList.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -665,7 +665,7 @@
 apply (subst LList_corec, force)
 done
 
-lemma llist_corec: 
+lemma llist_corec [nitpick_const_simp]: 
     "llist_corec a f =   
      (case f a of None => LNil | Some(z,w) => LCons z (llist_corec w f))"
 apply (unfold llist_corec_def LNil_def LCons_def)
@@ -774,10 +774,11 @@
 
 subsection{* The functional @{text lmap} *}
 
-lemma lmap_LNil [simp]: "lmap f LNil = LNil"
+lemma lmap_LNil [simp, nitpick_const_simp]: "lmap f LNil = LNil"
 by (rule lmap_def [THEN def_llist_corec, THEN trans], simp)
 
-lemma lmap_LCons [simp]: "lmap f (LCons M N) = LCons (f M) (lmap f N)"
+lemma lmap_LCons [simp, nitpick_const_simp]:
+"lmap f (LCons M N) = LCons (f M) (lmap f N)"
 by (rule lmap_def [THEN def_llist_corec, THEN trans], simp)
 
 
@@ -792,7 +793,7 @@
 
 subsection{* iterates -- @{text llist_fun_equalityI} cannot be used! *}
 
-lemma iterates: "iterates f x = LCons x (iterates f (f x))"
+lemma iterates [nitpick_const_simp]: "iterates f x = LCons x (iterates f (f x))"
 by (rule iterates_def [THEN def_llist_corec, THEN trans], simp)
 
 lemma lmap_iterates [simp]: "lmap f (iterates f x) = iterates f (f x)"
@@ -847,18 +848,18 @@
 
 subsection{* @{text lappend} -- its two arguments cause some complications! *}
 
-lemma lappend_LNil_LNil [simp]: "lappend LNil LNil = LNil"
+lemma lappend_LNil_LNil [simp, nitpick_const_simp]: "lappend LNil LNil = LNil"
 apply (simp add: lappend_def)
 apply (rule llist_corec [THEN trans], simp)
 done
 
-lemma lappend_LNil_LCons [simp]: 
+lemma lappend_LNil_LCons [simp, nitpick_const_simp]: 
     "lappend LNil (LCons l l') = LCons l (lappend LNil l')"
 apply (simp add: lappend_def)
 apply (rule llist_corec [THEN trans], simp)
 done
 
-lemma lappend_LCons [simp]: 
+lemma lappend_LCons [simp, nitpick_const_simp]: 
     "lappend (LCons l l') N = LCons l (lappend l' N)"
 apply (simp add: lappend_def)
 apply (rule llist_corec [THEN trans], simp)
--- a/src/HOL/Inductive.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Inductive.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -267,26 +267,6 @@
   Ball_def Bex_def
   induct_rulify_fallback
 
-ML {*
-val def_lfp_unfold = @{thm def_lfp_unfold}
-val def_gfp_unfold = @{thm def_gfp_unfold}
-val def_lfp_induct = @{thm def_lfp_induct}
-val def_coinduct = @{thm def_coinduct}
-val inf_bool_eq = @{thm inf_bool_eq} RS @{thm eq_reflection}
-val inf_fun_eq = @{thm inf_fun_eq} RS @{thm eq_reflection}
-val sup_bool_eq = @{thm sup_bool_eq} RS @{thm eq_reflection}
-val sup_fun_eq = @{thm sup_fun_eq} RS @{thm eq_reflection}
-val le_boolI = @{thm le_boolI}
-val le_boolI' = @{thm le_boolI'}
-val le_funI = @{thm le_funI}
-val le_boolE = @{thm le_boolE}
-val le_funE = @{thm le_funE}
-val le_boolD = @{thm le_boolD}
-val le_funD = @{thm le_funD}
-val le_bool_def = @{thm le_bool_def} RS @{thm eq_reflection}
-val le_fun_def = @{thm le_fun_def} RS @{thm eq_reflection}
-*}
-
 use "Tools/inductive.ML"
 setup Inductive.setup
 
--- a/src/HOL/IsaMakefile	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/IsaMakefile	Thu Sep 24 17:26:05 2009 +0200
@@ -210,7 +210,7 @@
 
 MAIN_DEPENDENCIES = $(PLAIN_DEPENDENCIES) \
   ATP_Linkup.thy \
-  Code_Eval.thy \
+  Code_Evaluation.thy \
   Code_Numeral.thy \
   Equiv_Relations.thy \
   Groebner_Basis.thy \
@@ -909,7 +909,7 @@
   ex/Sudoku.thy ex/Tarski.thy \
   ex/Termination.thy ex/Transfer_Ex.thy ex/Unification.thy ex/document/root.bib		\
   ex/document/root.tex ex/set.thy ex/svc_funcs.ML ex/svc_test.thy \
-  ex/Predicate_Compile.thy ex/predicate_compile.ML ex/Predicate_Compile_ex.thy
+  ex/Predicate_Compile.thy Tools/Predicate_Compile/predicate_compile_core.ML ex/Predicate_Compile_ex.thy
 	@$(ISABELLE_TOOL) usedir $(OUT)/HOL ex
 
 
--- a/src/HOL/Lattices.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Lattices.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -127,10 +127,10 @@
 lemma inf_left_idem[simp]: "x \<sqinter> (x \<sqinter> y) = x \<sqinter> y"
   by (rule antisym) (auto intro: le_infI2)
 
-lemma inf_absorb1[simp]: "x \<sqsubseteq> y \<Longrightarrow> x \<sqinter> y = x"
+lemma inf_absorb1: "x \<sqsubseteq> y \<Longrightarrow> x \<sqinter> y = x"
   by (rule antisym) auto
 
-lemma inf_absorb2[simp]: "y \<sqsubseteq> x \<Longrightarrow> x \<sqinter> y = y"
+lemma inf_absorb2: "y \<sqsubseteq> x \<Longrightarrow> x \<sqinter> y = y"
   by (rule antisym) auto
 
 lemma inf_left_commute: "x \<sqinter> (y \<sqinter> z) = y \<sqinter> (x \<sqinter> z)"
@@ -155,10 +155,10 @@
 lemma sup_left_idem[simp]: "x \<squnion> (x \<squnion> y) = x \<squnion> y"
   by (rule antisym) (auto intro: le_supI2)
 
-lemma sup_absorb1[simp]: "y \<sqsubseteq> x \<Longrightarrow> x \<squnion> y = x"
+lemma sup_absorb1: "y \<sqsubseteq> x \<Longrightarrow> x \<squnion> y = x"
   by (rule antisym) auto
 
-lemma sup_absorb2[simp]: "x \<sqsubseteq> y \<Longrightarrow> x \<squnion> y = y"
+lemma sup_absorb2: "x \<sqsubseteq> y \<Longrightarrow> x \<squnion> y = y"
   by (rule antisym) auto
 
 lemma sup_left_commute: "x \<squnion> (y \<squnion> z) = y \<squnion> (x \<squnion> z)"
@@ -229,11 +229,11 @@
 
 lemma less_infI1:
   "a \<sqsubset> x \<Longrightarrow> a \<sqinter> b \<sqsubset> x"
-  by (auto simp add: less_le intro: le_infI1)
+  by (auto simp add: less_le inf_absorb1 intro: le_infI1)
 
 lemma less_infI2:
   "b \<sqsubset> x \<Longrightarrow> a \<sqinter> b \<sqsubset> x"
-  by (auto simp add: less_le intro: le_infI2)
+  by (auto simp add: less_le inf_absorb2 intro: le_infI2)
 
 end
 
--- a/src/HOL/Library/Code_Char.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Library/Code_Char.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -5,7 +5,7 @@
 header {* Code generation of pretty characters (and strings) *}
 
 theory Code_Char
-imports List Code_Eval Main
+imports List Code_Evaluation Main
 begin
 
 code_type char
@@ -32,7 +32,7 @@
   (OCaml "!((_ : char) = _)")
   (Haskell infixl 4 "==")
 
-code_const "Code_Eval.term_of \<Colon> char \<Rightarrow> term"
+code_const "Code_Evaluation.term_of \<Colon> char \<Rightarrow> term"
   (Eval "HOLogic.mk'_char/ (IntInf.fromInt/ (Char.ord/ _))")
 
 end
--- a/src/HOL/Library/Code_Integer.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Library/Code_Integer.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -100,7 +100,7 @@
 
 text {* Evaluation *}
 
-code_const "Code_Eval.term_of \<Colon> int \<Rightarrow> term"
+code_const "Code_Evaluation.term_of \<Colon> int \<Rightarrow> term"
   (Eval "HOLogic.mk'_number/ HOLogic.intT")
 
 end
\ No newline at end of file
--- a/src/HOL/Library/Coinductive_List.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Library/Coinductive_List.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -260,7 +260,7 @@
   qed
 qed
 
-lemma llist_corec [code]:
+lemma llist_corec [code, nitpick_const_simp]:
   "llist_corec a f =
     (case f a of None \<Rightarrow> LNil | Some (z, w) \<Rightarrow> LCons z (llist_corec w f))"
 proof (cases "f a")
@@ -656,8 +656,9 @@
   qed
 qed
 
-lemma lmap_LNil [simp]: "lmap f LNil = LNil"
-  and lmap_LCons [simp]: "lmap f (LCons M N) = LCons (f M) (lmap f N)"
+lemma lmap_LNil [simp, nitpick_const_simp]: "lmap f LNil = LNil"
+  and lmap_LCons [simp, nitpick_const_simp]:
+  "lmap f (LCons M N) = LCons (f M) (lmap f N)"
   by (simp_all add: lmap_def llist_corec)
 
 lemma lmap_compose [simp]: "lmap (f o g) l = lmap f (lmap g l)"
@@ -728,9 +729,9 @@
   qed
 qed
 
-lemma lappend_LNil_LNil [simp]: "lappend LNil LNil = LNil"
-  and lappend_LNil_LCons [simp]: "lappend LNil (LCons l l') = LCons l (lappend LNil l')"
-  and lappend_LCons [simp]: "lappend (LCons l l') m = LCons l (lappend l' m)"
+lemma lappend_LNil_LNil [simp, nitpick_const_simp]: "lappend LNil LNil = LNil"
+  and lappend_LNil_LCons [simp, nitpick_const_simp]: "lappend LNil (LCons l l') = LCons l (lappend LNil l')"
+  and lappend_LCons [simp, nitpick_const_simp]: "lappend (LCons l l') m = LCons l (lappend l' m)"
   by (simp_all add: lappend_def llist_corec)
 
 lemma lappend_LNil1 [simp]: "lappend LNil l = l"
@@ -754,7 +755,7 @@
   iterates :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a llist" where
   "iterates f a = llist_corec a (\<lambda>x. Some (x, f x))"
 
-lemma iterates: "iterates f x = LCons x (iterates f (f x))"
+lemma iterates [nitpick_const_simp]: "iterates f x = LCons x (iterates f (f x))"
   apply (unfold iterates_def)
   apply (subst llist_corec)
   apply simp
--- a/src/HOL/Library/Efficient_Nat.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Library/Efficient_Nat.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -415,9 +415,9 @@
 text {* Evaluation *}
 
 lemma [code, code del]:
-  "(Code_Eval.term_of \<Colon> nat \<Rightarrow> term) = Code_Eval.term_of" ..
+  "(Code_Evaluation.term_of \<Colon> nat \<Rightarrow> term) = Code_Evaluation.term_of" ..
 
-code_const "Code_Eval.term_of \<Colon> nat \<Rightarrow> term"
+code_const "Code_Evaluation.term_of \<Colon> nat \<Rightarrow> term"
   (SML "HOLogic.mk'_number/ HOLogic.natT")
 
 
--- a/src/HOL/Library/Executable_Set.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Library/Executable_Set.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -85,4 +85,6 @@
   card                ("{*Fset.card*}")
   fold                ("{*foldl o flip*}")
 
+hide (open) const subset eq_set Inter Union flip
+
 end
\ No newline at end of file
--- a/src/HOL/Library/Fin_Fun.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Library/Fin_Fun.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -311,17 +311,17 @@
 notation scomp (infixl "o\<rightarrow>" 60)
 
 definition (in term_syntax) valtermify_finfun_const ::
-  "'b\<Colon>typerep \<times> (unit \<Rightarrow> Code_Eval.term) \<Rightarrow> ('a\<Colon>typerep \<Rightarrow>\<^isub>f 'b) \<times> (unit \<Rightarrow> Code_Eval.term)" where
-  "valtermify_finfun_const y = Code_Eval.valtermify finfun_const {\<cdot>} y"
+  "'b\<Colon>typerep \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> ('a\<Colon>typerep \<Rightarrow>\<^isub>f 'b) \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
+  "valtermify_finfun_const y = Code_Evaluation.valtermify finfun_const {\<cdot>} y"
 
 definition (in term_syntax) valtermify_finfun_update_code ::
-  "'a\<Colon>typerep \<times> (unit \<Rightarrow> Code_Eval.term) \<Rightarrow> 'b\<Colon>typerep \<times> (unit \<Rightarrow> Code_Eval.term) \<Rightarrow> ('a \<Rightarrow>\<^isub>f 'b) \<times> (unit \<Rightarrow> Code_Eval.term) \<Rightarrow> ('a \<Rightarrow>\<^isub>f 'b) \<times> (unit \<Rightarrow> Code_Eval.term)" where
-  "valtermify_finfun_update_code x y f = Code_Eval.valtermify finfun_update_code {\<cdot>} f {\<cdot>} x {\<cdot>} y"
+  "'a\<Colon>typerep \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> 'b\<Colon>typerep \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> ('a \<Rightarrow>\<^isub>f 'b) \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> ('a \<Rightarrow>\<^isub>f 'b) \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
+  "valtermify_finfun_update_code x y f = Code_Evaluation.valtermify finfun_update_code {\<cdot>} f {\<cdot>} x {\<cdot>} y"
 
 instantiation finfun :: (random, random) random
 begin
 
-primrec random_finfun_aux :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> Random.seed \<Rightarrow> ('a \<Rightarrow>\<^isub>f 'b \<times> (unit \<Rightarrow> Code_Eval.term)) \<times> Random.seed" where
+primrec random_finfun_aux :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> Random.seed \<Rightarrow> ('a \<Rightarrow>\<^isub>f 'b \<times> (unit \<Rightarrow> Code_Evaluation.term)) \<times> Random.seed" where
     "random_finfun_aux 0 j = Quickcheck.collapse (Random.select_weight
        [(1, Quickcheck.random j o\<rightarrow> (\<lambda>y. Pair (valtermify_finfun_const y)))])"
   | "random_finfun_aux (Suc_code_numeral i) j = Quickcheck.collapse (Random.select_weight
--- a/src/HOL/Library/Nested_Environment.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Library/Nested_Environment.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -567,6 +567,6 @@
 qed simp_all
 
 lemma [code, code del]:
-  "(Code_Eval.term_of :: ('a::{term_of, type}, 'b::{term_of, type}, 'c::{term_of, type}) env \<Rightarrow> term) = Code_Eval.term_of" ..
+  "(Code_Evaluation.term_of :: ('a::{term_of, type}, 'b::{term_of, type}, 'c::{term_of, type}) env \<Rightarrow> term) = Code_Evaluation.term_of" ..
 
 end
--- a/src/HOL/Library/Sum_Of_Squares.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Library/Sum_Of_Squares.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -10,6 +10,7 @@
 uses
   ("positivstellensatz.ML")  (* duplicate use!? -- cf. Euclidian_Space.thy *)
   ("Sum_Of_Squares/sum_of_squares.ML")
+  ("Sum_Of_Squares/positivstellensatz_tools.ML")
   ("Sum_Of_Squares/sos_wrapper.ML")
 begin
 
@@ -22,112 +23,142 @@
   of a minute for one sos call, because sos calls CSDP repeatedly.  If
   you install CSDP locally, sos calls typically takes only a few
   seconds.
+  sos generates a certificate which can be used to repeat the proof
+  without calling an external prover.
 *}
 
 text {* setup sos tactic *}
 
 use "positivstellensatz.ML"
+use "Sum_Of_Squares/positivstellensatz_tools.ML"
 use "Sum_Of_Squares/sum_of_squares.ML"
 use "Sum_Of_Squares/sos_wrapper.ML"
 
 setup SosWrapper.setup
 
-text {* Tests -- commented since they work only when csdp is installed
-  or take too long with remote csdps *}
+text {* Tests *}
+
+lemma "(3::real) * x + 7 * a < 4 & 3 < 2 * x \<Longrightarrow> a < 0"
+by (sos_cert "((R<1 + (((A<1 * R<1) * (R<2 * [1]^2)) + (((A<0 * R<1) * (R<3 * [1]^2)) + ((A<=0 * R<1) * (R<14 * [1]^2))))))")
+
+lemma "a1 >= 0 & a2 >= 0 \<and> (a1 * a1 + a2 * a2 = b1 * b1 + b2 * b2 + 2) \<and> (a1 * b1 + a2 * b2 = 0) --> a1 * a2 - b1 * b2 >= (0::real)"
+by (sos_cert "(((A<0 * R<1) + (([~1/2*a1*b2 + ~1/2*a2*b1] * A=0) + (([~1/2*a1*a2 + 1/2*b1*b2] * A=1) + (((A<0 * R<1) * ((R<1/2 * [b2]^2) + (R<1/2 * [b1]^2))) + ((A<=0 * (A<=1 * R<1)) * ((R<1/2 * [b2]^2) + ((R<1/2 * [b1]^2) + ((R<1/2 * [a2]^2) + (R<1/2 * [a1]^2))))))))))")
 
-(*
-lemma "(3::real) * x + 7 * a < 4 & 3 < 2 * x \<Longrightarrow> a < 0" by sos
+lemma "(3::real) * x + 7 * a < 4 & 3 < 2 * x --> a < 0"
+by (sos_cert "((R<1 + (((A<1 * R<1) * (R<2 * [1]^2)) + (((A<0 * R<1) * (R<3 * [1]^2)) + ((A<=0 * R<1) * (R<14 * [1]^2))))))")
+
+lemma "(0::real) <= x & x <= 1 & 0 <= y & y <= 1  --> x^2 + y^2 < 1 |(x - 1)^2 + y^2 < 1 | x^2 + (y - 1)^2 < 1 | (x - 1)^2 + (y - 1)^2 < 1"
+by (sos_cert "((R<1 + (((A<=3 * (A<=4 * R<1)) * (R<1 * [1]^2)) + (((A<=2 * (A<=7 * R<1)) * (R<1 * [1]^2)) + (((A<=1 * (A<=6 * R<1)) * (R<1 * [1]^2)) + ((A<=0 * (A<=5 * R<1)) * (R<1 * [1]^2)))))))")
 
-lemma "a1 >= 0 & a2 >= 0 \<and> (a1 * a1 + a2 * a2 = b1 * b1 + b2 * b2 + 2) \<and>
-  (a1 * b1 + a2 * b2 = 0) --> a1 * a2 - b1 * b2 >= (0::real)" by sos
+lemma "(0::real) <= x & 0 <= y & 0 <= z & x + y + z <= 3 --> x * y + x * z + y * z >= 3 * x * y * z"
+by (sos_cert "(((A<0 * R<1) + (((A<0 * R<1) * (R<1/2 * [1]^2)) + (((A<=2 * R<1) * (R<1/2 * [~1*x + y]^2)) + (((A<=1 * R<1) * (R<1/2 * [~1*x + z]^2)) + (((A<=1 * (A<=2 * (A<=3 * R<1))) * (R<1/2 * [1]^2)) + (((A<=0 * R<1) * (R<1/2 * [~1*y + z]^2)) + (((A<=0 * (A<=2 * (A<=3 * R<1))) * (R<1/2 * [1]^2)) + ((A<=0 * (A<=1 * (A<=3 * R<1))) * (R<1/2 * [1]^2))))))))))")
+
+lemma "((x::real)^2 + y^2 + z^2 = 1) --> (x + y + z)^2 <= 3"
+by (sos_cert "(((A<0 * R<1) + (([~3] * A=0) + (R<1 * ((R<2 * [~1/2*x + ~1/2*y + z]^2) + (R<3/2 * [~1*x + y]^2))))))")
 
-lemma "(3::real) * x + 7 * a < 4 & 3 < 2 * x --> a < 0" by sos
+lemma "(w^2 + x^2 + y^2 + z^2 = 1) --> (w + x + y + z)^2 <= (4::real)"
+by (sos_cert "(((A<0 * R<1) + (([~4] * A=0) + (R<1 * ((R<3 * [~1/3*w + ~1/3*x + ~1/3*y + z]^2) + ((R<8/3 * [~1/2*w + ~1/2*x + y]^2) + (R<2 * [~1*w + x]^2)))))))")
 
-lemma "(0::real) <= x & x <= 1 & 0 <= y & y <= 1 -->
-  x^2 + y^2 < 1 |(x - 1)^2 + y^2 < 1 | x^2 + (y - 1)^2 < 1 | (x - 1)^2 + (y - 1)^2 < 1" by sos
+lemma "(x::real) >= 1 & y >= 1 --> x * y >= x + y - 1"
+by (sos_cert "(((A<0 * R<1) + ((A<=0 * (A<=1 * R<1)) * (R<1 * [1]^2))))")
+
+lemma "(x::real) > 1 & y > 1 --> x * y > x + y - 1"
+by (sos_cert "((((A<0 * A<1) * R<1) + ((A<=0 * R<1) * (R<1 * [1]^2))))") 
 
-lemma "(0::real) <= x & 0 <= y & 0 <= z & x + y + z <= 3 -->
-  x * y + x * z + y * z >= 3 * x * y * z" by sos
+lemma "abs(x) <= 1 --> abs(64 * x^7 - 112 * x^5 + 56 * x^3 - 7 * x) <= (1::real)"
+by (sos_cert "((((A<0 * R<1) + ((A<=1 * R<1) * (R<1 * [~8*x^3 + ~4*x^2 + 4*x + 1]^2)))) & ((((A<0 * A<1) * R<1) + ((A<=1 * (A<0 * R<1)) * (R<1 * [8*x^3 + ~4*x^2 + ~4*x + 1]^2)))))")
+
+(* ------------------------------------------------------------------------- *)
+(* One component of denominator in dodecahedral example.                     *)
+(* ------------------------------------------------------------------------- *)
 
-lemma "((x::real)^2 + y^2 + z^2 = 1) --> (x + y + z)^2 <= 3" by sos
+lemma "2 <= x & x <= 125841 / 50000 & 2 <= y & y <= 125841 / 50000 & 2 <= z & z <= 125841 / 50000 --> 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z) >= (0::real)"
+by (sos_cert "(((A<0 * R<1) + ((R<1 * ((R<5749028157/5000000000 * [~25000/222477*x + ~25000/222477*y + ~25000/222477*z + 1]^2) + ((R<864067/1779816 * [419113/864067*x + 419113/864067*y + z]^2) + ((R<320795/864067 * [419113/1283180*x + y]^2) + (R<1702293/5132720 * [x]^2))))) + (((A<=4 * (A<=5 * R<1)) * (R<3/2 * [1]^2)) + (((A<=3 * (A<=5 * R<1)) * (R<1/2 * [1]^2)) + (((A<=2 * (A<=4 * R<1)) * (R<1 * [1]^2)) + (((A<=2 * (A<=3 * R<1)) * (R<3/2 * [1]^2)) + (((A<=1 * (A<=5 * R<1)) * (R<1/2 * [1]^2)) + (((A<=1 * (A<=3 * R<1)) * (R<1/2 * [1]^2)) + (((A<=0 * (A<=4 * R<1)) * (R<1 * [1]^2)) + (((A<=0 * (A<=2 * R<1)) * (R<1 * [1]^2)) + ((A<=0 * (A<=1 * R<1)) * (R<3/2 * [1]^2)))))))))))))")
 
-lemma "(w^2 + x^2 + y^2 + z^2 = 1) --> (w + x + y + z)^2 <= (4::real)" by sos
+(* ------------------------------------------------------------------------- *)
+(* Over a larger but simpler interval.                                       *)
+(* ------------------------------------------------------------------------- *)
 
-lemma "(x::real) >= 1 & y >= 1 --> x * y >= x + y - 1" by sos
+lemma "(2::real) <= x & x <= 4 & 2 <= y & y <= 4 & 2 <= z & z <= 4 --> 0 <= 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z)"
+by (sos_cert "((R<1 + ((R<1 * ((R<1 * [~1/6*x + ~1/6*y + ~1/6*z + 1]^2) + ((R<1/18 * [~1/2*x + ~1/2*y + z]^2) + (R<1/24 * [~1*x + y]^2)))) + (((A<0 * R<1) * (R<1/12 * [1]^2)) + (((A<=4 * (A<=5 * R<1)) * (R<1/6 * [1]^2)) + (((A<=2 * (A<=4 * R<1)) * (R<1/6 * [1]^2)) + (((A<=2 * (A<=3 * R<1)) * (R<1/6 * [1]^2)) + (((A<=0 * (A<=4 * R<1)) * (R<1/6 * [1]^2)) + (((A<=0 * (A<=2 * R<1)) * (R<1/6 * [1]^2)) + ((A<=0 * (A<=1 * R<1)) * (R<1/6 * [1]^2)))))))))))")
 
-lemma "(x::real) > 1 & y > 1 --> x * y > x + y - 1" by sos; 
+(* ------------------------------------------------------------------------- *)
+(* We can do 12. I think 12 is a sharp bound; see PP's certificate.          *)
+(* ------------------------------------------------------------------------- *)
+
+lemma "2 <= (x::real) & x <= 4 & 2 <= y & y <= 4 & 2 <= z & z <= 4 --> 12 <= 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z)"
+by (sos_cert "(((A<0 * R<1) + (((A<=4 * R<1) * (R<2/3 * [1]^2)) + (((A<=4 * (A<=5 * R<1)) * (R<1 * [1]^2)) + (((A<=3 * (A<=4 * R<1)) * (R<1/3 * [1]^2)) + (((A<=2 * R<1) * (R<2/3 * [1]^2)) + (((A<=2 * (A<=5 * R<1)) * (R<1/3 * [1]^2)) + (((A<=2 * (A<=4 * R<1)) * (R<8/3 * [1]^2)) + (((A<=2 * (A<=3 * R<1)) * (R<1 * [1]^2)) + (((A<=1 * (A<=4 * R<1)) * (R<1/3 * [1]^2)) + (((A<=1 * (A<=2 * R<1)) * (R<1/3 * [1]^2)) + (((A<=0 * R<1) * (R<2/3 * [1]^2)) + (((A<=0 * (A<=5 * R<1)) * (R<1/3 * [1]^2)) + (((A<=0 * (A<=4 * R<1)) * (R<8/3 * [1]^2)) + (((A<=0 * (A<=3 * R<1)) * (R<1/3 * [1]^2)) + (((A<=0 * (A<=2 * R<1)) * (R<8/3 * [1]^2)) + ((A<=0 * (A<=1 * R<1)) * (R<1 * [1]^2))))))))))))))))))")
 
-lemma "abs(x) <= 1 --> abs(64 * x^7 - 112 * x^5 + 56 * x^3 - 7 * x) <= (1::real)" by sos  
+(* ------------------------------------------------------------------------- *)
+(* Inequality from sci.math (see "Leon-Sotelo, por favor").                  *)
+(* ------------------------------------------------------------------------- *)
 
+lemma "0 <= (x::real) & 0 <= y & (x * y = 1) --> x + y <= x^2 + y^2"
+by (sos_cert "(((A<0 * R<1) + (([1] * A=0) + (R<1 * ((R<1 * [~1/2*x + ~1/2*y + 1]^2) + (R<3/4 * [~1*x + y]^2))))))") 
 
-text {* One component of denominator in dodecahedral example. *}
+lemma "0 <= (x::real) & 0 <= y & (x * y = 1) --> x * y * (x + y) <= x^2 + y^2"
+by (sos_cert "(((A<0 * R<1) + (([~1*x + ~1*y + 1] * A=0) + (R<1 * ((R<1 * [~1/2*x + ~1/2*y + 1]^2) + (R<3/4 * [~1*x + y]^2))))))") 
 
-lemma "2 <= x & x <= 125841 / 50000 & 2 <= y & y <= 125841 / 50000 & 2 <= z &
-  z <= 125841 / 50000 --> 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z) >= (0::real)" by sos
+lemma "0 <= (x::real) & 0 <= y --> x * y * (x + y)^2 <= (x^2 + y^2)^2"
+by (sos_cert "(((A<0 * R<1) + (R<1 * ((R<1 * [~1/2*x^2 + y^2 + ~1/2*x*y]^2) + (R<3/4 * [~1*x^2 + x*y]^2)))))")
 
+lemma "(0::real) <= a & 0 <= b & 0 <= c & c * (2 * a + b)^3/ 27 <= x \<longrightarrow> c * a^2 * b <= x"
+by (sos_cert "(((A<0 * R<1) + (((A<=3 * R<1) * (R<1 * [1]^2)) + (((A<=1 * (A<=2 * R<1)) * (R<1/27 * [~1*a + b]^2)) + ((A<=0 * (A<=2 * R<1)) * (R<8/27 * [~1*a + b]^2))))))")
+ 
+lemma "(0::real) < x --> 0 < 1 + x + x^2"
+by (sos_cert "((R<1 + ((R<1 * (R<1 * [x]^2)) + (((A<0 * R<1) * (R<1 * [1]^2)) + ((A<=0 * R<1) * (R<1 * [1]^2))))))")
 
-text {* Over a larger but simpler interval. *}
+lemma "(0::real) <= x --> 0 < 1 + x + x^2"
+by (sos_cert "((R<1 + ((R<1 * (R<1 * [x]^2)) + (((A<=1 * R<1) * (R<1 * [1]^2)) + ((A<=0 * R<1) * (R<1 * [1]^2))))))")
 
-lemma "(2::real) <= x & x <= 4 & 2 <= y & y <= 4 & 2 <= z &
-  z <= 4 --> 0 <= 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z)" by sos
+lemma "(0::real) < 1 + x^2"
+by (sos_cert "((R<1 + ((R<1 * (R<1 * [x]^2)) + ((A<=0 * R<1) * (R<1 * [1]^2)))))")
+
+lemma "(0::real) <= 1 + 2 * x + x^2"
+by (sos_cert "(((A<0 * R<1) + (R<1 * (R<1 * [x + 1]^2))))")
 
-text {* We can do 12. I think 12 is a sharp bound; see PP's certificate. *}
+lemma "(0::real) < 1 + abs x"
+by (sos_cert "((R<1 + (((A<=1 * R<1) * (R<1/2 * [1]^2)) + ((A<=0 * R<1) * (R<1/2 * [1]^2)))))")
 
-lemma "2 <= (x::real) & x <= 4 & 2 <= y & y <= 4 & 2 <= z & z <= 4 -->
-  12 <= 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z)" by sos
+lemma "(0::real) < 1 + (1 + x)^2 * (abs x)"
+by (sos_cert "(((R<1 + (((A<=1 * R<1) * (R<1 * [1]^2)) + ((A<=0 * R<1) * (R<1 * [x + 1]^2))))) & ((R<1 + (((A<0 * R<1) * (R<1 * [x + 1]^2)) + ((A<=0 * R<1) * (R<1 * [1]^2))))))")
 
 
-text {* Inequality from sci.math (see "Leon-Sotelo, por favor"). *}
 
-lemma "0 <= (x::real) & 0 <= y & (x * y = 1) --> x + y <= x^2 + y^2" by sos 
-
-lemma "0 <= (x::real) & 0 <= y & (x * y = 1) --> x * y * (x + y) <= x^2 + y^2" by sos 
-
-lemma "0 <= (x::real) & 0 <= y --> x * y * (x + y)^2 <= (x^2 + y^2)^2" by sos
+lemma "abs ((1::real) + x^2) = (1::real) + x^2"
+by (sos_cert "(() & (((R<1 + ((R<1 * (R<1 * [x]^2)) + ((A<1 * R<1) * (R<1/2 * [1]^2))))) & ((R<1 + ((R<1 * (R<1 * [x]^2)) + ((A<0 * R<1) * (R<1 * [1]^2)))))))")
+lemma "(3::real) * x + 7 * a < 4 \<and> 3 < 2 * x \<longrightarrow> a < 0"
+by (sos_cert "((R<1 + (((A<1 * R<1) * (R<2 * [1]^2)) + (((A<0 * R<1) * (R<3 * [1]^2)) + ((A<=0 * R<1) * (R<14 * [1]^2))))))")
 
-lemma "(0::real) <= a & 0 <= b & 0 <= c & c * (2 * a + b)^3/ 27 <= x \<longrightarrow> c * a^2 * b <= x" by sos
- 
-lemma "(0::real) < x --> 0 < 1 + x + x^2" by sos
-
-lemma "(0::real) <= x --> 0 < 1 + x + x^2" by sos
-
-lemma "(0::real) < 1 + x^2" by sos
-
-lemma "(0::real) <= 1 + 2 * x + x^2" by sos
-
-lemma "(0::real) < 1 + abs x" by sos
-
-lemma "(0::real) < 1 + (1 + x)^2 * (abs x)" by sos
+lemma "(0::real) < x --> 1 < y --> y * x <= z --> x < z"
+by (sos_cert "((((A<0 * A<1) * R<1) + (((A<=1 * R<1) * (R<1 * [1]^2)) + ((A<=0 * R<1) * (R<1 * [1]^2)))))")
+lemma "(1::real) < x --> x^2 < y --> 1 < y"
+by (sos_cert "((((A<0 * A<1) * R<1) + ((R<1 * ((R<1/10 * [~2*x + y + 1]^2) + (R<1/10 * [~1*x + y]^2))) + (((A<1 * R<1) * (R<1/2 * [1]^2)) + (((A<0 * R<1) * (R<1 * [x]^2)) + (((A<=0 * R<1) * ((R<1/10 * [x + 1]^2) + (R<1/10 * [x]^2))) + (((A<=0 * (A<1 * R<1)) * (R<1/5 * [1]^2)) + ((A<=0 * (A<0 * R<1)) * (R<1/5 * [1]^2)))))))))")
+lemma "(b::real)^2 < 4 * a * c --> ~(a * x^2 + b * x + c = 0)"
+by (sos_cert "(((A<0 * R<1) + (R<1 * (R<1 * [2*a*x + b]^2))))")
+lemma "(b::real)^2 < 4 * a * c --> ~(a * x^2 + b * x + c = 0)"
+by (sos_cert "(((A<0 * R<1) + (R<1 * (R<1 * [2*a*x + b]^2))))")
+lemma "((a::real) * x^2 + b * x + c = 0) --> b^2 >= 4 * a * c"
+by (sos_cert "(((A<0 * R<1) + (R<1 * (R<1 * [2*a*x + b]^2))))")
+lemma "(0::real) <= b & 0 <= c & 0 <= x & 0 <= y & (x^2 = c) & (y^2 = a^2 * c + b) --> a * c <= y * x"
+by (sos_cert "(((A<0 * (A<0 * R<1)) + (((A<=2 * (A<=3 * (A<0 * R<1))) * (R<2 * [1]^2)) + ((A<=0 * (A<=1 * R<1)) * (R<1 * [1]^2)))))")
+lemma "abs(x - z) <= e & abs(y - z) <= e & 0 <= u & 0 <= v & (u + v = 1) --> abs((u * x + v * y) - z) <= (e::real)"
+by (sos_cert "((((A<0 * R<1) + (((A<=3 * (A<=6 * R<1)) * (R<1 * [1]^2)) + ((A<=1 * (A<=5 * R<1)) * (R<1 * [1]^2))))) & ((((A<0 * A<1) * R<1) + (((A<=3 * (A<=5 * (A<0 * R<1))) * (R<1 * [1]^2)) + ((A<=1 * (A<=4 * (A<0 * R<1))) * (R<1 * [1]^2))))))")
 
 
-lemma "abs ((1::real) + x^2) = (1::real) + x^2" by sos
-lemma "(3::real) * x + 7 * a < 4 \<and> 3 < 2 * x \<longrightarrow> a < 0" by sos
-
-lemma "(0::real) < x --> 1 < y --> y * x <= z --> x < z" by sos
-lemma "(1::real) < x --> x^2 < y --> 1 < y" by sos
-lemma "(b::real)^2 < 4 * a * c --> ~(a * x^2 + b * x + c = 0)" by sos
-lemma "(b::real)^2 < 4 * a * c --> ~(a * x^2 + b * x + c = 0)" by sos
-lemma "((a::real) * x^2 + b * x + c = 0) --> b^2 >= 4 * a * c" by sos
-lemma "(0::real) <= b & 0 <= c & 0 <= x & 0 <= y & (x^2 = c) & (y^2 = a^2 * c + b) --> a * c <= y * x" by sos
-lemma "abs(x - z) <= e & abs(y - z) <= e & 0 <= u & 0 <= v & (u + v = 1) -->
-  abs((u * x + v * y) - z) <= (e::real)" by sos
-
-(*
-lemma "((x::real) - y - 2 * x^4 = 0) & 0 <= x & x <= 2 & 0 <= y & y <= 3 -->
-  y^2 - 7 * y - 12 * x + 17 >= 0" by sos  -- {* Too hard?*}
-*)
+(* lemma "((x::real) - y - 2 * x^4 = 0) & 0 <= x & x <= 2 & 0 <= y & y <= 3 --> y^2 - 7 * y - 12 * x + 17 >= 0" by sos *) (* Too hard?*)
 
 lemma "(0::real) <= x --> (1 + x + x^2)/(1 + x^2) <= 1 + x"
-  by sos
+by (sos_cert "(((((A<0 * A<1) * R<1) + ((A<=0 * (A<0 * R<1)) * (R<1 * [x]^2)))) & ((R<1 + ((R<1 * (R<1 * [x]^2)) + ((A<0 * R<1) * (R<1 * [1]^2))))))")
 
 lemma "(0::real) <= x --> 1 - x <= 1 / (1 + x + x^2)"
-  by sos
+by (sos_cert "(((R<1 + (([~4/3] * A=0) + ((R<1 * ((R<1/3 * [3/2*x + 1]^2) + (R<7/12 * [x]^2))) + ((A<=0 * R<1) * (R<1/3 * [1]^2)))))) & (((((A<0 * A<1) * R<1) + ((A<=0 * (A<0 * R<1)) * (R<1 * [x]^2)))) & ((R<1 + ((R<1 * (R<1 * [x]^2)) + (((A<0 * R<1) * (R<1 * [1]^2)) + ((A<=0 * R<1) * (R<1 * [1]^2))))))))")
 
 lemma "(x::real) <= 1 / 2 --> - x - 2 * x^2 <= - x / (1 - x)"
-  by sos
+by (sos_cert "((((A<0 * A<1) * R<1) + ((A<=0 * (A<0 * R<1)) * (R<1 * [x]^2))))")
 
-lemma "4*r^2 = p^2 - 4*q & r >= (0::real) & x^2 + p*x + q = 0 -->
-  2*(x::real) = - p + 2*r | 2*x = -p - 2*r" by sos
-*)
+lemma "4*r^2 = p^2 - 4*q & r >= (0::real) & x^2 + p*x + q = 0 --> 2*(x::real) = - p + 2*r | 2*x = -p - 2*r"
+by (sos_cert "((((((A<0 * A<1) * R<1) + ([~4] * A=0))) & ((((A<0 * A<1) * R<1) + ([4] * A=0)))) & (((((A<0 * A<1) * R<1) + ([4] * A=0))) & ((((A<0 * A<1) * R<1) + ([~4] * A=0)))))")
 
 end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Sum_Of_Squares/positivstellensatz_tools.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -0,0 +1,158 @@
+(* Title:      positivstellensatz_tools.ML
+   Author:     Philipp Meyer, TU Muenchen
+
+Functions for generating a certificate from a positivstellensatz and vice versa
+*)
+
+signature POSITIVSTELLENSATZ_TOOLS =
+sig
+  val pss_tree_to_cert : RealArith.pss_tree -> string
+
+  val cert_to_pss_tree : Proof.context -> string -> RealArith.pss_tree
+
+end
+
+
+structure PositivstellensatzTools : POSITIVSTELLENSATZ_TOOLS =
+struct
+
+open RealArith FuncUtil
+
+(*** certificate generation ***)
+
+fun string_of_rat r =
+  let
+    val (nom, den) = Rat.quotient_of_rat r
+  in
+    if den = 1 then string_of_int nom
+    else string_of_int nom ^ "/" ^ string_of_int den
+  end
+
+(* map polynomials to strings *)
+
+fun string_of_varpow x k =
+  let
+    val term = term_of x
+    val name = case term of
+      Free (n, _) => n
+    | _ => error "Term in monomial not free variable"
+  in
+    if k = 1 then name else name ^ "^" ^ string_of_int k 
+  end
+
+fun string_of_monomial m = 
+ if Ctermfunc.is_undefined m then "1" 
+ else 
+  let 
+   val m' = dest_monomial m
+   val vps = fold_rev (fn (x,k) => cons (string_of_varpow x k)) m' [] 
+  in foldr1 (fn (s, t) => s ^ "*" ^ t) vps
+  end
+
+fun string_of_cmonomial (m,c) =
+  if Ctermfunc.is_undefined m then string_of_rat c
+  else if c = Rat.one then string_of_monomial m
+  else (string_of_rat c) ^ "*" ^ (string_of_monomial m);
+
+fun string_of_poly p = 
+ if Monomialfunc.is_undefined p then "0" 
+ else
+  let 
+   val cms = map string_of_cmonomial
+     (sort (prod_ord monomial_order (K EQUAL)) (Monomialfunc.graph p))
+  in foldr1 (fn (t1, t2) => t1 ^ " + " ^ t2) cms
+  end;
+
+fun pss_to_cert (Axiom_eq i) = "A=" ^ string_of_int i
+  | pss_to_cert (Axiom_le i) = "A<=" ^ string_of_int i
+  | pss_to_cert (Axiom_lt i) = "A<" ^ string_of_int i
+  | pss_to_cert (Rational_eq r) = "R=" ^ string_of_rat r
+  | pss_to_cert (Rational_le r) = "R<=" ^ string_of_rat r
+  | pss_to_cert (Rational_lt r) = "R<" ^ string_of_rat r
+  | pss_to_cert (Square p) = "[" ^ string_of_poly p ^ "]^2"
+  | pss_to_cert (Eqmul (p, pss)) = "([" ^ string_of_poly p ^ "] * " ^ pss_to_cert pss ^ ")"
+  | pss_to_cert (Sum (pss1, pss2)) = "(" ^ pss_to_cert pss1 ^ " + " ^ pss_to_cert pss2 ^ ")"
+  | pss_to_cert (Product (pss1, pss2)) = "(" ^ pss_to_cert pss1 ^ " * " ^ pss_to_cert pss2 ^ ")"
+
+fun pss_tree_to_cert Trivial = "()"
+  | pss_tree_to_cert (Cert pss) = "(" ^ pss_to_cert pss ^ ")"
+  | pss_tree_to_cert (Branch (t1, t2)) = "(" ^ pss_tree_to_cert t1 ^ " & " ^ pss_tree_to_cert t2 ^ ")"
+
+(*** certificate parsing ***)
+
+(* basic parser *)
+
+val str = Scan.this_string
+
+val number = Scan.repeat1 (Scan.one Symbol.is_ascii_digit >>
+  (fn s => ord s - ord "0")) >>
+  foldl1 (fn (n, d) => n * 10 + d)
+
+val nat = number
+val int = Scan.optional (str "~" >> K ~1) 1 -- nat >> op *;
+val rat = int --| str "/" -- int >> Rat.rat_of_quotient
+val rat_int = rat || int >> Rat.rat_of_int
+
+(* polynomial parser *)
+
+fun repeat_sep s f = f ::: Scan.repeat (str s |-- f)
+
+val parse_id = Scan.one Symbol.is_letter ::: Scan.many Symbol.is_letdig >> implode
+
+fun parse_varpow ctxt = parse_id -- Scan.optional (str "^" |-- nat) 1 >>
+  (fn (x, k) => (cterm_of (Context.theory_of_proof ctxt) (Free (x, @{typ real})), k)) 
+
+fun parse_monomial ctxt = repeat_sep "*" (parse_varpow ctxt) >>
+  foldl (uncurry Ctermfunc.update) Ctermfunc.undefined
+
+fun parse_cmonomial ctxt =
+  rat_int --| str "*" -- (parse_monomial ctxt) >> swap ||
+  (parse_monomial ctxt) >> (fn m => (m, Rat.one)) ||
+  rat_int >> (fn r => (Ctermfunc.undefined, r))
+
+fun parse_poly ctxt = repeat_sep "+" (parse_cmonomial ctxt) >>
+  foldl (uncurry Monomialfunc.update) Monomialfunc.undefined
+
+(* positivstellensatz parser *)
+
+val parse_axiom =
+  (str "A=" |-- int >> Axiom_eq) ||
+  (str "A<=" |-- int >> Axiom_le) ||
+  (str "A<" |-- int >> Axiom_lt)
+
+val parse_rational =
+  (str "R=" |-- rat_int >> Rational_eq) ||
+  (str "R<=" |-- rat_int >> Rational_le) ||
+  (str "R<" |-- rat_int >> Rational_lt)
+
+fun parse_cert ctxt input =
+  let
+    val pc = parse_cert ctxt
+    val pp = parse_poly ctxt
+  in
+  (parse_axiom ||
+   parse_rational ||
+   str "[" |-- pp --| str "]^2" >> Square ||
+   str "([" |-- pp --| str "]*" -- pc --| str ")" >> Eqmul ||
+   str "(" |-- pc --| str "*" -- pc --| str ")" >> Product ||
+   str "(" |-- pc --| str "+" -- pc --| str ")" >> Sum) input
+  end
+
+fun parse_cert_tree ctxt input =
+  let
+    val pc = parse_cert ctxt
+    val pt = parse_cert_tree ctxt
+  in
+  (str "()" >> K Trivial ||
+   str "(" |-- pc --| str ")" >> Cert ||
+   str "(" |-- pt --| str "&" -- pt --| str ")" >> Branch) input
+  end
+
+(* scanner *)
+
+fun cert_to_pss_tree ctxt input_str = Symbol.scanner "bad certificate" (parse_cert_tree ctxt)
+  (filter_out Symbol.is_blank (Symbol.explode input_str))
+
+end
+
+
--- a/src/HOL/Library/Sum_Of_Squares/sos_wrapper.ML	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Library/Sum_Of_Squares/sos_wrapper.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -136,13 +136,32 @@
     run_solver name (Path.explode cmd) find_failure
   end
 
+(* certificate output *)
+
+fun output_line cert = "To repeat this proof with a certifiate use this command:\n" ^
+        (Markup.markup Markup.sendback) ("by (sos_cert \"" ^ cert ^ "\")")
+
+val print_cert = Output.writeln o output_line o PositivstellensatzTools.pss_tree_to_cert
+
 (* setup tactic *)
 
-fun sos_solver name = (SIMPLE_METHOD' o (Sos.sos_tac (call_solver name))) 
+fun parse_cert (input as (ctxt, _)) = 
+  (Scan.lift OuterParse.string >>
+    PositivstellensatzTools.cert_to_pss_tree (Context.proof_of ctxt)) input
+
+fun sos_solver print method = (SIMPLE_METHOD' o (Sos.sos_tac print method)) 
 
-val sos_method = Scan.option (Scan.lift OuterParse.xname) >> sos_solver
+val sos_method =
+   Scan.lift (Scan.option OuterParse.xname) >> (fn n => Sos.Prover (call_solver n)) >>
+   sos_solver print_cert
 
-val setup = Method.setup @{binding sos} sos_method
-  "Prove universal problems over the reals using sums of squares"
+val sos_cert_method =
+  parse_cert >> Sos.Certificate >> sos_solver ignore
+
+val setup =
+     Method.setup @{binding sos} sos_method
+     "Prove universal problems over the reals using sums of squares"
+  #> Method.setup @{binding sos_cert} sos_cert_method
+     "Prove universal problems over the reals using sums of squares with certificates"
 
 end
--- a/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -8,7 +8,12 @@
 signature SOS =
 sig
 
-  val sos_tac : (string -> string) -> Proof.context -> int -> Tactical.tactic
+  datatype proof_method =
+    Certificate of RealArith.pss_tree
+  | Prover of (string -> string)
+
+  val sos_tac : (RealArith.pss_tree -> unit) ->
+    proof_method -> Proof.context -> int -> tactic
 
   val debugging : bool ref;
   
@@ -18,6 +23,8 @@
 structure Sos : SOS = 
 struct
 
+open FuncUtil;
+
 val rat_0 = Rat.zero;
 val rat_1 = Rat.one;
 val rat_2 = Rat.two;
@@ -59,6 +66,10 @@
 
 exception Failure of string;
 
+datatype proof_method =
+    Certificate of RealArith.pss_tree
+  | Prover of (string -> string)
+
 (* Turn a rational into a decimal string with d sig digits.                  *)
 
 local
@@ -93,23 +104,11 @@
 
 (* The main types.                                                           *)
 
-fun strict_ord ord (x,y) = case ord (x,y) of LESS => LESS | _ => GREATER
-
-structure Intpairfunc = FuncFun(type key = int*int val ord = prod_ord int_ord int_ord);
-
 type vector = int* Rat.rat Intfunc.T;
 
 type matrix = (int*int)*(Rat.rat Intpairfunc.T);
 
-type monomial = int Ctermfunc.T;
-
-val cterm_ord = (fn (s,t) => TermOrd.fast_term_ord(term_of s, term_of t))
- fun monomial_ord (m1,m2) = list_ord (prod_ord cterm_ord int_ord) (Ctermfunc.graph m1, Ctermfunc.graph m2)
-structure Monomialfunc = FuncFun(type key = monomial val ord = monomial_ord)
-
-type poly = Rat.rat Monomialfunc.T;
-
- fun iszero (k,r) = r =/ rat_0;
+fun iszero (k,r) = r =/ rat_0;
 
 fun fold_rev2 f l1 l2 b =
   case (l1,l2) of
@@ -346,11 +345,6 @@
   sort humanorder_varpow (Ctermfunc.graph m2))
 end;
 
-fun fold1 f l =  case l of
-   []     => error "fold1"
- | [x]    => x
- | (h::t) => f h (fold1 f t);
-
 (* Conversions to strings.                                                   *)
 
 fun string_of_vector min_size max_size (v:vector) =
@@ -359,7 +353,7 @@
   let 
    val n = max min_size (min n_raw max_size) 
    val xs = map (Rat.string_of_rat o (fn i => Intfunc.tryapplyd (snd v) i rat_0)) (1 upto n) 
-  in "[" ^ fold1 (fn s => fn t => s ^ ", " ^ t) xs ^
+  in "[" ^ foldr1 (fn (s, t) => s ^ ", " ^ t) xs ^
   (if n_raw > max_size then ", ...]" else "]")
   end
  end;
@@ -370,7 +364,7 @@
   val i = min max_size i_raw 
   val j = min max_size j_raw
   val rstr = map (fn k => string_of_vector j j (row k m)) (1 upto i) 
- in "["^ fold1 (fn s => fn t => s^";\n "^t) rstr ^
+ in "["^ foldr1 (fn (s, t) => s^";\n "^t) rstr ^
   (if j > max_size then "\n ...]" else "]")
  end;
 
@@ -396,7 +390,7 @@
  if Ctermfunc.is_undefined m then "1" else
  let val vps = fold_rev (fn (x,k) => fn a => string_of_varpow x k :: a)
   (sort humanorder_varpow (Ctermfunc.graph m)) [] 
- in fold1 (fn s => fn t => s^"*"^t) vps
+ in foldr1 (fn (s, t) => s^"*"^t) vps
  end;
 
 fun string_of_cmonomial (c,m) =
@@ -404,7 +398,7 @@
  else if c =/ rat_1 then string_of_monomial m
  else Rat.string_of_rat c ^ "*" ^ string_of_monomial m;;
 
-fun string_of_poly (p:poly) =
+fun string_of_poly p =
  if Monomialfunc.is_undefined p then "<<0>>" else
  let 
   val cms = sort (fn ((m1,_),(m2,_)) => humanorder_monomial m1  m2) (Monomialfunc.graph p)
@@ -478,10 +472,9 @@
  let 
   val n = dim v
   val strs = map (decimalize 20 o (fn i => Intfunc.tryapplyd (snd v) i rat_0)) (1 upto n) 
- in fold1 (fn x => fn y => x ^ " " ^ y) strs ^ "\n"
+ in foldr1 (fn (x, y) => x ^ " " ^ y) strs ^ "\n"
  end;
 
-fun increasing f ord (x,y) = ord (f x, f y);
 fun triple_int_ord ((a,b,c),(a',b',c')) = 
  prod_ord int_ord (prod_ord int_ord int_ord) 
     ((a,(b,c)),(a',(b',c')));
@@ -989,7 +982,7 @@
  let val alts =
   map (fn k => let val oths = enumerate_monomials (d - k) (tl vars) 
                in map (fn ks => if k = 0 then ks else Ctermfunc.update (hd vars, k) ks) oths end) (0 upto d) 
- in fold1 (curry op @) alts
+ in foldr1 op @ alts
  end;
 
 (* Enumerate products of distinct input polys with degree <= d.              *)
@@ -1040,7 +1033,7 @@
  in
   string_of_int m ^ "\n" ^
   string_of_int nblocks ^ "\n" ^
-  (fold1 (fn s => fn t => s^" "^t) (map string_of_int blocksizes)) ^
+  (foldr1 (fn (s, t) => s^" "^t) (map string_of_int blocksizes)) ^
   "\n" ^
   sdpa_of_vector obj ^
   fold_rev2 (fn k => fn m => fn a => sdpa_of_blockdiagonal (k - 1) m ^ a)
@@ -1080,11 +1073,6 @@
   fun tryfind f = tryfind_with "tryfind" f
 end
 
-(*
-fun tryfind f [] = error "tryfind"
-  | tryfind f (x::xs) = (f x handle ERROR _ => tryfind f xs);
-*)
-
 (* Positiv- and Nullstellensatz. Flag "linf" forces a linear representation. *)
 
  
@@ -1210,61 +1198,17 @@
 fun deepen f n = 
   (writeln ("Searching with depth limit " ^ string_of_int n) ; (f n handle Failure s => (writeln ("failed with message: " ^ s) ; deepen f (n+1))))
 
-(* The ordering so we can create canonical HOL polynomials.                  *)
 
-fun dest_monomial mon = sort (increasing fst cterm_ord) (Ctermfunc.graph mon);
-
-fun monomial_order (m1,m2) =
- if Ctermfunc.is_undefined m2 then LESS 
- else if Ctermfunc.is_undefined m1 then GREATER 
- else
-  let val mon1 = dest_monomial m1 
-      val mon2 = dest_monomial m2
-      val deg1 = fold (curry op + o snd) mon1 0
-      val deg2 = fold (curry op + o snd) mon2 0 
-  in if deg1 < deg2 then GREATER else if deg1 > deg2 then LESS
-     else list_ord (prod_ord cterm_ord int_ord) (mon1,mon2)
-  end;
-
-fun dest_poly p =
-  map (fn (m,c) => (c,dest_monomial m))
-      (sort (prod_ord monomial_order (K EQUAL)) (Monomialfunc.graph p));
-
-(* Map back polynomials and their composites to HOL.                         *)
+(* Map back polynomials and their composites to a positivstellensatz.        *)
 
 local
  open Thm Numeral RealArith
 in
 
-fun cterm_of_varpow x k = if k = 1 then x else capply (capply @{cterm "op ^ :: real => _"} x) 
-  (mk_cnumber @{ctyp nat} k)
-
-fun cterm_of_monomial m = 
- if Ctermfunc.is_undefined m then @{cterm "1::real"} 
- else 
-  let 
-   val m' = dest_monomial m
-   val vps = fold_rev (fn (x,k) => cons (cterm_of_varpow x k)) m' [] 
-  in fold1 (fn s => fn t => capply (capply @{cterm "op * :: real => _"} s) t) vps
-  end
-
-fun cterm_of_cmonomial (m,c) = if Ctermfunc.is_undefined m then cterm_of_rat c
-    else if c = Rat.one then cterm_of_monomial m
-    else capply (capply @{cterm "op *::real => _"} (cterm_of_rat c)) (cterm_of_monomial m);
-
-fun cterm_of_poly p = 
- if Monomialfunc.is_undefined p then @{cterm "0::real"} 
- else
-  let 
-   val cms = map cterm_of_cmonomial
-     (sort (prod_ord monomial_order (K EQUAL)) (Monomialfunc.graph p))
-  in fold1 (fn t1 => fn t2 => capply(capply @{cterm "op + :: real => _"} t1) t2) cms
-  end;
-
-fun cterm_of_sqterm (c,p) = Product(Rational_lt c,Square(cterm_of_poly p));
+fun cterm_of_sqterm (c,p) = Product(Rational_lt c,Square p);
 
 fun cterm_of_sos (pr,sqs) = if null sqs then pr
-  else Product(pr,fold1 (fn a => fn b => Sum(a,b)) (map cterm_of_sqterm sqs));
+  else Product(pr,foldr1 (fn (a, b) => Sum(a,b)) (map cterm_of_sqterm sqs));
 
 end
 
@@ -1275,14 +1219,14 @@
   fun simple_cterm_ord t u = TermOrd.fast_term_ord (term_of t, term_of u) = LESS
 in
   (* FIXME: Replace tryfind by get_first !! *)
-fun real_nonlinear_prover prover ctxt =
+fun real_nonlinear_prover proof_method ctxt =
  let 
   val {add,mul,neg,pow,sub,main} =  Normalizer.semiring_normalizers_ord_wrapper ctxt
       (valOf (NormalizerData.match ctxt @{cterm "(0::real) + 1"})) 
      simple_cterm_ord
   val (real_poly_add_conv,real_poly_mul_conv,real_poly_neg_conv,
        real_poly_pow_conv,real_poly_sub_conv,real_poly_conv) = (add,mul,neg,pow,sub,main)
-  fun mainf  translator (eqs,les,lts) = 
+  fun mainf cert_choice translator (eqs,les,lts) = 
   let 
    val eq0 = map (poly_of_term o dest_arg1 o concl) eqs
    val le0 = map (poly_of_term o dest_arg o concl) les
@@ -1303,33 +1247,49 @@
                      else raise Failure "trivial_axiom: Not a trivial axiom"
      | _ => error "trivial_axiom: Not a trivial axiom"
    in 
-  ((let val th = tryfind trivial_axiom (keq @ klep @ kltp)
-   in fconv_rule (arg_conv (arg1_conv real_poly_conv) then_conv field_comp_conv) th end)
-   handle Failure _ => (
-    let 
-     val pol = fold_rev poly_mul (map fst ltp) (poly_const Rat.one)
-     val leq = lep @ ltp
-     fun tryall d =
-      let val e = multidegree pol
-          val k = if e = 0 then 0 else d div e
-          val eq' = map fst eq 
-      in tryfind (fn i => (d,i,real_positivnullstellensatz_general prover false d eq' leq
-                            (poly_neg(poly_pow pol i))))
-              (0 upto k)
-      end
-    val (d,i,(cert_ideal,cert_cone)) = deepen tryall 0
-    val proofs_ideal =
-      map2 (fn q => fn (p,ax) => Eqmul(cterm_of_poly q,ax)) cert_ideal eq
-    val proofs_cone = map cterm_of_sos cert_cone
-    val proof_ne = if null ltp then Rational_lt Rat.one else
-      let val p = fold1 (fn s => fn t => Product(s,t)) (map snd ltp) 
-      in  funpow i (fn q => Product(p,q)) (Rational_lt Rat.one)
-      end
-    val proof = fold1 (fn s => fn t => Sum(s,t))
-                           (proof_ne :: proofs_ideal @ proofs_cone) 
-    in writeln "Translating proof certificate to HOL";
-       translator (eqs,les,lts) proof
-    end))
+  (let val th = tryfind trivial_axiom (keq @ klep @ kltp)
+   in
+    (fconv_rule (arg_conv (arg1_conv real_poly_conv) then_conv field_comp_conv) th, Trivial)
+   end)
+   handle Failure _ => 
+     (let val proof =
+       (case proof_method of Certificate certs =>
+         (* choose certificate *)
+         let
+           fun chose_cert [] (Cert c) = c
+             | chose_cert (Left::s) (Branch (l, _)) = chose_cert s l
+             | chose_cert (Right::s) (Branch (_, r)) = chose_cert s r
+             | chose_cert _ _ = error "certificate tree in invalid form"
+         in
+           chose_cert cert_choice certs
+         end
+       | Prover prover =>
+         (* call prover *)
+         let 
+          val pol = fold_rev poly_mul (map fst ltp) (poly_const Rat.one)
+          val leq = lep @ ltp
+          fun tryall d =
+           let val e = multidegree pol
+               val k = if e = 0 then 0 else d div e
+               val eq' = map fst eq 
+           in tryfind (fn i => (d,i,real_positivnullstellensatz_general prover false d eq' leq
+                                 (poly_neg(poly_pow pol i))))
+                   (0 upto k)
+           end
+         val (d,i,(cert_ideal,cert_cone)) = deepen tryall 0
+         val proofs_ideal =
+           map2 (fn q => fn (p,ax) => Eqmul(q,ax)) cert_ideal eq
+         val proofs_cone = map cterm_of_sos cert_cone
+         val proof_ne = if null ltp then Rational_lt Rat.one else
+           let val p = foldr1 (fn (s, t) => Product(s,t)) (map snd ltp) 
+           in  funpow i (fn q => Product(p,q)) (Rational_lt Rat.one)
+           end
+         in 
+           foldr1 (fn (s, t) => Sum(s,t)) (proof_ne :: proofs_ideal @ proofs_cone) 
+         end)
+     in
+        (translator (eqs,les,lts) proof, Cert proof)
+     end)
    end
  in mainf end
 end
@@ -1396,7 +1356,7 @@
          orelse g aconvc @{cterm "op < :: real => _"} 
        then arg_conv cv ct else arg1_conv cv ct
     end
-  fun mainf translator =
+  fun mainf cert_choice translator =
    let 
     fun substfirst(eqs,les,lts) =
       ((let 
@@ -1407,7 +1367,7 @@
                                    aconvc @{cterm "0::real"}) (map modify eqs),
                                    map modify les,map modify lts)
        end)
-       handle Failure  _ => real_nonlinear_prover prover ctxt translator (rev eqs, rev les, rev lts))
+       handle Failure  _ => real_nonlinear_prover prover ctxt cert_choice translator (rev eqs, rev les, rev lts))
     in substfirst
    end
 
@@ -1417,7 +1377,8 @@
 
 (* Overall function. *)
 
-fun real_sos prover ctxt t = gen_prover_real_arith ctxt (real_nonlinear_subst_prover prover ctxt) t;
+fun real_sos prover ctxt =
+  gen_prover_real_arith ctxt (real_nonlinear_subst_prover prover ctxt)
 end;
 
 (* A tactic *)
@@ -1429,8 +1390,6 @@
    end
 | _ => ([],ct)
 
-fun core_sos_conv prover ctxt t = Drule.arg_cong_rule @{cterm Trueprop} (real_sos prover ctxt (Thm.dest_arg t) RS @{thm Eq_TrueI})
-
 val known_sos_constants = 
   [@{term "op ==>"}, @{term "Trueprop"}, 
    @{term "op -->"}, @{term "op &"}, @{term "op |"}, 
@@ -1458,17 +1417,19 @@
   val _ = if exists (fn ((_,T)) => not (T = @{typ "real"})) fs 
           then error "SOS: not sos. Variables with type not real" else ()
   val vs = Term.add_vars t []
-  val _ = if exists (fn ((_,T)) => not (T = @{typ "real"})) fs 
+  val _ = if exists (fn ((_,T)) => not (T = @{typ "real"})) vs 
           then error "SOS: not sos. Variables with type not real" else ()
   val ukcs = subtract (fn (t,p) => Const p aconv t) kcts (Term.add_consts t [])
   val _ = if  null ukcs then () 
               else error ("SOSO: Unknown constants in Subgoal:" ^ commas (map fst ukcs))
 in () end
 
-fun core_sos_tac prover ctxt = CSUBGOAL (fn (ct, i) => 
+fun core_sos_tac print_cert prover ctxt = CSUBGOAL (fn (ct, i) => 
   let val _ = check_sos known_sos_constants ct
       val (avs, p) = strip_all ct
-      val th = standard (fold_rev forall_intr avs (real_sos prover ctxt (Thm.dest_arg p)))
+      val (ths, certificates) = real_sos prover ctxt (Thm.dest_arg p)
+      val th = standard (fold_rev forall_intr avs ths)
+      val _ = print_cert certificates
   in rtac th i end);
 
 fun default_SOME f NONE v = SOME v
@@ -1506,7 +1467,7 @@
 
 fun elim_denom_tac ctxt i = REPEAT (elim_one_denom_tac ctxt i);
 
-fun sos_tac prover ctxt = ObjectLogic.full_atomize_tac THEN' elim_denom_tac ctxt THEN' core_sos_tac prover ctxt
+fun sos_tac print_cert prover ctxt = ObjectLogic.full_atomize_tac THEN' elim_denom_tac ctxt THEN' core_sos_tac print_cert prover ctxt
 
 
 end;
--- a/src/HOL/Library/normarith.ML	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Library/normarith.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -15,7 +15,7 @@
 structure NormArith : NORM_ARITH = 
 struct
 
- open Conv Thm;
+ open Conv Thm FuncUtil;
  val bool_eq = op = : bool *bool -> bool
   fun dest_ratconst t = case term_of t of
    Const(@{const_name divide}, _)$a$b => Rat.rat_of_quotient(HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd)
@@ -330,13 +330,13 @@
    val zerodests = filter
         (fn t => null (Ctermfunc.dom (vector_lincomb t))) (map snd rawdests)
 
-  in RealArith.real_linear_prover translator
+  in fst (RealArith.real_linear_prover translator
         (map (fn t => instantiate ([(tv_n, ctyp_of_term t)],[]) pth_zero)
             zerodests,
         map (fconv_rule (try_conv (More_Conv.top_sweep_conv (K norm_canon_conv) ctxt) then_conv
                        arg_conv (arg_conv real_poly_conv))) ges',
         map (fconv_rule (try_conv (More_Conv.top_sweep_conv (K norm_canon_conv) ctxt) then_conv 
-                       arg_conv (arg_conv real_poly_conv))) gts)
+                       arg_conv (arg_conv real_poly_conv))) gts))
   end
 in val real_vector_combo_prover = real_vector_combo_prover
 end;
@@ -389,9 +389,9 @@
    val (th1,th2) = conj_pair(rawrule th)
   in th1::fconv_rule (arg_conv (arg_conv real_poly_neg_conv)) th2::acc
   end
-in fun real_vector_prover ctxt translator (eqs,ges,gts) =
-     real_vector_ineq_prover ctxt translator
-         (fold_rev (splitequation ctxt) eqs ges,gts)
+in fun real_vector_prover ctxt _ translator (eqs,ges,gts) =
+     (real_vector_ineq_prover ctxt translator
+         (fold_rev (splitequation ctxt) eqs ges,gts), RealArith.Trivial)
 end;
 
   fun init_conv ctxt = 
@@ -400,7 +400,7 @@
    then_conv field_comp_conv 
    then_conv nnf_conv
 
- fun pure ctxt = RealArith.gen_prover_real_arith ctxt (real_vector_prover ctxt);
+ fun pure ctxt = fst o RealArith.gen_prover_real_arith ctxt (real_vector_prover ctxt);
  fun norm_arith ctxt ct = 
   let 
    val ctxt' = Variable.declare_term (term_of ct) ctxt
--- a/src/HOL/Library/positivstellensatz.ML	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Library/positivstellensatz.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -1,10 +1,11 @@
-(* Title:      Library/positivstellensatz
+(* Title:      Library/Sum_Of_Squares/positivstellensatz
    Author:     Amine Chaieb, University of Cambridge
    Description: A generic arithmetic prover based on Positivstellensatz certificates --- 
     also implements Fourrier-Motzkin elimination as a special case Fourrier-Motzkin elimination.
 *)
 
 (* A functor for finite mappings based on Tables *)
+
 signature FUNC = 
 sig
  type 'a T
@@ -75,27 +76,54 @@
 end
 end;
 
+(* Some standard functors and utility functions for them *)
+
+structure FuncUtil =
+struct
+
+fun increasing f ord (x,y) = ord (f x, f y);
+
 structure Intfunc = FuncFun(type key = int val ord = int_ord);
+structure Ratfunc = FuncFun(type key = Rat.rat val ord = Rat.ord);
+structure Intpairfunc = FuncFun(type key = int*int val ord = prod_ord int_ord int_ord);
 structure Symfunc = FuncFun(type key = string val ord = fast_string_ord);
 structure Termfunc = FuncFun(type key = term val ord = TermOrd.fast_term_ord);
-structure Ctermfunc = FuncFun(type key = cterm val ord = (fn (s,t) => TermOrd.fast_term_ord(term_of s, term_of t)));
+
+val cterm_ord = (fn (s,t) => TermOrd.fast_term_ord(term_of s, term_of t))
+
+structure Ctermfunc = FuncFun(type key = cterm val ord = cterm_ord);
+
+type monomial = int Ctermfunc.T;
 
-structure Ratfunc = FuncFun(type key = Rat.rat val ord = Rat.ord);
+fun monomial_ord (m1,m2) = list_ord (prod_ord cterm_ord int_ord) (Ctermfunc.graph m1, Ctermfunc.graph m2)
+
+structure Monomialfunc = FuncFun(type key = monomial val ord = monomial_ord)
 
+type poly = Rat.rat Monomialfunc.T;
+
+(* The ordering so we can create canonical HOL polynomials.                  *)
 
-    (* Some useful derived rules *)
-fun deduct_antisym_rule tha thb = 
-    equal_intr (implies_intr (cprop_of thb) tha) 
-     (implies_intr (cprop_of tha) thb);
+fun dest_monomial mon = sort (increasing fst cterm_ord) (Ctermfunc.graph mon);
 
-fun prove_hyp tha thb = 
-  if exists (curry op aconv (concl_of tha)) (#hyps (rep_thm thb)) 
-  then equal_elim (symmetric (deduct_antisym_rule tha thb)) tha else thb;
+fun monomial_order (m1,m2) =
+ if Ctermfunc.is_undefined m2 then LESS 
+ else if Ctermfunc.is_undefined m1 then GREATER 
+ else
+  let val mon1 = dest_monomial m1 
+      val mon2 = dest_monomial m2
+      val deg1 = fold (curry op + o snd) mon1 0
+      val deg2 = fold (curry op + o snd) mon2 0 
+  in if deg1 < deg2 then GREATER else if deg1 > deg2 then LESS
+     else list_ord (prod_ord cterm_ord int_ord) (mon1,mon2)
+  end;
 
+end
 
+(* positivstellensatz datatype and prover generation *)
 
 signature REAL_ARITH = 
 sig
+  
   datatype positivstellensatz =
    Axiom_eq of int
  | Axiom_le of int
@@ -103,34 +131,41 @@
  | Rational_eq of Rat.rat
  | Rational_le of Rat.rat
  | Rational_lt of Rat.rat
- | Square of cterm
- | Eqmul of cterm * positivstellensatz
+ | Square of FuncUtil.poly
+ | Eqmul of FuncUtil.poly * positivstellensatz
  | Sum of positivstellensatz * positivstellensatz
  | Product of positivstellensatz * positivstellensatz;
 
+datatype pss_tree = Trivial | Cert of positivstellensatz | Branch of pss_tree * pss_tree
+
+datatype tree_choice = Left | Right
+
+type prover = tree_choice list -> 
+  (thm list * thm list * thm list -> positivstellensatz -> thm) ->
+  thm list * thm list * thm list -> thm * pss_tree
+type cert_conv = cterm -> thm * pss_tree
+
 val gen_gen_real_arith :
-  Proof.context -> (Rat.rat -> Thm.cterm) * conv * conv * conv * 
-   conv * conv * conv * conv * conv * conv * 
-    ( (thm list * thm list * thm list -> positivstellensatz -> thm) ->
-        thm list * thm list * thm list -> thm) -> conv
-val real_linear_prover : 
-  (thm list * thm list * thm list -> positivstellensatz -> thm) ->
-   thm list * thm list * thm list -> thm
+  Proof.context -> (Rat.rat -> cterm) * conv * conv * conv *
+   conv * conv * conv * conv * conv * conv * prover -> cert_conv
+val real_linear_prover : (thm list * thm list * thm list -> positivstellensatz -> thm) ->
+  thm list * thm list * thm list -> thm * pss_tree
 
 val gen_real_arith : Proof.context ->
-   (Rat.rat -> cterm) * conv * conv * conv * conv * conv * conv * conv *
-   ( (thm list * thm list * thm list -> positivstellensatz -> thm) ->
-       thm list * thm list * thm list -> thm) -> conv
-val gen_prover_real_arith : Proof.context ->
-   ((thm list * thm list * thm list -> positivstellensatz -> thm) ->
-     thm list * thm list * thm list -> thm) -> conv
-val real_arith : Proof.context -> conv
+  (Rat.rat -> cterm) * conv * conv * conv * conv * conv * conv * conv * prover -> cert_conv
+
+val gen_prover_real_arith : Proof.context -> prover -> cert_conv
+
+val is_ratconst : cterm -> bool
+val dest_ratconst : cterm -> Rat.rat
+val cterm_of_rat : Rat.rat -> cterm
+
 end
 
-structure RealArith (* : REAL_ARITH *)=
+structure RealArith : REAL_ARITH =
 struct
 
- open Conv Thm;;
+ open Conv Thm FuncUtil;;
 (* ------------------------------------------------------------------------- *)
 (* Data structure for Positivstellensatz refutations.                        *)
 (* ------------------------------------------------------------------------- *)
@@ -142,12 +177,18 @@
  | Rational_eq of Rat.rat
  | Rational_le of Rat.rat
  | Rational_lt of Rat.rat
- | Square of cterm
- | Eqmul of cterm * positivstellensatz
+ | Square of FuncUtil.poly
+ | Eqmul of FuncUtil.poly * positivstellensatz
  | Sum of positivstellensatz * positivstellensatz
  | Product of positivstellensatz * positivstellensatz;
          (* Theorems used in the procedure *)
 
+datatype pss_tree = Trivial | Cert of positivstellensatz | Branch of pss_tree * pss_tree
+datatype tree_choice = Left | Right
+type prover = tree_choice list -> 
+  (thm list * thm list * thm list -> positivstellensatz -> thm) ->
+  thm list * thm list * thm list -> thm * pss_tree
+type cert_conv = cterm -> thm * pss_tree
 
 val my_eqs = ref ([] : thm list);
 val my_les = ref ([] : thm list);
@@ -164,6 +205,16 @@
 val my_poly_add_conv = ref no_conv;
 val my_poly_mul_conv = ref no_conv;
 
+
+    (* Some useful derived rules *)
+fun deduct_antisym_rule tha thb = 
+    equal_intr (implies_intr (cprop_of thb) tha) 
+     (implies_intr (cprop_of tha) thb);
+
+fun prove_hyp tha thb = 
+  if exists (curry op aconv (concl_of tha)) (#hyps (rep_thm thb)) 
+  then equal_elim (symmetric (deduct_antisym_rule tha thb)) tha else thb;
+
 fun conjunctions th = case try Conjunction.elim th of
    SOME (th1,th2) => (conjunctions th1) @ conjunctions th2
  | NONE => [th];
@@ -292,7 +343,6 @@
  | Abs (_,_,t') => find_cterm p (Thm.dest_abs NONE t |> snd)
  | _ => raise CTERM ("find_cterm",[t]);
 
-
     (* Some conversions-related stuff which has been forbidden entrance into Pure/conv.ML*)
 fun instantiate_cterm' ty tms = Drule.cterm_rule (Drule.instantiate' ty tms)
 fun is_comb t = case (term_of t) of _$_ => true | _ => false;
@@ -300,6 +350,39 @@
 fun is_binop ct ct' = ct aconvc (Thm.dest_fun (Thm.dest_fun ct'))
   handle CTERM _ => false;
 
+
+(* Map back polynomials to HOL.                         *)
+
+local
+ open Thm Numeral
+in
+
+fun cterm_of_varpow x k = if k = 1 then x else capply (capply @{cterm "op ^ :: real => _"} x) 
+  (mk_cnumber @{ctyp nat} k)
+
+fun cterm_of_monomial m = 
+ if Ctermfunc.is_undefined m then @{cterm "1::real"} 
+ else 
+  let 
+   val m' = dest_monomial m
+   val vps = fold_rev (fn (x,k) => cons (cterm_of_varpow x k)) m' [] 
+  in foldr1 (fn (s, t) => capply (capply @{cterm "op * :: real => _"} s) t) vps
+  end
+
+fun cterm_of_cmonomial (m,c) = if Ctermfunc.is_undefined m then cterm_of_rat c
+    else if c = Rat.one then cterm_of_monomial m
+    else capply (capply @{cterm "op *::real => _"} (cterm_of_rat c)) (cterm_of_monomial m);
+
+fun cterm_of_poly p = 
+ if Monomialfunc.is_undefined p then @{cterm "0::real"} 
+ else
+  let 
+   val cms = map cterm_of_cmonomial
+     (sort (prod_ord monomial_order (K EQUAL)) (Monomialfunc.graph p))
+  in foldr1 (fn (t1, t2) => capply(capply @{cterm "op + :: real => _"} t1) t2) cms
+  end;
+
+end;
     (* A general real arithmetic prover *)
 
 fun gen_gen_real_arith ctxt (mk_numeric,
@@ -368,8 +451,8 @@
       | Rational_lt x => eqT_elim(numeric_gt_conv(capply @{cterm Trueprop} 
                       (capply (capply @{cterm "op <::real => _"} @{cterm "0::real"})
                         (mk_numeric x))))
-      | Square t => square_rule t
-      | Eqmul(t,p) => emul_rule t (translate p)
+      | Square pt => square_rule (cterm_of_poly pt)
+      | Eqmul(pt,p) => emul_rule (cterm_of_poly pt) (translate p)
       | Sum(p1,p2) => add_rule (translate p1) (translate p2)
       | Product(p1,p2) => mul_rule (translate p1) (translate p2)
    in fconv_rule (first_conv [numeric_ge_conv, numeric_gt_conv, numeric_eq_conv, all_conv]) 
@@ -394,13 +477,13 @@
        val _ = if c aconvc (concl th2) then () else error "disj_cases : conclusions not alpha convertible"
    in implies_elim (implies_elim (implies_elim (instantiate' [] (map SOME [p,q,c]) @{thm disjE}) th) (implies_intr (capply @{cterm Trueprop} p) th1)) (implies_intr (capply @{cterm Trueprop} q) th2)
    end
- fun overall dun ths = case ths of
+ fun overall cert_choice dun ths = case ths of
   [] =>
    let 
     val (eq,ne) = List.partition (is_req o concl) dun
      val (le,nl) = List.partition (is_ge o concl) ne
      val lt = filter (is_gt o concl) nl 
-    in prover hol_of_positivstellensatz (eq,le,lt) end
+    in prover (rev cert_choice) hol_of_positivstellensatz (eq,le,lt) end
  | th::oths =>
    let 
     val ct = concl th 
@@ -408,13 +491,13 @@
     if is_conj ct  then
      let 
       val (th1,th2) = conj_pair th in
-      overall dun (th1::th2::oths) end
+      overall cert_choice dun (th1::th2::oths) end
     else if is_disj ct then
       let 
-       val th1 = overall dun (assume (capply @{cterm Trueprop} (dest_arg1 ct))::oths)
-       val th2 = overall dun (assume (capply @{cterm Trueprop} (dest_arg ct))::oths)
-      in disj_cases th th1 th2 end
-   else overall (th::dun) oths
+       val (th1, cert1) = overall (Left::cert_choice) dun (assume (capply @{cterm Trueprop} (dest_arg1 ct))::oths)
+       val (th2, cert2) = overall (Right::cert_choice) dun (assume (capply @{cterm Trueprop} (dest_arg ct))::oths)
+      in (disj_cases th th1 th2, Branch (cert1, cert2)) end
+   else overall cert_choice (th::dun) oths
   end
   fun dest_binary b ct = if is_binop b ct then dest_binop ct 
                          else raise CTERM ("dest_binary",[b,ct])
@@ -496,16 +579,16 @@
   val nct = capply @{cterm Trueprop} (capply @{cterm "Not"} ct)
   val th0 = (init_conv then_conv arg_conv nnf_norm_conv') nct
   val tm0 = dest_arg (rhs_of th0)
-  val th = if tm0 aconvc @{cterm False} then equal_implies_1_rule th0 else
+  val (th, certificates) = if tm0 aconvc @{cterm False} then (equal_implies_1_rule th0, Trivial) else
    let 
     val (evs,bod) = strip_exists tm0
     val (avs,ibod) = strip_forall bod
     val th1 = Drule.arg_cong_rule @{cterm Trueprop} (fold mk_forall avs (absremover ibod))
-    val th2 = overall [] [specl avs (assume (rhs_of th1))]
+    val (th2, certs) = overall [] [] [specl avs (assume (rhs_of th1))]
     val th3 = fold simple_choose evs (prove_hyp (equal_elim th1 (assume (capply @{cterm Trueprop} bod))) th2)
-   in  Drule.implies_intr_hyps (prove_hyp (equal_elim th0 (assume nct)) th3)
+   in (Drule.implies_intr_hyps (prove_hyp (equal_elim th0 (assume nct)) th3), certs)
    end
-  in implies_elim (instantiate' [] [SOME ct] pth_final) th
+  in (implies_elim (instantiate' [] [SOME ct] pth_final) th, certificates)
  end
 in f
 end;
@@ -580,7 +663,7 @@
          val k = (Rat.neg d) */ Rat.abs c // c
          val e' = linear_cmul k e
          val t' = linear_cmul (Rat.abs c) t
-         val p' = Eqmul(cterm_of_rat k,p)
+         val p' = Eqmul(Monomialfunc.onefunc (Ctermfunc.undefined, k),p)
          val q' = Product(Rational_lt(Rat.abs c),q) 
         in (linear_add e' t',Sum(p',q')) 
         end 
@@ -632,7 +715,7 @@
   val le_pols' = le_pols @ map (fn v => Ctermfunc.onefunc (v,Rat.one)) aliens
   val (_,proof) = linear_prover (eq_pols,le_pols',lt_pols)
   val le' = le @ map (fn a => instantiate' [] [SOME (dest_arg a)] @{thm real_of_nat_ge_zero}) aliens 
- in (translator (eq,le',lt) proof) : thm
+ in ((translator (eq,le',lt) proof), Trivial)
  end
 end;
 
@@ -698,5 +781,4 @@
     main,neg,add,mul, prover)
 end;
 
-fun real_arith ctxt = gen_prover_real_arith ctxt real_linear_prover;
 end
--- a/src/HOL/Lim.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Lim.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -84,6 +84,8 @@
 lemma LIM_const [simp]: "(%x. k) -- x --> k"
 by (simp add: LIM_def)
 
+lemma LIM_cong_limit: "\<lbrakk> f -- x --> L ; K = L \<rbrakk> \<Longrightarrow> f -- x --> K" by simp
+
 lemma LIM_add:
   fixes f g :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
   assumes f: "f -- a --> L" and g: "g -- a --> M"
@@ -544,7 +546,7 @@
     case True thus ?thesis using `0 < s` by auto
   next
     case False hence "s / 2 \<ge> (x - b) / 2" by auto
-    hence "?x = (x + b) / 2" by(simp add:field_simps)
+    hence "?x = (x + b) / 2" by (simp add: field_simps min_max.inf_absorb2)
     thus ?thesis using `b < x` by auto
   qed
   hence "0 \<le> f ?x" using all_le `?x < x` by auto
--- a/src/HOL/MicroJava/BV/Effect.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/MicroJava/BV/Effect.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MicroJava/BV/Effect.thy
-    ID:         $Id$
     Author:     Gerwin Klein
     Copyright   2000 Technische Universitaet Muenchen
 *)
@@ -391,7 +390,7 @@
   with Pair 
   have "?app s \<Longrightarrow> ?P s" by (simp only:)
   moreover
-  have "?P s \<Longrightarrow> ?app s" by (unfold app_def) (clarsimp)
+  have "?P s \<Longrightarrow> ?app s" by (clarsimp simp add: min_max.inf_absorb2)
   ultimately
   show ?thesis by (rule iffI) 
 qed 
--- a/src/HOL/MicroJava/BV/LBVSpec.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/MicroJava/BV/LBVSpec.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MicroJava/BV/LBVSpec.thy
-    ID:         $Id$
     Author:     Gerwin Klein
     Copyright   1999 Technische Universitaet Muenchen
 *)
@@ -293,7 +292,7 @@
   shows "wtl (take (pc+1) is) c 0 s = wtc c pc (wtl (take pc is) c 0 s)"
 proof -
   from suc have "take (pc+1) is=(take pc is)@[is!pc]" by (simp add: take_Suc)
-  with suc wtl show ?thesis by (simp)
+  with suc wtl show ?thesis by (simp add: min_max.inf_absorb2)
 qed
 
 lemma (in lbv) wtl_all:
@@ -308,7 +307,7 @@
   with all have take: "?wtl (take pc is@i#r) \<noteq> \<top>" by simp 
   from pc have "is!pc = drop pc is ! 0" by simp
   with Cons have "is!pc = i" by simp
-  with take pc show ?thesis by (auto)
+  with take pc show ?thesis by (auto simp add: min_max.inf_absorb2)
 qed
 
 section "preserves-type"
--- a/src/HOL/MicroJava/Comp/CorrCompTp.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/MicroJava/Comp/CorrCompTp.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MicroJava/Comp/CorrCompTp.thy
-    ID:         $Id$
     Author:     Martin Strecker
 *)
 
@@ -1058,7 +1057,7 @@
 lemma bc_mt_corresp_New: "\<lbrakk>is_class cG cname \<rbrakk>
   \<Longrightarrow> bc_mt_corresp [New cname] (pushST [Class cname]) (ST, LT) cG rT mxr (Suc 0)"
 apply (simp add: bc_mt_corresp_def pushST_def wt_instr_altern_def
-    max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def)
+    max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def min_max.sup_absorb2)
 apply (intro strip)
 apply (rule conjI)
 apply (rule check_type_mono, assumption, simp)
@@ -1069,7 +1068,7 @@
   bc_mt_corresp [Pop] (popST (Suc 0)) (T # ST, LT) cG rT mxr (Suc 0)"
   apply (simp add: bc_mt_corresp_def popST_def wt_instr_altern_def eff_def norm_eff_def)
   apply (simp add: max_ssize_def ssize_sto_def max_of_list_def) 
-  apply (simp add: check_type_simps)
+  apply (simp add: check_type_simps min_max.sup_absorb1)
   apply clarify
   apply (rule_tac x="(length ST)" in exI)
   apply simp+
@@ -1092,7 +1091,7 @@
   \<Longrightarrow> bc_mt_corresp [LitPush val] (pushST [T]) sttp cG rT mxr (Suc 0)"
 apply (subgoal_tac "\<exists> ST LT. sttp= (ST, LT)", (erule exE)+)
   apply (simp add: bc_mt_corresp_def pushST_def wt_instr_altern_def
-              max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def)
+              max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def min_max.sup_absorb2)
   apply (intro strip)
   apply (rule conjI)
   apply (rule check_type_mono, assumption, simp)
@@ -1111,7 +1110,7 @@
   \<Longrightarrow> bc_mt_corresp [LitPush val] (pushST [T']) sttp cG rT mxr (Suc 0)"
 apply (subgoal_tac "\<exists> ST LT. sttp= (ST, LT)", (erule exE)+)
   apply (simp add: bc_mt_corresp_def pushST_def wt_instr_altern_def
-              max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def)
+              max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def min_max.sup_absorb2)
   apply (intro strip)
   apply (rule conjI)
   apply (rule check_type_mono, assumption, simp)
@@ -1127,7 +1126,7 @@
   \<Longrightarrow> bc_mt_corresp [Load i] 
          (\<lambda>(ST, LT). pushST [ok_val (LT ! i)] (ST, LT)) (ST, LT) cG rT mxr (Suc 0)"
 apply (simp add: bc_mt_corresp_def pushST_def wt_instr_altern_def
-            max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def)
+            max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def min_max.sup_absorb2)
   apply (intro strip)
   apply (rule conjI)
   apply (rule check_type_mono, assumption, simp)
@@ -1148,10 +1147,10 @@
 lemma bc_mt_corresp_Store_init: "\<lbrakk> i < length LT \<rbrakk>
   \<Longrightarrow> bc_mt_corresp [Store i] (storeST i T) (T # ST, LT) cG rT mxr (Suc 0)"
  apply (simp add: bc_mt_corresp_def storeST_def wt_instr_altern_def eff_def norm_eff_def)
-  apply (simp add: max_ssize_def  max_of_list_def )
+  apply (simp add: max_ssize_def  max_of_list_def)
   apply (simp add: ssize_sto_def)
   apply (intro strip)
-apply (simp add: check_type_simps)
+apply (simp add: check_type_simps min_max.sup_absorb1)
 apply clarify
 apply (rule conjI)
 apply (rule_tac x="(length ST)" in exI)
@@ -1159,14 +1158,13 @@
 done
 
 
-
 lemma bc_mt_corresp_Store: "\<lbrakk> i < length LT; cG  \<turnstile>  LT[i := OK T] <=l LT \<rbrakk>
   \<Longrightarrow> bc_mt_corresp [Store i] (popST (Suc 0)) (T # ST, LT) cG rT mxr (Suc 0)"
   apply (simp add: bc_mt_corresp_def popST_def wt_instr_altern_def eff_def norm_eff_def)
   apply (simp add: sup_state_conv)
   apply (simp add: max_ssize_def max_of_list_def ssize_sto_def)
  apply (intro strip)
-apply (simp add: check_type_simps)
+apply (simp add: check_type_simps min_max.sup_absorb1)
 apply clarify
 apply (rule_tac x="(length ST)" in exI)
 apply simp+
@@ -1176,7 +1174,7 @@
 lemma bc_mt_corresp_Dup: "
   bc_mt_corresp [Dup] dupST (T # ST, LT) cG rT mxr (Suc 0)"
  apply (simp add: bc_mt_corresp_def dupST_def wt_instr_altern_def
-             max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def)
+             max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def min_max.sup_absorb2)
   apply (intro strip)
   apply (rule conjI)
   apply (rule check_type_mono, assumption, simp)
@@ -1189,7 +1187,7 @@
 lemma bc_mt_corresp_Dup_x1: "
   bc_mt_corresp [Dup_x1] dup_x1ST (T1 # T2 # ST, LT) cG rT mxr (Suc 0)"
   apply (simp add: bc_mt_corresp_def dup_x1ST_def wt_instr_altern_def
-              max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def)
+              max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def min_max.sup_absorb2)
   apply (intro strip)
   apply (rule conjI)
   apply (rule check_type_mono, assumption, simp)
@@ -1206,7 +1204,7 @@
          (PrimT Integer # PrimT Integer # ST, LT) cG rT mxr (Suc 0)"
   apply (simp add: bc_mt_corresp_def replST_def wt_instr_altern_def eff_def norm_eff_def)
   apply (simp add: max_ssize_def max_of_list_def ssize_sto_def)
-  apply (simp add: check_type_simps)
+  apply (simp add: check_type_simps min_max.sup_absorb1)
   apply clarify
   apply (rule_tac x="Suc (length ST)" in exI)
   apply simp+
@@ -1249,7 +1247,7 @@
   apply (simp add: max_ssize_def max_of_list_def ssize_sto_def)
 
   apply (intro strip)
-apply (simp add: check_type_simps)
+apply (simp add: check_type_simps min_max.sup_absorb1)
 apply clarify
 apply (rule_tac x="Suc (length ST)" in exI)
 apply simp+
--- a/src/HOL/OrderedGroup.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/OrderedGroup.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -1075,17 +1075,16 @@
 lemma nprt_0[simp]: "nprt 0 = 0" by (simp add: nprt_def)
 
 lemma pprt_eq_id [simp, noatp]: "0 \<le> x \<Longrightarrow> pprt x = x"
-by (simp add: pprt_def sup_aci)
-
+  by (simp add: pprt_def sup_aci sup_absorb1)
 
 lemma nprt_eq_id [simp, noatp]: "x \<le> 0 \<Longrightarrow> nprt x = x"
-by (simp add: nprt_def inf_aci)
+  by (simp add: nprt_def inf_aci inf_absorb1)
 
 lemma pprt_eq_0 [simp, noatp]: "x \<le> 0 \<Longrightarrow> pprt x = 0"
-by (simp add: pprt_def sup_aci)
+  by (simp add: pprt_def sup_aci sup_absorb2)
 
 lemma nprt_eq_0 [simp, noatp]: "0 \<le> x \<Longrightarrow> nprt x = 0"
-by (simp add: nprt_def inf_aci)
+  by (simp add: nprt_def inf_aci inf_absorb2)
 
 lemma sup_0_imp_0: "sup a (- a) = 0 \<Longrightarrow> a = 0"
 proof -
@@ -1119,7 +1118,7 @@
   "0 \<le> a + a \<longleftrightarrow> 0 \<le> a"
 proof
   assume "0 <= a + a"
-  hence a:"inf (a+a) 0 = 0" by (simp add: inf_commute)
+  hence a:"inf (a+a) 0 = 0" by (simp add: inf_commute inf_absorb1)
   have "(inf a 0)+(inf a 0) = inf (inf (a+a) 0) a" (is "?l=_")
     by (simp add: add_sup_inf_distribs inf_aci)
   hence "?l = 0 + inf a 0" by (simp add: a, simp add: inf_commute)
@@ -1135,7 +1134,7 @@
   assume assm: "a + a = 0"
   then have "a + a + - a = - a" by simp
   then have "a + (a + - a) = - a" by (simp only: add_assoc)
-  then have a: "- a = a" by simp (*FIXME tune proof*)
+  then have a: "- a = a" by simp
   show "a = 0" apply (rule antisym)
   apply (unfold neg_le_iff_le [symmetric, of a])
   unfolding a apply simp
@@ -1275,7 +1274,7 @@
 proof -
   note add_le_cancel_right [of a a "- a", symmetric, simplified]
   moreover note add_le_cancel_right [of "-a" a a, symmetric, simplified]
-  then show ?thesis by (auto simp: sup_max)
+  then show ?thesis by (auto simp: sup_max min_max.sup_absorb1 min_max.sup_absorb2)
 qed
 
 lemma abs_if_lattice:
--- a/src/HOL/Predicate.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Predicate.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -797,6 +797,10 @@
    (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
+
 ML {*
 signature PREDICATE =
 sig
--- a/src/HOL/Quickcheck.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Quickcheck.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -3,7 +3,7 @@
 header {* A simple counterexample generator *}
 
 theory Quickcheck
-imports Random Code_Eval
+imports Random Code_Evaluation
 uses ("Tools/quickcheck_generators.ML")
 begin
 
@@ -24,7 +24,7 @@
 
 definition
   "random i = Random.range 2 o\<rightarrow>
-    (\<lambda>k. Pair (if k = 0 then Code_Eval.valtermify False else Code_Eval.valtermify True))"
+    (\<lambda>k. Pair (if k = 0 then Code_Evaluation.valtermify False else Code_Evaluation.valtermify True))"
 
 instance ..
 
@@ -34,7 +34,7 @@
 begin
 
 definition random_itself :: "code_numeral \<Rightarrow> Random.seed \<Rightarrow> ('a itself \<times> (unit \<Rightarrow> term)) \<times> Random.seed" where
-  "random_itself _ = Pair (Code_Eval.valtermify TYPE('a))"
+  "random_itself _ = Pair (Code_Evaluation.valtermify TYPE('a))"
 
 instance ..
 
@@ -44,7 +44,7 @@
 begin
 
 definition
-  "random _ = Random.select chars o\<rightarrow> (\<lambda>c. Pair (c, \<lambda>u. Code_Eval.term_of c))"
+  "random _ = Random.select chars o\<rightarrow> (\<lambda>c. Pair (c, \<lambda>u. Code_Evaluation.term_of c))"
 
 instance ..
 
@@ -54,7 +54,7 @@
 begin
 
 definition 
-  "random _ = Pair (STR '''', \<lambda>u. Code_Eval.term_of (STR ''''))"
+  "random _ = Pair (STR '''', \<lambda>u. Code_Evaluation.term_of (STR ''''))"
 
 instance ..
 
@@ -63,10 +63,10 @@
 instantiation nat :: random
 begin
 
-definition random_nat :: "code_numeral \<Rightarrow> Random.seed \<Rightarrow> (nat \<times> (unit \<Rightarrow> Code_Eval.term)) \<times> Random.seed" where
+definition random_nat :: "code_numeral \<Rightarrow> Random.seed \<Rightarrow> (nat \<times> (unit \<Rightarrow> Code_Evaluation.term)) \<times> Random.seed" where
   "random_nat i = Random.range (i + 1) o\<rightarrow> (\<lambda>k. Pair (
      let n = Code_Numeral.nat_of k
-     in (n, \<lambda>_. Code_Eval.term_of n)))"
+     in (n, \<lambda>_. Code_Evaluation.term_of n)))"
 
 instance ..
 
@@ -78,7 +78,7 @@
 definition
   "random i = Random.range (2 * i + 1) o\<rightarrow> (\<lambda>k. Pair (
      let j = (if k \<ge> i then Code_Numeral.int_of (k - i) else - Code_Numeral.int_of (i - k))
-     in (j, \<lambda>_. Code_Eval.term_of j)))"
+     in (j, \<lambda>_. Code_Evaluation.term_of j)))"
 
 instance ..
 
@@ -95,7 +95,7 @@
 
 definition random_fun_lift :: "(Random.seed \<Rightarrow> ('b \<times> (unit \<Rightarrow> term)) \<times> Random.seed)
   \<Rightarrow> Random.seed \<Rightarrow> (('a\<Colon>term_of \<Rightarrow> 'b\<Colon>typerep) \<times> (unit \<Rightarrow> term)) \<times> Random.seed" where
-  "random_fun_lift f = random_fun_aux TYPEREP('a) TYPEREP('b) (op =) Code_Eval.term_of f Random.split_seed"
+  "random_fun_lift f = random_fun_aux TYPEREP('a) TYPEREP('b) (op =) Code_Evaluation.term_of f Random.split_seed"
 
 instantiation "fun" :: ("{eq, term_of}", random) random
 begin
--- a/src/HOL/Rational.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Rational.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -1002,8 +1002,8 @@
   by simp
 
 definition (in term_syntax)
-  valterm_fract :: "int \<times> (unit \<Rightarrow> Code_Eval.term) \<Rightarrow> int \<times> (unit \<Rightarrow> Code_Eval.term) \<Rightarrow> rat \<times> (unit \<Rightarrow> Code_Eval.term)" where
-  [code_unfold]: "valterm_fract k l = Code_Eval.valtermify Fract {\<cdot>} k {\<cdot>} l"
+  valterm_fract :: "int \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> int \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> rat \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
+  [code_unfold]: "valterm_fract k l = Code_Evaluation.valtermify Fract {\<cdot>} k {\<cdot>} l"
 
 notation fcomp (infixl "o>" 60)
 notation scomp (infixl "o\<rightarrow>" 60)
@@ -1014,7 +1014,7 @@
 definition
   "Quickcheck.random i = Quickcheck.random i o\<rightarrow> (\<lambda>num. Random.range i o\<rightarrow> (\<lambda>denom. Pair (
      let j = Code_Numeral.int_of (denom + 1)
-     in valterm_fract num (j, \<lambda>u. Code_Eval.term_of j))))"
+     in valterm_fract num (j, \<lambda>u. Code_Evaluation.term_of j))))"
 
 instance ..
 
--- a/src/HOL/RealDef.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/RealDef.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -1128,8 +1128,8 @@
   by (simp add: of_rat_divide)
 
 definition (in term_syntax)
-  valterm_ratreal :: "rat \<times> (unit \<Rightarrow> Code_Eval.term) \<Rightarrow> real \<times> (unit \<Rightarrow> Code_Eval.term)" where
-  [code_unfold]: "valterm_ratreal k = Code_Eval.valtermify Ratreal {\<cdot>} k"
+  valterm_ratreal :: "rat \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> real \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
+  [code_unfold]: "valterm_ratreal k = Code_Evaluation.valtermify Ratreal {\<cdot>} k"
 
 notation fcomp (infixl "o>" 60)
 notation scomp (infixl "o\<rightarrow>" 60)
--- a/src/HOL/Tools/Datatype/datatype.ML	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Tools/Datatype/datatype.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -26,7 +26,7 @@
   val info_of_case : theory -> string -> info option
   val interpretation : (config -> string list -> theory -> theory) -> theory -> theory
   val distinct_simproc : simproc
-  val make_case :  Proof.context -> bool -> string list -> term ->
+  val make_case :  Proof.context -> DatatypeCase.config -> string list -> term ->
     (term * term) list -> term * (term * (int * bool)) list
   val strip_case : Proof.context -> bool -> term -> (term * (term * term) list) option
   val read_typ: theory ->
--- a/src/HOL/Tools/Datatype/datatype_case.ML	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Tools/Datatype/datatype_case.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -7,8 +7,9 @@
 
 signature DATATYPE_CASE =
 sig
+  datatype config = Error | Warning | Quiet;
   val make_case: (string -> DatatypeAux.info option) ->
-    Proof.context -> bool -> string list -> term -> (term * term) list ->
+    Proof.context -> config -> string list -> term -> (term * term) list ->
     term * (term * (int * bool)) list
   val dest_case: (string -> DatatypeAux.info option) -> bool ->
     string list -> term -> (term * (term * term) list) option
@@ -23,6 +24,8 @@
 structure DatatypeCase : DATATYPE_CASE =
 struct
 
+datatype config = Error | Warning | Quiet;
+
 exception CASE_ERROR of string * int;
 
 fun match_type thy pat ob = Sign.typ_match thy (pat, ob) Vartab.empty;
@@ -260,7 +263,7 @@
         else x :: xs)
     | _ => I) pat [];
 
-fun gen_make_case ty_match ty_inst type_of tab ctxt err used x clauses =
+fun gen_make_case ty_match ty_inst type_of tab ctxt config used x clauses =
   let
     fun string_of_clause (pat, rhs) = Syntax.string_of_term ctxt
       (Syntax.const "_case1" $ pat $ rhs);
@@ -285,7 +288,7 @@
     val originals = map (row_of_pat o #2) rows
     val _ = case originals \\ finals of
         [] => ()
-      | is => (if err then case_error else warning)
+        | is => (case config of Error => case_error | Warning => warning | Quiet => fn _ => {})
           ("The following clauses are redundant (covered by preceding clauses):\n" ^
            cat_lines (map (string_of_clause o nth clauses) is));
   in
@@ -338,7 +341,8 @@
       fun dest_case2 (Const ("_case2", _) $ t $ u) = t :: dest_case2 u
         | dest_case2 t = [t];
       val (cases, cnstrts) = split_list (map dest_case1 (dest_case2 u));
-      val (case_tm, _) = make_case_untyped (tab_of thy) ctxt err []
+      val (case_tm, _) = make_case_untyped (tab_of thy) ctxt
+        (if err then Error else Warning) []
         (fold (fn tT => fn t => Syntax.const "_constrain" $ t $ tT)
            (flat cnstrts) t) cases;
     in case_tm end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Predicate_Compile/pred_compile_aux.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -0,0 +1,100 @@
+(* Author: Lukas Bulwahn, TU Muenchen
+
+Auxilary functions for predicate compiler
+*)
+
+structure Predicate_Compile_Aux =
+struct
+
+(* syntactic functions *)
+ 
+fun is_equationlike_term (Const ("==", _) $ _ $ _) = true
+  | is_equationlike_term (Const ("Trueprop", _) $ (Const ("op =", _) $ _ $ _)) = true
+  | is_equationlike_term _ = false
+  
+val is_equationlike = is_equationlike_term o prop_of 
+
+fun is_pred_equation_term (Const ("==", _) $ u $ v) =
+  (fastype_of u = @{typ bool}) andalso (fastype_of v = @{typ bool})
+  | is_pred_equation_term _ = false
+  
+val is_pred_equation = is_pred_equation_term o prop_of 
+
+fun is_intro_term constname t =
+  case fst (strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl t))) of
+    Const (c, _) => c = constname
+  | _ => false
+  
+fun is_intro constname t = is_intro_term constname (prop_of t)
+
+fun is_pred thy constname =
+  let
+    val T = (Sign.the_const_type thy constname)
+  in body_type T = @{typ "bool"} end;
+  
+
+fun is_predT (T as Type("fun", [_, _])) = (snd (strip_type T) = HOLogic.boolT)
+  | is_predT _ = false
+
+  
+(*** check if a term contains only constructor functions ***)
+fun is_constrt thy =
+  let
+    val cnstrs = flat (maps
+      (map (fn (_, (Tname, _, cs)) => map (apsnd (rpair Tname o length)) cs) o #descr o snd)
+      (Symtab.dest (Datatype.get_all thy)));
+    fun check t = (case strip_comb t of
+        (Free _, []) => true
+      | (Const (s, T), ts) => (case (AList.lookup (op =) cnstrs s, body_type T) of
+            (SOME (i, Tname), Type (Tname', _)) => length ts = i andalso Tname = Tname' andalso forall check ts
+          | _ => false)
+      | _ => false)
+  in check end;  
+  
+fun strip_ex (Const ("Ex", _) $ Abs (x, T, t)) =
+  let
+    val (xTs, t') = strip_ex t
+  in
+    ((x, T) :: xTs, t')
+  end
+  | strip_ex t = ([], t)
+
+fun focus_ex t nctxt =
+  let
+    val ((xs, Ts), t') = apfst split_list (strip_ex t) 
+    val (xs', nctxt') = Name.variants xs nctxt;
+    val ps' = xs' ~~ Ts;
+    val vs = map Free ps';
+    val t'' = Term.subst_bounds (rev vs, t');
+  in ((ps', t''), nctxt') end;
+
+
+
+
+(*
+fun map_atoms f intro = 
+fun fold_atoms f intro =
+*)
+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
+      | _ => 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 equals_conv lhs_cv rhs_cv ct =
+  case Thm.term_of ct of
+    Const ("==", _) $ _ $ _ => Conv.arg_conv cv ct  
+  | _ => error "equals_conv"  
+*)
+
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Predicate_Compile/pred_compile_data.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -0,0 +1,223 @@
+(* Author: Lukas Bulwahn, TU Muenchen
+
+Book-keeping datastructure for the predicate compiler
+
+*)
+signature PRED_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 normalize_equation : theory -> thm -> thm
+end;
+
+structure Pred_Compile_Data : PRED_COMPILE_DATA =
+struct
+
+open Predicate_Compile_Aux;
+
+structure Data = TheoryDataFun
+(
+  type T =
+    {const_spec_table : thm list Symtab.table};
+  val empty =
+    {const_spec_table = Symtab.empty};
+  val copy = I;
+  val extend = I;
+  fun merge _
+    ({const_spec_table = const_spec_table1},
+     {const_spec_table = const_spec_table2}) =
+     {const_spec_table = Symtab.merge (K true) (const_spec_table1, const_spec_table2)}
+);
+
+fun mk_data c = {const_spec_table = c}
+fun map_data f {const_spec_table = c} = mk_data (f c)
+
+type specification_table = thm list Symtab.table
+
+fun defining_const_of_introrule_term t =
+  let
+    val _ $ u = Logic.strip_imp_concl t
+    val (pred, all_args) = strip_comb u
+  in case pred of
+    Const (c, T) => c
+    | _ => raise TERM ("defining_const_of_introrule_term failed: Not a constant", [t])
+  end
+
+val defining_const_of_introrule = defining_const_of_introrule_term o prop_of
+
+(*TODO*)
+fun is_introlike_term t = true
+
+val is_introlike = is_introlike_term o prop_of
+
+fun check_equation_format_term (t as (Const ("==", _) $ u $ v)) =
+  (case strip_comb u of
+    (Const (c, T), args) =>
+      if (length (binder_types T) = length args) then
+        true
+      else
+        raise TERM ("check_equation_format_term failed: Number of arguments mismatch", [t])
+  | _ => raise TERM ("check_equation_format_term failed: Not a constant", [t]))
+  | check_equation_format_term t =
+    raise TERM ("check_equation_format_term failed: Not an equation", [t])
+
+val check_equation_format = check_equation_format_term o prop_of
+
+fun defining_const_of_equation_term (t as (Const ("==", _) $ u $ v)) =
+  (case fst (strip_comb u) of
+    Const (c, _) => c
+  | _ => raise TERM ("defining_const_of_equation_term failed: Not a constant", [t]))
+  | defining_const_of_equation_term t =
+    raise TERM ("defining_const_of_equation_term failed: Not an equation", [t])
+
+val defining_const_of_equation = defining_const_of_equation_term o prop_of
+
+(* Normalizing equations *)
+
+fun mk_meta_equation th =
+  case prop_of th of
+    Const ("Trueprop", _) $ (Const ("op =", _) $ _ $ _) => th RS @{thm eq_reflection}
+  | _ => th
+
+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;
+
+fun declare_names s xs ctxt =
+  let
+    val res = Name.names ctxt s xs
+  in (res, fold Name.declare (map fst res) ctxt) end
+  
+fun split_all_pairs thy th =
+  let
+    val ctxt = ProofContext.init thy
+    val ((_, [th']), ctxt') = Variable.import true [th] ctxt
+    val t = prop_of th'
+    val frees = Term.add_frees t [] 
+    val freenames = Term.add_free_names t []
+    val nctxt = Name.make_context freenames
+    fun mk_tuple_rewrites (x, T) nctxt =
+      let
+        val Ts = HOLogic.flatten_tupleT T
+        val (xTs, nctxt') = declare_names x Ts nctxt
+        val paths = HOLogic.flat_tupleT_paths T
+      in ((Free (x, T), HOLogic.mk_ptuple paths T (map Free xTs)), nctxt') end
+    val (rewr, _) = fold_map mk_tuple_rewrites frees nctxt 
+    val t' = Pattern.rewrite_term thy rewr [] t
+    val tac = setmp quick_and_dirty true (SkipProof.cheat_tac thy)
+    val th'' = Goal.prove ctxt (Term.add_free_names t' []) [] t' (fn {...} => tac)
+    val th''' = LocalDefs.unfold ctxt [@{thm split_conv}] th''
+  in
+    th'''
+  end;
+
+fun normalize_equation thy th =
+  mk_meta_equation th
+	|> Pred_Compile_Set.unfold_set_notation
+  |> full_fun_cong_expand
+  |> split_all_pairs thy
+  |> tap check_equation_format
+
+fun inline_equations thy th = Conv.fconv_rule (Simplifier.rewrite
+((Simplifier.theory_context thy Simplifier.empty_ss) addsimps (Predicate_Compile_Inline_Defs.get (ProofContext.init thy)))) th
+
+fun store_thm_in_table ignore_consts thy th=
+  let
+    val th = AxClass.unoverload thy th
+      |> inline_equations 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 error "store_thm: unexpected definition format"
+  in
+    if not (member (op =) ignore_consts const) then
+      Symtab.cons_list (const, th)
+    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_Const_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)))
+    val table = Symtab.empty
+      |> store [] Predicate_Compile_Alternative_Defs.get
+    val ignore_consts = Symtab.keys table
+  in
+    table   
+    |> store ignore_consts Predicate_Compile_Preproc_Const_Defs.get
+    |> store ignore_consts Nitpick_Const_Simps.get
+    |> store ignore_consts Nitpick_Ind_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
+  | 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 &"}]
+
+val special_cases = member (op =) [@{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},
+  @{const_name Int.Bit1},
+  @{const_name Int.Pls},
+@{const_name "Int.zero_int_inst.zero_int"},
+@{const_name "List.filter"}]
+
+fun obtain_specification_graph table constname =
+  let
+    fun is_nondefining_constname c = member (op =) logic_operator_names c
+    val is_defining_constname = member (op =) (Symtab.keys table)
+    fun defiants_of specs =
+      fold (Term.add_const_names o prop_of) specs []
+      |> filter is_defining_constname
+      |> filter_out special_cases
+    fun extend constname =
+      let
+        val specs = get_specification table constname
+      in (specs, defiants_of specs) end;
+  in
+    Graph.extend extend constname Graph.empty
+  end;
+  
+  
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Predicate_Compile/pred_compile_fun.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -0,0 +1,424 @@
+(* Author: Lukas Bulwahn, TU Muenchen
+
+Preprocessing functions to predicates
+*)
+
+signature PREDICATE_COMPILE_FUN =
+sig
+  val define_predicates : (string * thm list) list -> theory -> theory
+  val rewrite_intro : theory -> thm -> thm list
+  val setup_oracle : theory -> theory
+end;
+
+structure Predicate_Compile_Fun : PREDICATE_COMPILE_FUN =
+struct
+
+
+(* Oracle for preprocessing  *)
+
+val (oracle : (string * (cterm -> thm)) option ref) = ref NONE;
+
+fun the_oracle () =
+  case !oracle of
+    NONE => error "Oracle is not setup"
+  | SOME (_, oracle) => oracle
+             
+val setup_oracle = Thm.add_oracle (Binding.name "pred_compile_preprocessing", I) #->
+  (fn ora => fn thy => let val _ = (oracle := SOME ora) in thy end)
+  
+  
+fun is_funtype (Type ("fun", [_, _])) = true
+  | is_funtype _ = false;
+
+fun is_Type (Type _) = true
+  | is_Type _ = false
+
+(* returns true if t is an application of an datatype constructor *)
+(* which then consequently would be splitted *)
+(* else false *)
+(*
+fun is_constructor thy t =
+  if (is_Type (fastype_of t)) then
+    (case DatatypePackage.get_datatype thy ((fst o dest_Type o fastype_of) t) of
+      NONE => false
+    | SOME info => (let
+      val constr_consts = flat (map (fn (_, (_, _, constrs)) => map fst constrs) (#descr info))
+      val (c, _) = strip_comb t
+      in (case c of
+        Const (name, _) => name mem_string constr_consts
+        | _ => false) end))
+  else false
+*)
+
+(* must be exported in code.ML *)
+fun is_constr thy = is_some o Code.get_datatype_of_constr thy;
+
+(* Table from constant name (string) to term of inductive predicate *)
+structure Pred_Compile_Preproc = TheoryDataFun
+(
+  type T = string Symtab.table;
+  val empty = Symtab.empty;
+  val copy = I;
+  val extend = I;
+  fun merge _ = Symtab.merge (op =);
+)
+
+fun defined thy = Symtab.defined (Pred_Compile_Preproc.get thy) 
+
+
+fun transform_ho_typ (T as Type ("fun", _)) =
+  let
+    val (Ts, T') = strip_type T
+  in if T' = @{typ "bool"} then T else (Ts @ [T']) ---> HOLogic.boolT end
+| transform_ho_typ t = t
+
+fun transform_ho_arg arg = 
+  case (fastype_of arg) of
+    (T as Type ("fun", _)) =>
+      (case arg of
+        Free (name, _) => Free (name, transform_ho_typ T)
+      | _ => error "I am surprised")
+| _ => arg
+
+fun pred_type T =
+  let
+    val (Ts, T') = strip_type T
+    val Ts' = map transform_ho_typ Ts
+  in
+    (Ts' @ [T']) ---> HOLogic.boolT
+  end;
+
+(* FIXME: create new predicate name -- does not avoid nameclashing *)
+fun pred_of f =
+  let
+    val (name, T) = dest_Const f
+  in
+    if (body_type T = @{typ bool}) then
+      (Free (Long_Name.base_name name ^ "P", T))
+    else
+      (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 =
+  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;
+
+
+(* creates the list of premises for every intro rule *)
+(* theory -> term -> (string list, term list list) *)
+
+fun dest_code_eqn eqn = let
+  val (lhs, rhs) = Logic.dest_equals (Logic.unvarify (Thm.prop_of eqn))
+  val (func, args) = strip_comb lhs
+in ((func, args), rhs) end;
+
+fun string_of_typ T = Syntax.string_of_typ_global @{theory} T
+
+fun string_of_term t =
+  case t of
+    Const (c, T) => "Const (" ^ c ^ ", " ^ string_of_typ T ^ ")"
+  | Free (c, T) => "Free (" ^ c ^ ", " ^ string_of_typ T ^ ")"
+  | Var ((c, i), T) => "Var ((" ^ c ^ ", " ^ string_of_int i ^ "), " ^ string_of_typ T ^ ")"
+  | Bound i => "Bound " ^ string_of_int i
+  | Abs (x, T, t) => "Abs (" ^ x ^ ", " ^ string_of_typ T ^ ", " ^ string_of_term t ^ ")"
+  | t1 $ t2 => "(" ^ string_of_term t1 ^ ") $ (" ^ string_of_term t2 ^ ")"
+  
+fun ind_package_get_nparams thy name =
+  case try (Inductive.the_inductive (ProofContext.init thy)) name of
+    SOME (_, result) => length (Inductive.params_of (#raw_induct result))
+  | NONE => error ("No such predicate: " ^ quote name) 
+
+(* TODO: does not work with higher order functions yet *)
+fun mk_rewr_eq (func, pred) =
+  let
+    val (argTs, resT) = (strip_type (fastype_of func))
+    val nctxt =
+      Name.make_context (Term.fold_aterms (fn Free (x, _) => insert (op =) x | _ => I) (func $ pred) [])
+    val (argnames, nctxt') = Name.variants (replicate (length argTs) "a") nctxt
+    val ([resname], nctxt'') = Name.variants ["r"] nctxt'
+    val args = map Free (argnames ~~ argTs)
+    val res = Free (resname, resT)
+  in Logic.mk_equals
+      (HOLogic.mk_eq (res, list_comb (func, args)), list_comb (pred, args @ [res]))
+  end;
+
+fun has_split_rule_cname @{const_name "nat_case"} = true
+  | has_split_rule_cname @{const_name "list_case"} = true
+  | has_split_rule_cname _ = false
+  
+fun has_split_rule_term thy (Const (@{const_name "nat_case"}, _)) = true 
+  | has_split_rule_term thy (Const (@{const_name "list_case"}, _)) = true 
+  | has_split_rule_term thy _ = false
+
+fun has_split_rule_term' thy (Const (@{const_name "If"}, _)) = true
+  | has_split_rule_term' thy (Const (@{const_name "Let"}, _)) = true
+  | has_split_rule_term' thy c = has_split_rule_term thy c
+  
+fun prepare_split_thm ctxt split_thm =
+    (split_thm RS @{thm iffD2})
+    |> LocalDefs.unfold ctxt [@{thm atomize_conjL[symmetric]},
+      @{thm atomize_all[symmetric]}, @{thm atomize_imp[symmetric]}]
+
+fun find_split_thm thy (Const (name, typ)) =
+  let
+    fun split_name str =
+      case first_field "." str
+        of (SOME (field, rest)) => field :: split_name rest
+         | NONE => [str]
+    val splitted_name = split_name name
+  in
+    if length splitted_name > 0 andalso
+       String.isSuffix "_case" (List.last splitted_name)
+    then
+      (List.take (splitted_name, length splitted_name - 1)) @ ["split"]
+      |> String.concatWith "."
+      |> PureThy.get_thm thy
+      |> SOME
+      handle ERROR msg => NONE
+    else NONE
+  end
+  | find_split_thm _ _ = NONE
+
+fun find_split_thm' thy (Const (@{const_name "If"}, _)) = SOME @{thm split_if}
+  | find_split_thm' thy (Const (@{const_name "Let"}, _)) = SOME @{thm refl} (* TODO *)
+  | find_split_thm' thy c = find_split_thm thy c
+
+fun strip_all t = (Term.strip_all_vars t, Term.strip_all_body t)
+
+fun folds_map f xs y =
+  let
+    fun folds_map' acc [] y = [(rev acc, y)]
+      | folds_map' acc (x :: xs) y =
+        maps (fn (x, y) => folds_map' (x :: acc) xs y) (f x y)
+    in
+      folds_map' [] xs y
+    end;
+
+fun mk_prems thy (lookup_pred, get_nparams) t (names, prems) =
+  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))]
+      else [(lookup_pred t, (names, prems))]
+    | mk_prems' (t as Free (f, T)) (names, prems) = 
+      [(lookup_pred t, (names, prems))]
+    | mk_prems' t (names, prems) =
+      if Predicate_Compile_Aux.is_constrt thy t then
+        [(t, (names, prems))]
+      else
+        if has_split_rule_term' thy (fst (strip_comb t)) then
+          let
+            val (f, args) = strip_comb t
+            val split_thm = prepare_split_thm (ProofContext.init thy) (the (find_split_thm' thy f))
+            (* TODO: contextify things - this line is to unvarify the split_thm *)
+            (*val ((_, [isplit_thm]), _) = Variable.import true [split_thm] (ProofContext.init thy)*)
+            val (assms, concl) = Logic.strip_horn (Thm.prop_of split_thm)
+            val (P, [split_t]) = strip_comb (HOLogic.dest_Trueprop concl) 
+            val subst = Pattern.match thy (split_t, t) (Vartab.empty, Vartab.empty)
+            val (_, split_args) = strip_comb split_t
+            val match = split_args ~~ args
+            fun mk_prems_of_assm assm =
+              let
+                val (vTs, assm') = strip_all (Envir.beta_norm (Envir.subst_term subst assm))
+                val var_names = Name.variant_list names (map fst vTs)
+                val vars = map Free (var_names ~~ (map snd vTs))
+                val (prems', pre_res) = Logic.strip_horn (subst_bounds (rev vars, assm'))
+                val (_, [inner_t]) = strip_comb (HOLogic.dest_Trueprop pre_res)
+              in
+                mk_prems' inner_t (var_names @ names, prems' @ prems)
+              end
+          in
+            maps mk_prems_of_assm assms
+          end
+        else
+          let 
+            val (f, args) = strip_comb t
+            val resname = Name.variant names "res"
+            val resvar = Free (resname, body_type (fastype_of t))
+            val names' = resname :: names
+            fun mk_prems'' (t as Const (c, _)) =
+              if is_constr thy c orelse (is_none (try lookup_pred t)) then
+                folds_map mk_prems' args (names', prems) |>
+                map
+                  (fn (argvs, (names'', prems')) =>
+                  let
+                    val prem = HOLogic.mk_Trueprop (HOLogic.mk_eq (resvar, list_comb (f, argvs)))
+                  in (names'', prem :: prems') end)
+              else
+                let
+                  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
+                in
+                  folds_map mk_prems' args (names', prems)
+                  |> map (fn (argvs, (names'', prems')) =>
+                    let
+                      val prem = HOLogic.mk_Trueprop (list_comb (pred, params' @ argvs @ [resvar]))
+                    in (names'', prem :: prems') end)
+                end
+            | mk_prems'' (t as Free (_, _)) =
+                let
+                  (* higher order argument call *)
+                  val pred = lookup_pred t
+                in
+                  folds_map mk_prems' args (resname :: names, prems)
+                  |> map (fn (argvs, (names', prems')) =>
+                     let
+                       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)
+          in
+            map (pair resvar) (mk_prems'' f)
+          end
+  in
+    mk_prems' t (names, prems)
+  end;
+
+(* 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
+  else
+  let
+    val consts = map fst specs
+    val eqns = maps snd specs
+    (*val eqns = maps (Predicate_Compile_Preproc_Data.get_specification thy) consts*)
+      (* create prednames *)
+    val ((funs, argss), rhss) = map_split dest_code_eqn eqns |>> split_list
+    val argss' = map (map transform_ho_arg) argss
+    val pnames = map dest_Free (distinct (op =) (maps (filter (is_funtype o fastype_of)) argss'))
+    val preds = map pred_of funs
+    val prednames = map (fst o dest_Free) preds
+    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)) =
+      (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)) =
+        if member op = (map fst pnames) name then
+          Free (name, transform_ho_typ T)
+        else
+          Free (name, T)
+      | lookup_pred t =
+         error ("lookup function is not defined for " ^ Syntax.string_of_term_global thy t)
+     
+        (* mapping from term (predicate term, not function term!) to int *)
+    fun get_nparams (Const (name, _)) =
+      the_default 0 (try (ind_package_get_nparams thy) name)
+    | get_nparams (Free (name, _)) =
+        (if member op = prednames name then
+          length pnames
+        else 0)
+    | get_nparams t = error ("No parameters for " ^ (Syntax.string_of_term_global thy t))
+  
+    (* create intro rules *)
+  
+    fun mk_intros ((func, pred), (args, rhs)) =
+      if (body_type (fastype_of func) = @{typ bool}) then
+       (*TODO: preprocess predicate definition of rhs *)
+        [Logic.list_implies ([HOLogic.mk_Trueprop rhs], HOLogic.mk_Trueprop (list_comb (pred, args)))]
+      else
+        let
+          val names = Term.add_free_names rhs []
+        in mk_prems thy (lookup_pred, get_nparams) rhs (names, [])
+          |> map (fn (resultt, (names', prems)) =>
+            Logic.list_implies (prems, HOLogic.mk_Trueprop (list_comb (pred, args @ [resultt]))))
+        end
+    fun mk_rewr_thm (func, pred) = @{thm refl}
+  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
+        if is_some (try (map (cterm_of thy)) intr_ts) then
+          let
+            val (ind_result, thy') =
+              Inductive.add_inductive_global (serial_string ())
+                {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}
+                (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (distinct (op =) (map dest_Free preds)))
+                pnames
+                (map (fn x => (Attrib.empty_binding, x)) intr_ts)
+                [] thy
+            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
+        else
+          thy
+      end
+  end
+
+(* preprocessing intro rules - uses oracle *)
+
+(* theory -> thm -> thm *)
+fun rewrite_intro thy intro =
+  let
+    fun lookup_pred (Const (name, T)) =
+      (case (Symtab.lookup (Pred_Compile_Preproc.get thy) name) of
+        SOME c => Const (c, pred_type T)
+      | NONE => error ("Function " ^ name ^ " is not inductified"))
+    | lookup_pred (Free (name, T)) = Free (name, T)
+    | lookup_pred _ = error "lookup function is not defined!"
+
+    fun get_nparams (Const (name, _)) =
+      the_default 0 (try (ind_package_get_nparams thy) name)
+    | get_nparams (Free _) = 0
+    | 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 =
+      let
+        val t = (HOLogic.dest_Trueprop prem)
+        val (lit, mk_lit) = case try HOLogic.dest_not t of
+            SOME t => (t, HOLogic.mk_not)
+          | NONE => (t, I)
+        val (P, args) = (strip_comb lit) 
+      in
+        folds_map (
+          fn t => if (is_funtype (fastype_of t)) then (fn x => [(t, x)])
+            else mk_prems thy (lookup_pred, get_nparams) t) args (names, [])
+        |> map (fn (resargs, (names', prems')) =>
+          let
+            val prem' = HOLogic.mk_Trueprop (mk_lit (list_comb (P, resargs)))
+          in (prem'::prems', names') end)
+      end
+    val intro_ts' = folds_map rewrite prems frees
+      |> maps (fn (prems', frees') =>
+        rewrite concl frees'
+        |> map (fn (concl'::conclprems, _) =>
+          Logic.list_implies ((flat prems') @ conclprems, concl')))
+    val _ = Output.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; 
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Predicate_Compile/pred_compile_pred.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -0,0 +1,138 @@
+(* Author: Lukas Bulwahn, TU Muenchen
+
+Preprocessing definitions of predicates to introduction rules
+*)
+
+signature PREDICATE_COMPILE_PRED =
+sig
+  (* preprocesses an equation to a set of intro rules; defines new constants *)
+  (*
+  val preprocess_pred_equation : thm -> theory -> thm list * theory
+  *)
+  val preprocess : string -> theory -> (thm list list * theory) 
+  (* output is the term list of clauses of an unknown predicate *)
+  val preprocess_term : term -> theory -> (term list * theory)
+  
+  (*val rewrite : thm -> thm*)
+  
+end;
+(* : PREDICATE_COMPILE_PREPROC_PRED *)
+structure Predicate_Compile_Pred =
+struct
+
+open Predicate_Compile_Aux
+
+fun is_compound ((Const ("Not", _)) $ t) =
+    error "is_compound: Negation should not occur; preprocessing is defect"
+  | is_compound ((Const ("Ex", _)) $ _) = true
+  | is_compound ((Const (@{const_name "op |"}, _)) $ _ $ _) = true
+  | is_compound ((Const ("op &", _)) $ _ $ _) =
+    error "is_compound: Conjunction should not occur; preprocessing is defect"
+  | is_compound _ = false
+
+fun flatten constname atom (defs, thy) =
+  if is_compound atom then
+    let
+      val constname = Name.variant (map (Long_Name.base_name o fst) defs)
+        ((Long_Name.base_name constname) ^ "_aux")
+      val full_constname = Sign.full_bname thy constname
+      val (params, args) = List.partition (is_predT o fastype_of)
+        (map Free (Term.add_frees atom []))
+      val constT = map fastype_of (params @ args) ---> HOLogic.boolT
+      val lhs = list_comb (Const (full_constname, constT), params @ args)
+      val def = Logic.mk_equals (lhs, atom)
+      val ([definition], thy') = thy
+        |> Sign.add_consts_i [(Binding.name constname, constT, NoSyn)]
+        |> PureThy.add_defs false [((Binding.name (constname ^ "_def"), def), [])]
+    in
+      (lhs, ((full_constname, [definition]) :: defs, thy'))
+    end
+  else
+    (atom, (defs, thy))
+
+fun flatten_intros constname intros thy =
+  let
+    val ctxt = ProofContext.init thy
+    val ((_, intros), ctxt') = Variable.import true intros ctxt
+    val (intros', (local_defs, thy')) = (fold_map o fold_map_atoms)
+      (flatten constname) (map prop_of intros) ([], thy)
+    val tac = fn {...} => setmp quick_and_dirty true (SkipProof.cheat_tac thy')
+    val intros'' = map (fn t => Goal.prove ctxt' [] [] t tac) intros'
+      |> Variable.export ctxt' ctxt
+  in
+    (intros'', (local_defs, thy'))
+  end
+
+(* TODO: same function occurs in inductive package *)
+fun select_disj 1 1 = []
+  | select_disj _ 1 = [rtac @{thm disjI1}]
+  | select_disj n i = (rtac @{thm disjI2})::(select_disj (n - 1) (i - 1));
+
+fun introrulify thy ths = 
+  let
+    val ctxt = ProofContext.init thy
+    val ((_, ths'), ctxt') = Variable.import true ths ctxt
+    fun introrulify' th =
+      let
+        val (lhs, rhs) = Logic.dest_equals (prop_of th)
+        val frees = Term.add_free_names rhs []
+        val disjuncts = HOLogic.dest_disj rhs
+        val nctxt = Name.make_context frees
+        fun mk_introrule t =
+          let
+            val ((ps, t'), nctxt') = focus_ex t nctxt
+            val prems = map HOLogic.mk_Trueprop (HOLogic.dest_conj t')
+          in
+            (ps, Logic.list_implies (prems, HOLogic.mk_Trueprop lhs))
+          end
+        val x = ((cterm_of thy) o the_single o snd o strip_comb o HOLogic.dest_Trueprop o fst o
+          Logic.dest_implies o prop_of) @{thm exI}
+        fun prove_introrule (index, (ps, introrule)) =
+          let
+            val tac = Simplifier.simp_tac (HOL_basic_ss addsimps [th]) 1
+              THEN EVERY1 (select_disj (length disjuncts) (index + 1)) 
+              THEN (EVERY (map (fn y =>
+                rtac (Drule.cterm_instantiate [(x, cterm_of thy (Free y))] @{thm exI}) 1) ps))
+              THEN REPEAT_DETERM (rtac @{thm conjI} 1 THEN atac 1)
+              THEN TRY (atac 1)
+          in
+            Goal.prove ctxt' (map fst ps) [] introrule (fn {...} => tac)
+          end
+      in
+        map_index prove_introrule (map mk_introrule disjuncts)
+      end
+  in maps introrulify' ths' |> Variable.export ctxt' ctxt end
+
+val rewrite =
+  Simplifier.simplify (HOL_basic_ss addsimps [@{thm Ball_def}, @{thm Bex_def}])
+  #> Simplifier.simplify (HOL_basic_ss addsimps [@{thm all_not_ex}]) 
+  #> Conv.fconv_rule nnf_conv 
+  #> Simplifier.simplify (HOL_basic_ss addsimps [@{thm ex_disj_distrib}])
+
+val rewrite_intros =
+  Simplifier.simplify (HOL_basic_ss addsimps @{thms HOL.simp_thms(9)})
+
+fun preprocess (constname, specs) thy =
+  let
+    val ctxt = ProofContext.init thy
+      val intros =
+      if forall is_pred_equation specs then 
+        introrulify thy (map rewrite specs)
+      else if forall (is_intro constname) specs then
+        map rewrite_intros specs
+      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'')
+  end;
+
+fun preprocess_term t thy = error "preprocess_pred_term: to implement" 
+  
+  
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Predicate_Compile/pred_compile_quickcheck.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -0,0 +1,93 @@
+(* Author: Lukas Bulwahn, TU Muenchen
+
+A quickcheck generator based on the predicate compiler
+*)
+
+signature PRED_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) ref
+end;
+
+structure Pred_Compile_Quickcheck : PRED_COMPILE_QUICKCHECK =
+struct
+
+val test_ref = 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
+val mk_predT = #mk_predT (dest_compfuns Predicate_Compile_Core.pred_compfuns)
+val mk_rpredT = #mk_predT (dest_compfuns Predicate_Compile_Core.rpred_compfuns)
+val mk_return = #mk_single (dest_compfuns Predicate_Compile_Core.rpred_compfuns)
+val mk_bind = #mk_bind (dest_compfuns Predicate_Compile_Core.rpred_compfuns)
+val lift_pred = #lift_pred (dest_compfuns Predicate_Compile_Core.rpred_compfuns)
+
+fun mk_split_lambda [] t = Abs ("u", HOLogic.unitT, t)
+  | mk_split_lambda [x] t = lambda x t
+  | mk_split_lambda xs t =
+  let
+    fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t))
+      | mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t))
+  in
+    mk_split_lambda' xs t
+  end;
+
+fun strip_imp_prems (Const("op -->", _) $ A $ B) = A :: strip_imp_prems B
+  | strip_imp_prems _ = [];
+
+fun strip_imp_concl (Const("op -->", _) $ A $ B) = strip_imp_concl B
+  | strip_imp_concl A = A : term;
+
+fun strip_horn A = (strip_imp_prems A, strip_imp_concl A);
+
+fun quickcheck ctxt t =
+  let
+    val _ = tracing ("Starting quickcheck with " ^ (Syntax.string_of_term ctxt t))
+    val ctxt' = ProofContext.theory (Context.copy_thy) ctxt
+    val thy = (ProofContext.theory_of ctxt') 
+    val (vs, t') = strip_abs t
+    val vs' = Variable.variant_frees ctxt' [] vs
+    val t'' = subst_bounds (map Free (rev vs'), t')
+    val (prems, concl) = strip_horn t''
+    val constname = "pred_compile_quickcheck"
+    val full_constname = Sign.full_bname thy constname
+    val constT = map snd vs' ---> @{typ bool}
+    val thy' = Sign.add_consts_i [(Binding.name constname, constT, NoSyn)] thy
+    val t = Logic.list_implies
+      (map HOLogic.mk_Trueprop (prems @ [HOLogic.mk_not concl]),
+       HOLogic.mk_Trueprop (list_comb (Const (full_constname, constT), map Free vs')))
+    val tac = fn {...} => setmp quick_and_dirty true (SkipProof.cheat_tac thy')
+    val intro = Goal.prove (ProofContext.init thy') (map fst vs') [] t tac
+    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
+    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
+        let
+          val name = Predicate_Compile_Core.sizelim_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"
+    val qc_term = Abs ("size", @{typ code_numeral}, mk_bind (prog,
+      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)
+      (fn proc => fn g => fn s => g s #>> (Predicate.map o map) proc)
+      thy'' qc_term []
+  in
+    ((compile #> Random_Engine.run) #> (Option.map fst o Predicate.yield))
+  end
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Predicate_Compile/pred_compile_set.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -0,0 +1,51 @@
+(* Author: Lukas Bulwahn, TU Muenchen
+
+Preprocessing sets to predicates
+*)
+
+signature PRED_COMPILE_SET =
+sig
+(*
+  val preprocess_intro : thm -> theory -> thm * theory
+  val preprocess_clause : term -> theory -> term * theory
+*)
+  val unfold_set_notation : thm -> thm;
+end;
+
+structure Pred_Compile_Set : PRED_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_notation = Simplifier.rewrite_rule unfold_set_lemmas
+
+(*
+fun find_set_theorems ctxt cname =
+  let
+    val _ = cname 
+*)
+
+(*
+fun preprocess_term t ctxt =
+  case t of
+    Const ("op :", _) $ x $ A => 
+      case strip_comb A of
+        (Const (cname, T), params) =>
+          let 
+            val _ = find_set_theorems
+          in
+            (t, ctxt)
+          end
+      | _ => (t, ctxt)  
+  | _ => (t, ctxt)
+*) 
+(*
+fun preprocess_intro th thy =
+  let
+    val cnames = find_heads_of_prems
+    val set_cname = filter (has_set_definition
+    val _ = define_preds
+    val _ = prep_pred_def
+  in
+*)
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -0,0 +1,91 @@
+(* Author: Lukas Bulwahn, TU Muenchen
+
+*)
+signature PREDICATE_COMPILE =
+sig
+  val setup : theory -> theory
+  val preprocess : string -> theory -> theory
+end;
+
+structure Predicate_Compile : PREDICATE_COMPILE =
+struct
+
+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 =
+  let
+    val get_specs = map (fn k => (k, Graph.get_node gr k))
+    val _ = priority ("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 _ = Output.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''
+  in
+    thy'''
+  end;
+
+fun preprocess const thy =
+  let
+    val _ = Output.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 _ = Output.tracing (commas (Graph.all_succs gr [const]))
+    val gr = Graph.subgraph (member (op =) (Graph.all_succs gr [const])) gr
+  in fold_rev (preprocess_strong_conn_constnames 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
+
+val setup = Predicate_Compile_Fun.setup_oracle #> Predicate_Compile_Core.setup
+
+val _ = List.app OuterKeyword.keyword ["inductify_all", "rpred"]
+
+local structure P = OuterParse
+in
+
+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)
+
+end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -0,0 +1,2425 @@
+(* Author: Lukas Bulwahn, TU Muenchen
+
+(Prototype of) A compiler from predicates specified by intro/elim rules
+to equations.
+*)
+
+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
+  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 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 generator_modes_of: theory -> string -> mode list
+  val generator_name_of : theory -> string -> mode -> string
+  val string_of_mode : mode -> string
+  val intros_of: theory -> string -> thm list
+  val nparams_of: theory -> string -> int
+  val add_intro: thm -> theory -> theory
+  val set_elim: thm -> theory -> theory
+  val set_nparams : string -> int -> theory -> theory
+  val print_stored_rules: theory -> unit
+  val print_all_modes: theory -> unit
+  val do_proofs: bool ref
+  val mk_casesrule : Proof.context -> int -> thm list -> term
+  val analyze_compr: theory -> term -> term
+  val eval_ref: (unit -> term Predicate.pred) option ref
+  val add_equations : string list -> theory -> theory
+  val code_pred_intros_attrib : attribute
+  (* used by Quickcheck_Generator *) 
+  (*val funT_of : mode -> typ -> typ
+  val mk_if_pred : term -> term
+  val mk_Eval : term * term -> term*)
+  val mk_tupleT : typ list -> typ
+(*  val mk_predT :  typ -> typ *)
+  (* temporary for testing of the compilation *)
+  datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term |
+    GeneratorPrem of term list * term | Generator of (string * typ);
+ (* val prepare_intrs: theory -> string list ->
+    (string * typ) list * int * string list * string list * (string * mode list) list *
+    (string * (term list * indprem list) list) list * (string * (int option list * int)) list*)
+  datatype compilation_funs = CompilationFuns of {
+    mk_predT : typ -> typ,
+    dest_predT : typ -> typ,
+    mk_bot : typ -> term,
+    mk_single : term -> term,
+    mk_bind : term * term -> term,
+    mk_sup : term * term -> term,
+    mk_if : term -> term,
+    mk_not : term -> term,
+    mk_map : typ -> typ -> term -> term -> term,
+    lift_pred : term -> term
+  };  
+  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
+end;
+
+structure Predicate_Compile_Core : PREDICATE_COMPILE_CORE =
+struct
+
+open Predicate_Compile_Aux;
+(** auxiliary **)
+
+(* debug stuff *)
+
+fun tracing s = (if ! Toplevel.debug then Output.tracing s else ());
+
+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 => (Output.tracing msg; Seq.single st)); *)
+
+val do_proofs = ref true;
+
+(* reference to preprocessing of InductiveSet package *)
+
+val ind_set_codegen_preproc = (fn thy => I) (*Inductive_Set.codegen_preproc;*)
+
+(** fundamentals **)
+
+(* syntactic operations *)
+
+fun mk_eq (x, xs) =
+  let fun mk_eqs _ [] = []
+        | mk_eqs a (b::cs) =
+            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
+    val U = fastype_of u
+    val [A] = binder_types T
+    val D = body_type U 
+  in 
+    Const (@{const_name "scomp"}, T --> U --> A --> D) $ t $ u
+  end;
+
+fun dest_funT (Type ("fun",[S, T])) = (S, T)
+  | dest_funT T = raise TYPE ("dest_funT", [T], [])
+ 
+fun mk_fun_comp (t, u) =
+  let
+    val (_, B) = dest_funT (fastype_of t)
+    val (C, A) = dest_funT (fastype_of u)
+  in
+    Const(@{const_name "Fun.comp"}, (A --> B) --> (C --> A) --> C --> B) $ t $ u
+  end;
+
+fun dest_randomT (Type ("fun", [@{typ Random.seed},
+  Type ("*", [Type ("*", [T, @{typ "unit => Code_Evaluation.term"}]) ,@{typ Random.seed}])])) = T
+  | dest_randomT T = raise TYPE ("dest_randomT", [T], [])
+
+(* destruction of intro rules *)
+
+(* FIXME: look for other place where this functionality was used before *)
+fun strip_intro_concl nparams intro = let
+  val _ $ u = Logic.strip_imp_concl intro
+  val (pred, all_args) = strip_comb u
+  val (params, args) = chop nparams all_args
+in (pred, (params, args)) end
+
+(** data structures **)
+
+type smode = (int * int list option) list
+type mode = smode option list * smode;
+datatype tmode = Mode of mode * smode * tmode option list;
+
+fun gen_split_smode (mk_tuple, strip_tuple) smode ts =
+  let
+    fun split_tuple' _ _ [] = ([], [])
+    | split_tuple' is i (t::ts) =
+      (if i mem is then apfst else apsnd) (cons t)
+        (split_tuple' is (i+1) ts)
+    fun split_tuple is t = split_tuple' is 1 (strip_tuple t)
+    fun split_smode' _ _ [] = ([], [])
+      | split_smode' smode i (t::ts) =
+        (if i mem (map fst smode) then
+          case (the (AList.lookup (op =) smode i)) of
+            NONE => apfst (cons t)
+            | SOME is =>
+              let
+                val (ts1, ts2) = split_tuple is t
+                fun cons_tuple ts = if null ts then I else cons (mk_tuple ts)
+                in (apfst (cons_tuple ts1)) o (apsnd (cons_tuple ts2)) end
+          else apsnd (cons t))
+        (split_smode' smode (i+1) ts)
+  in split_smode' smode 1 ts end
+
+val split_smode = gen_split_smode (HOLogic.mk_tuple, HOLogic.strip_tuple)   
+val split_smodeT = gen_split_smode (HOLogic.mk_tupleT, HOLogic.strip_tupleT)
+
+fun gen_split_mode split_smode (iss, is) ts =
+  let
+    val (t1, t2) = chop (length iss) ts 
+  in (t1, split_smode is t2) end
+
+val split_mode = gen_split_mode split_smode
+val split_modeT = gen_split_mode split_smodeT
+
+fun string_of_smode js =
+    commas (map
+      (fn (i, is) =>
+        string_of_int i ^ (case is of NONE => ""
+    | SOME is => "p" ^ enclose "[" "]" (commas (map string_of_int is)))) js)
+
+fun string_of_mode (iss, is) = space_implode " -> " (map
+  (fn NONE => "X"
+    | SOME js => enclose "[" "]" (string_of_smode js))
+       (iss @ [SOME is]));
+
+fun string_of_tmode (Mode (predmode, termmode, param_modes)) =
+  "predmode: " ^ (string_of_mode predmode) ^ 
+  (if null param_modes then "" else
+    "; " ^ "params: " ^ commas (map (the_default "NONE" o Option.map string_of_tmode) param_modes))
+
+(* 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);
+
+type moded_clause = term list * (indprem * tmode) list
+type 'a pred_mode_table = (string * (mode * 'a) list) list
+
+datatype predfun_data = PredfunData of {
+  name : string,
+  definition : thm,
+  intro : thm,
+  elim : thm
+};
+
+fun rep_predfun_data (PredfunData data) = data;
+fun mk_predfun_data (name, definition, intro, elim) =
+  PredfunData {name = name, definition = definition, intro = intro, elim = elim}
+
+datatype function_data = FunctionData of {
+  name : string,
+  equation : thm option (* is not used at all? *)
+};
+
+fun rep_function_data (FunctionData data) = data;
+fun mk_function_data (name, equation) =
+  FunctionData {name = name, equation = equation}
+
+datatype pred_data = PredData of {
+  intros : thm list,
+  elim : thm option,
+  nparams : int,
+  functions : (mode * predfun_data) list,
+  generators : (mode * function_data) list,
+  sizelim_functions : (mode * function_data) list 
+};
+
+fun rep_pred_data (PredData data) = data;
+fun mk_pred_data ((intros, elim, nparams), (functions, generators, sizelim_functions)) =
+  PredData {intros = intros, elim = elim, nparams = nparams,
+    functions = functions, generators = generators, sizelim_functions = sizelim_functions}
+fun map_pred_data f (PredData {intros, elim, nparams, functions, generators, sizelim_functions}) =
+  mk_pred_data (f ((intros, elim, nparams), (functions, generators, sizelim_functions)))
+  
+fun eq_option eq (NONE, NONE) = true
+  | eq_option eq (SOME x, SOME y) = eq (x, y)
+  | 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;
+  val empty = Graph.empty;
+  val copy = I;
+  val extend = I;
+  fun merge _ = Graph.merge eq_pred_data;
+);
+
+(* queries *)
+
+fun lookup_pred_data thy name =
+  Option.map rep_pred_data (try (Graph.get_node (PredData.get thy)) name)
+
+fun the_pred_data thy name = case lookup_pred_data thy name
+ of NONE => error ("No such predicate " ^ quote name)  
+  | SOME data => data;
+
+val is_registered = is_some oo lookup_pred_data 
+
+val all_preds_of = Graph.keys o PredData.get
+
+fun intros_of thy = map (Thm.transfer thy) o #intros o the_pred_data thy
+
+fun the_elim_of thy name = case #elim (the_pred_data thy name)
+ of NONE => error ("No elimination rule for predicate " ^ quote name)
+  | SOME thm => Thm.transfer thy thm 
+  
+val has_elim = is_some o #elim oo the_pred_data;
+
+val nparams_of = #nparams oo the_pred_data
+
+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 rpred_modes_of = (map fst) o #generators oo the_pred_data
+  
+fun all_modes_of thy = map (fn name => (name, modes_of thy name)) (all_preds_of thy) 
+
+val is_compiled = not o null o #functions oo the_pred_data
+
+fun lookup_predfun_data thy name mode =
+  Option.map rep_predfun_data (AList.lookup (op =)
+  (#functions (the_pred_data thy name)) mode)
+
+fun the_predfun_data thy name mode = case lookup_predfun_data thy name mode
+  of NONE => error ("No function defined for mode " ^ string_of_mode mode ^ " of predicate " ^ name)
+   | SOME data => data;
+
+val predfun_name_of = #name ooo the_predfun_data
+
+val predfun_definition_of = #definition ooo the_predfun_data
+
+val predfun_intro_of = #intro ooo the_predfun_data
+
+val predfun_elim_of = #elim ooo the_predfun_data
+
+fun lookup_generator_data thy name mode = 
+  Option.map rep_function_data (AList.lookup (op =)
+  (#generators (the_pred_data thy name)) mode)
+  
+fun the_generator_data thy name mode = case lookup_generator_data thy name mode
+  of NONE => error ("No generator defined for mode " ^ string_of_mode mode ^ " of predicate " ^ name)
+   | SOME data => data
+
+val generator_name_of = #name ooo the_generator_data
+
+val generator_modes_of = (map fst) o #generators oo the_pred_data
+
+fun all_generator_modes_of thy =
+  map (fn name => (name, generator_modes_of thy name)) (all_preds_of thy) 
+
+fun lookup_sizelim_function_data thy name mode =
+  Option.map rep_function_data (AList.lookup (op =)
+  (#sizelim_functions (the_pred_data thy name)) mode)
+
+fun the_sizelim_function_data thy name mode = case lookup_sizelim_function_data thy name mode
+  of NONE => error ("No size-limited function defined for mode " ^ string_of_mode mode
+    ^ " of predicate " ^ name)
+   | SOME data => data
+
+val sizelim_function_name_of = #name ooo the_sizelim_function_data
+
+(*val generator_modes_of = (map fst) o #generators oo the_pred_data*)
+     
+(* diagnostic display functions *)
+
+fun print_modes modes = Output.tracing ("Inferred modes:\n" ^
+  cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
+    string_of_mode ms)) modes));
+
+fun print_pred_mode_table string_of_entry thy pred_mode_table =
+  let
+    fun print_mode pred (mode, entry) =  "mode : " ^ (string_of_mode mode)
+      ^ (string_of_entry pred mode entry)  
+    fun print_pred (pred, modes) =
+      "predicate " ^ pred ^ ": " ^ cat_lines (map (print_mode pred) modes)
+    val _ = Output.tracing (cat_lines (map print_pred pred_mode_table))
+  in () end;
+
+fun string_of_moded_prem thy (Prem (ts, p), tmode) =
+    (Syntax.string_of_term_global thy (list_comb (p, ts))) ^
+    "(" ^ (string_of_tmode tmode) ^ ")"
+  | string_of_moded_prem thy (GeneratorPrem (ts, p), Mode (predmode, is, _)) =
+    (Syntax.string_of_term_global thy (list_comb (p, ts))) ^
+    "(generator_mode: " ^ (string_of_mode predmode) ^ ")"
+  | string_of_moded_prem thy (Generator (v, T), _) =
+    "Generator for " ^ v ^ " of Type " ^ (Syntax.string_of_typ_global thy T)
+  | string_of_moded_prem thy (Negprem (ts, p), Mode (_, is, _)) =
+    (Syntax.string_of_term_global thy (list_comb (p, ts))) ^
+    "(negative mode: " ^ string_of_smode is ^ ")"
+  | string_of_moded_prem thy (Sidecond t, Mode (_, is, _)) =
+    (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        
+    fun string_of_clause pred mode clauses =
+      cat_lines (map (fn (ts, prems) => (space_implode " --> "
+        (map (string_of_moded_prem thy) prems)) ^ " --> " ^ pred ^ " "
+        ^ (space_implode " " (map (Syntax.string_of_term_global thy) ts))) clauses)
+  in print_pred_mode_table string_of_clause thy end;
+
+fun print_compiled_terms thy =
+  print_pred_mode_table (fn _ => fn _ => Syntax.string_of_term_global thy) thy
+    
+fun print_stored_rules thy =
+  let
+    val preds = (Graph.keys o PredData.get) thy
+    fun print pred () = let
+      val _ = writeln ("predicate: " ^ pred)
+      val _ = writeln ("number of parameters: " ^ string_of_int (nparams_of thy pred))
+      val _ = writeln ("introrules: ")
+      val _ = fold (fn thm => fn u => writeln (Display.string_of_thm_global thy thm))
+        (rev (intros_of thy pred)) ()
+    in
+      if (has_elim thy pred) then
+        writeln ("elimrule: " ^ Display.string_of_thm_global thy (the_elim_of thy pred))
+      else
+        writeln ("no elimrule defined")
+    end
+  in
+    fold print preds ()
+  end;
+
+fun print_all_modes thy =
+  let
+    val _ = writeln ("Inferred modes:")
+    fun print (pred, modes) u =
+      let
+        val _ = writeln ("predicate: " ^ pred)
+        val _ = writeln ("modes: " ^ (commas (map string_of_mode modes)))
+      in u end  
+  in
+    fold print (all_modes_of thy) ()
+  end
+  
+(** preprocessing rules **)  
+
+fun imp_prems_conv cv ct =
+  case Thm.term_of ct of
+    Const ("==>", _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv) (imp_prems_conv cv) ct
+  | _ => Conv.all_conv ct
+
+fun Trueprop_conv cv ct =
+  case Thm.term_of ct of
+    Const ("Trueprop", _) $ _ => Conv.arg_conv cv ct  
+  | _ => error "Trueprop_conv"
+
+fun preprocess_intro thy rule =
+  Conv.fconv_rule
+    (imp_prems_conv
+      (Trueprop_conv (Conv.try_conv (Conv.rewr_conv (Thm.symmetric @{thm Predicate.eq_is_eq})))))
+    (Thm.transfer thy rule)
+
+fun preprocess_elim thy nparams elimrule =
+  let
+    val _ = Output.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 prems = Thm.prems_of elimrule
+    val nargs = length (snd (strip_comb (HOLogic.dest_Trueprop (hd prems)))) - nparams
+    fun preprocess_case t =
+     let
+       val params = Logic.strip_params t
+       val (assums1, assums2) = chop nargs (Logic.strip_assums_hyp t)
+       val assums_hyp' = assums1 @ (map replace_eqs assums2)
+     in
+       list_all (params, Logic.list_implies (assums_hyp', Logic.strip_assums_concl t))
+     end
+    val cases' = map preprocess_case (tl prems)
+    val elimrule' = Logic.list_implies ((hd prems) :: cases', Thm.concl_of elimrule)
+    (*val _ =  Output.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 _ = Output.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 _ = Output.tracing "Preprocessed elimination rule"
+  in
+    Thm.equal_elim bigeq elimrule
+  end;
+
+(* special case: predicate with no introduction rule *)
+fun noclause thy predname elim = let
+  val T = (Logic.unvarifyT o Sign.the_const_type thy) predname
+  val Ts = binder_types T
+  val names = Name.variant_list []
+        (map (fn i => "x" ^ (string_of_int i)) (1 upto (length Ts)))
+  val vs = map2 (curry Free) names Ts
+  val clausehd = HOLogic.mk_Trueprop (list_comb (Const (predname, T), vs))
+  val intro_t = Logic.mk_implies (@{prop False}, clausehd)
+  val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT))
+  val elim_t = Logic.list_implies ([clausehd, Logic.mk_implies (@{prop False}, P)], P)
+  val intro = Goal.prove (ProofContext.init thy) names [] intro_t
+        (fn {...} => etac @{thm FalseE} 1)
+  val elim = Goal.prove (ProofContext.init thy) ("P" :: names) [] elim_t
+        (fn {...} => etac elim 1) 
+in
+  ([intro], elim)
+end
+
+fun fetch_pred_data thy name =
+  case try (Inductive.the_inductive (ProofContext.init thy)) name of
+    SOME (info as (_, result)) => 
+      let
+        fun is_intro_of intro =
+          let
+            val (const, _) = strip_comb (HOLogic.dest_Trueprop (concl_of intro))
+          in (fst (dest_Const const) = name) end;      
+        val intros = ind_set_codegen_preproc thy ((map (preprocess_intro thy))
+          (filter is_intro_of (#intrs result)))
+        val pre_elim = nth (#elims result) (find_index (fn s => s = name) (#names (fst info)))
+        val nparams = length (Inductive.params_of (#raw_induct result))
+        val elim = singleton (ind_set_codegen_preproc thy) (preprocess_elim thy nparams pre_elim)
+        val (intros, elim) = if null intros then noclause thy name elim else (intros, elim)
+      in
+        mk_pred_data ((intros, SOME elim, nparams), ([], [], []))
+      end                                                                    
+  | NONE => error ("No such predicate: " ^ quote name)
+  
+(* updaters *)
+
+fun apfst3 f (x, y, z) =  (f x, y, z)
+fun apsnd3 f (x, y, z) =  (x, f y, z)
+fun aptrd3 f (x, y, z) =  (x, y, f z)
+
+fun add_predfun name mode data =
+  let
+    val add = (apsnd o apfst3 o cons) (mode, mk_predfun_data data)
+  in PredData.map (Graph.map_node name (map_pred_data add)) end
+
+fun is_inductive_predicate thy name =
+  is_some (try (Inductive.the_inductive (ProofContext.init thy)) name)
+
+fun depending_preds_of thy (key, value) =
+  let
+    val intros = (#intros o rep_pred_data) value
+  in
+    fold Term.add_const_names (map Thm.prop_of intros) []
+      |> filter (fn c => (not (c = key)) andalso (is_inductive_predicate thy c orelse is_registered thy c))
+  end;
+
+
+(* code dependency graph *)
+(*
+fun dependencies_of thy name =
+  let
+    val (intros, elim, nparams) = fetch_pred_data thy name 
+    val data = mk_pred_data ((intros, SOME elim, nparams), ([], [], []))
+    val keys = depending_preds_of thy intros
+  in
+    (data, keys)
+  end;
+*)
+(* guessing number of parameters *)
+fun find_indexes pred xs =
+  let
+    fun find is n [] = is
+      | find is n (x :: xs) = find (if pred x then (n :: is) else is) (n + 1) xs;
+  in rev (find [] 0 xs) end;
+
+fun guess_nparams T =
+  let
+    val argTs = binder_types T
+    val nparams = fold (curry Int.max)
+      (map (fn x => x + 1) (find_indexes is_predT argTs)) 0
+  in nparams end;
+
+fun add_intro thm thy = let
+   val (name, T) = dest_Const (fst (strip_intro_concl 0 (prop_of thm)))
+   fun cons_intro gr =
+     case try (Graph.get_node gr) name of
+       SOME pred_data => Graph.map_node name (map_pred_data
+         (apfst (fn (intro, elim, nparams) => (thm::intro, elim, nparams)))) gr
+     | NONE =>
+       let
+         val nparams = the_default (guess_nparams T)  (try (#nparams o rep_pred_data o (fetch_pred_data thy)) name)
+       in Graph.new_node (name, mk_pred_data (([thm], NONE, nparams), ([], [], []))) gr end;
+  in PredData.map cons_intro thy end
+
+fun set_elim thm = let
+    val (name, _) = dest_Const (fst 
+      (strip_comb (HOLogic.dest_Trueprop (hd (prems_of thm)))))
+    fun set (intros, _, nparams) = (intros, SOME thm, nparams)  
+  in PredData.map (Graph.map_node name (map_pred_data (apfst set))) end
+
+fun set_nparams name nparams = let
+    fun set (intros, elim, _ ) = (intros, elim, nparams) 
+  in PredData.map (Graph.map_node name (map_pred_data (apfst set))) end
+    
+fun register_predicate (pre_intros, pre_elim, nparams) thy =
+  let
+    val (name, _) = dest_Const (fst (strip_intro_concl nparams (prop_of (hd pre_intros))))
+    (* preprocessing *)
+    val intros = ind_set_codegen_preproc thy (map (preprocess_intro thy) pre_intros)
+    val elim = singleton (ind_set_codegen_preproc thy) (preprocess_elim thy nparams pre_elim)
+  in
+    if not (member (op =) (Graph.keys (PredData.get thy)) name) then
+      PredData.map
+        (Graph.new_node (name, mk_pred_data ((intros, SOME elim, nparams), ([], [], [])))) thy
+    else thy
+  end
+
+fun register_intros pre_intros thy =
+  let
+    val (c, T) = dest_Const (fst (strip_intro_concl 0 (prop_of (hd pre_intros))))
+    val _ = Output.tracing ("Registering introduction rules of " ^ c)
+    val _ = Output.tracing (commas (map (Display.string_of_thm_global thy) pre_intros))
+    val nparams = guess_nparams T
+    val pre_elim = 
+      (Drule.standard o (setmp quick_and_dirty true (SkipProof.make_thm thy)))
+      (mk_casesrule (ProofContext.init thy) nparams pre_intros)
+  in register_predicate (pre_intros, pre_elim, nparams) thy end
+
+fun set_generator_name pred mode name = 
+  let
+    val set = (apsnd o apsnd3 o cons) (mode, mk_function_data (name, NONE))
+  in
+    PredData.map (Graph.map_node pred (map_pred_data set))
+  end
+
+fun set_sizelim_function_name pred mode name = 
+  let
+    val set = (apsnd o aptrd3 o cons) (mode, mk_function_data (name, NONE))
+  in
+    PredData.map (Graph.map_node pred (map_pred_data set))
+  end
+
+(** data structures for generic compilation for different monads **)
+
+(* maybe rename functions more generic:
+  mk_predT -> mk_monadT; dest_predT -> dest_monadT
+  mk_single -> mk_return (?)
+*)
+datatype compilation_funs = CompilationFuns of {
+  mk_predT : typ -> typ,
+  dest_predT : typ -> typ,
+  mk_bot : typ -> term,
+  mk_single : term -> term,
+  mk_bind : term * term -> term,
+  mk_sup : term * term -> term,
+  mk_if : term -> term,
+  mk_not : term -> term,
+(*  funT_of : mode -> typ -> typ, *)
+(*  mk_fun_of : theory -> (string * typ) -> mode -> term, *) 
+  mk_map : typ -> typ -> term -> term -> term,
+  lift_pred : term -> term
+};
+
+fun mk_predT (CompilationFuns funs) = #mk_predT funs
+fun dest_predT (CompilationFuns funs) = #dest_predT funs
+fun mk_bot (CompilationFuns funs) = #mk_bot funs
+fun mk_single (CompilationFuns funs) = #mk_single funs
+fun mk_bind (CompilationFuns funs) = #mk_bind funs
+fun mk_sup (CompilationFuns funs) = #mk_sup funs
+fun mk_if (CompilationFuns funs) = #mk_if funs
+fun mk_not (CompilationFuns funs) = #mk_not funs
+(*fun funT_of (CompilationFuns funs) = #funT_of funs*)
+(*fun mk_fun_of (CompilationFuns funs) = #mk_fun_of funs*)
+fun mk_map (CompilationFuns funs) = #mk_map funs
+fun lift_pred (CompilationFuns funs) = #lift_pred funs
+
+fun funT_of compfuns (iss, is) T =
+  let
+    val Ts = binder_types T
+    val (paramTs, (inargTs, outargTs)) = split_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))
+  end;
+
+fun mk_fun_of compfuns thy (name, T) mode = 
+  Const (predfun_name_of thy name mode, funT_of compfuns mode T)
+
+
+structure PredicateCompFuns =
+struct
+
+fun mk_predT T = Type (@{type_name "Predicate.pred"}, [T])
+
+fun dest_predT (Type (@{type_name "Predicate.pred"}, [T])) = T
+  | dest_predT T = raise TYPE ("dest_predT", [T], []);
+
+fun mk_bot T = Const (@{const_name Orderings.bot}, mk_predT T);
+
+fun mk_single t =
+  let val T = fastype_of t
+  in Const(@{const_name Predicate.single}, T --> mk_predT T) $ t end;
+
+fun mk_bind (x, f) =
+  let val T as Type ("fun", [_, U]) = fastype_of f
+  in
+    Const (@{const_name Predicate.bind}, fastype_of x --> T --> U) $ x $ f
+  end;
+
+val mk_sup = HOLogic.mk_binop @{const_name sup};
+
+fun mk_if cond = Const (@{const_name Predicate.if_pred},
+  HOLogic.boolT --> mk_predT HOLogic.unitT) $ cond;
+
+fun mk_not t = let val T = mk_predT HOLogic.unitT
+  in Const (@{const_name Predicate.not_pred}, T --> T) $ t end
+
+fun mk_Enum f =
+  let val T as Type ("fun", [T', _]) = fastype_of f
+  in
+    Const (@{const_name Predicate.Pred}, T --> mk_predT T') $ f    
+  end;
+
+fun mk_Eval (f, x) =
+  let
+    val T = fastype_of x
+  in
+    Const (@{const_name Predicate.eval}, mk_predT T --> T --> HOLogic.boolT) $ f $ x
+  end;
+
+fun mk_map T1 T2 tf tp = Const (@{const_name Predicate.map},
+  (T1 --> T2) --> mk_predT T1 --> mk_predT T2) $ tf $ tp;
+
+val lift_pred = I
+
+val compfuns = CompilationFuns {mk_predT = mk_predT, dest_predT = dest_predT, mk_bot = mk_bot,
+  mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if, mk_not = mk_not,
+  mk_map = mk_map, lift_pred = lift_pred};
+
+end;
+
+structure RPredCompFuns =
+struct
+
+fun mk_rpredT T =
+  @{typ "Random.seed"} --> HOLogic.mk_prodT (PredicateCompFuns.mk_predT T, @{typ "Random.seed"})
+
+fun dest_rpredT (Type ("fun", [_,
+  Type (@{type_name "*"}, [Type (@{type_name "Predicate.pred"}, [T]), _])])) = T
+  | dest_rpredT T = raise TYPE ("dest_rpredT", [T], []); 
+
+fun mk_bot T = Const(@{const_name RPred.bot}, mk_rpredT T)
+
+fun mk_single t =
+  let
+    val T = fastype_of t
+  in
+    Const (@{const_name RPred.return}, T --> mk_rpredT T) $ t
+  end;
+
+fun mk_bind (x, f) =
+  let
+    val T as (Type ("fun", [_, U])) = fastype_of f
+  in
+    Const (@{const_name RPred.bind}, fastype_of x --> T --> U) $ x $ f
+  end
+
+val mk_sup = HOLogic.mk_binop @{const_name RPred.supp}
+
+fun mk_if cond = Const (@{const_name RPred.if_rpred},
+  HOLogic.boolT --> mk_rpredT HOLogic.unitT) $ cond;
+
+fun mk_not t = error "Negation is not defined for RPred"
+
+fun mk_map t = error "FIXME" (*FIXME*)
+
+fun lift_pred t =
+  let
+    val T = PredicateCompFuns.dest_predT (fastype_of t)
+    val lift_predT = PredicateCompFuns.mk_predT T --> mk_rpredT T 
+  in
+    Const (@{const_name "RPred.lift_pred"}, lift_predT) $ t  
+  end;
+
+val compfuns = CompilationFuns {mk_predT = mk_rpredT, dest_predT = dest_rpredT, mk_bot = mk_bot,
+    mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if, mk_not = mk_not,
+    mk_map = mk_map, lift_pred = lift_pred};
+
+end;
+(* for external use with interactive mode *)
+val pred_compfuns = PredicateCompFuns.compfuns
+val rpred_compfuns = RPredCompFuns.compfuns;
+
+fun lift_random random =
+  let
+    val T = dest_randomT (fastype_of random)
+  in
+    Const (@{const_name lift_random}, (@{typ Random.seed} -->
+      HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed})) --> 
+      RPredCompFuns.mk_rpredT T) $ random
+  end;
+
+fun sizelim_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 
+  in
+    (paramTs' @ inargTs @ [@{typ "code_numeral"}]) ---> (mk_predT compfuns (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_generator_of compfuns thy (name, T) mode = 
+  Const (generator_name_of thy name mode, sizelim_funT_of compfuns mode T)
+
+(* Mode analysis *)
+
+(*** check if a term contains only constructor functions ***)
+fun is_constrt thy =
+  let
+    val cnstrs = flat (maps
+      (map (fn (_, (Tname, _, cs)) => map (apsnd (rpair Tname o length)) cs) o #descr o snd)
+      (Symtab.dest (Datatype.get_all thy)));
+    fun check t = (case strip_comb t of
+        (Free _, []) => true
+      | (Const (s, T), ts) => (case (AList.lookup (op =) cnstrs s, body_type T) of
+            (SOME (i, Tname), Type (Tname', _)) => length ts = i andalso Tname = Tname' andalso forall check ts
+          | _ => false)
+      | _ => false)
+  in check end;
+
+(*** check if a type is an equality type (i.e. doesn't contain fun)
+  FIXME this is only an approximation ***)
+fun is_eqT (Type (s, Ts)) = s <> "fun" andalso forall is_eqT Ts
+  | is_eqT _ = true;
+
+fun term_vs tm = fold_aterms (fn Free (x, T) => cons x | _ => I) tm [];
+val terms_vs = distinct (op =) o maps term_vs;
+
+(** collect all Frees in a term (with duplicates!) **)
+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 [[]];
+     
+(* FIXME: should be in library - cprod = map_prod I *)
+fun cprod ([], ys) = []
+  | cprod (x :: xs, ys) = map (pair x) ys @ cprod (xs, ys);
+
+fun cprods xss = foldr (map op :: o cprod) [[]] xss;
+
+fun cprods_subset [] = [[]]
+  | cprods_subset (xs :: xss) =
+  let
+    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 default = [Mode (([], ks), ks, [])];
+    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 prfx = map (rpair NONE) (1 upto k)
+        in
+          if not (is_prefix op = prfx is) then [] else
+          let val is' = List.drop (is, k)
+          in 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)))
+          end
+        end)) (AList.lookup op = modes name)
+
+  in
+    case strip_comb (Envir.eta_contract t) of
+      (Const (name, _), args) => the_default default (mk_modes name args)
+    | (Var ((name, _), _), args) => the (mk_modes name args)
+    | (Free (name, _), args) => the (mk_modes name args)
+    | (Abs _, []) => error "Abs at param position" (* modes_of_param default modes t *)
+    | _ => default
+  end
+  
+fun select_mode_prem thy modes vs ps =
+  find_first (is_some o snd) (ps ~~ map
+    (fn Prem (us, t) => find_first (fn Mode (_, is, _) =>
+          let
+            val (in_ts, out_ts) = split_smode is us;
+            val (out_ts', in_ts') = List.partition (is_constrt thy) out_ts;
+            val vTs = maps term_vTs out_ts';
+            val dupTs = map snd (duplicates (op =) vTs) @
+              List.mapPartial (AList.lookup (op =) vTs) vs;
+          in
+            terms_vs (in_ts @ in_ts') subset vs andalso
+            forall (is_eqT o fastype_of) in_ts' andalso
+            term_vs t subset vs andalso
+            forall is_eqT dupTs
+          end)
+            (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
+            terms_vs us subset vs andalso
+            term_vs t subset vs)
+            (modes_of_term modes t handle Option =>
+               error ("Bad predicate: " ^ Syntax.string_of_term_global thy t))
+      | Sidecond t => if term_vs t subset vs then SOME (Mode (([], []), [], []))
+          else NONE
+      ) ps);
+
+fun fold_prem f (Prem (args, _)) = fold f args
+  | fold_prem f (Negprem (args, _)) = fold f args
+  | fold_prem f (Sidecond t) = f t
+
+fun all_subsets [] = [[]]
+  | all_subsets (x::xs) = let val xss' = all_subsets xs in xss' @ (map (cons x) xss') end
+
+fun generator vTs v = 
+  let
+    val T = the (AList.lookup (op =) vTs v)
+  in
+    (Generator (v, T), Mode (([], []), [], []))
+  end;
+
+fun gen_prem (Prem (us, t)) = GeneratorPrem (us, t)
+  | gen_prem (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 _ = Output.tracing ("param_vs:" ^ commas param_vs)
+  val _ = Output.tracing ("iss:" ^
+    commas (map (fn is => case is of SOME is => string_of_smode is | NONE => "NONE") iss))
+    *)
+    val modes' = modes @ List.mapPartial
+      (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
+        (param_vs ~~ iss);
+    val gen_modes' = gen_modes @ List.mapPartial
+      (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
+        (param_vs ~~ iss);  
+    val vTs = distinct (op =) ((fold o fold_prem) Term.add_frees ps (fold Term.add_frees ts []))
+    val prem_vs = distinct (op =) ((fold o fold_prem) Term.add_free_names ps [])
+    fun check_mode_prems acc_ps vs [] = SOME (acc_ps, vs)
+      | check_mode_prems acc_ps vs ps = (case select_mode_prem thy modes' vs ps of
+          NONE =>
+            (if with_generator then
+              (case select_mode_prem thy gen_modes' vs ps of
+                SOME (p as Prem _, SOME mode) => check_mode_prems ((gen_prem p, mode) :: acc_ps) 
+                  (case p of Prem (us, _) => vs union terms_vs us | _ => vs)
+                  (filter_out (equal p) ps)
+              | _ =>
+                  let 
+                    val all_generator_vs = all_subsets (prem_vs \\ vs) |> sort (int_ord o (pairself length))
+                  in
+                    case (find_first (fn generator_vs => is_some
+                      (select_mode_prem thy modes' (vs union generator_vs) ps)) all_generator_vs) of
+                      SOME generator_vs => check_mode_prems ((map (generator vTs) generator_vs) @ acc_ps)
+                        (vs union generator_vs) ps
+                    | NONE => let
+                    val _ = Output.tracing ("ps:" ^ (commas
+                    (map (fn p => string_of_moded_prem thy (p, Mode (([], []), [], []))) ps)))
+                  in (*error "mode analysis failed"*)NONE end
+                  end)
+            else
+              NONE)
+        | SOME (p, SOME mode) => check_mode_prems ((if with_generator then param_gen_prem param_vs p else p, mode) :: acc_ps) 
+            (case p of Prem (us, _) => vs union terms_vs us | _ => vs)
+            (filter_out (equal p) ps))
+    val (in_ts, in_ts') = List.partition (is_constrt thy) (fst (split_smode is ts));
+    val in_vs = terms_vs in_ts;
+    val concl_vs = terms_vs ts
+  in
+    if forall is_eqT (map snd (duplicates (op =) (maps term_vTs in_ts))) andalso
+    forall (is_eqT o fastype_of) in_ts' then
+      case check_mode_prems [] (param_vs union in_vs) ps of
+         NONE => NONE
+       | SOME (acc_ps, vs) =>
+         if with_generator then
+           SOME (ts, (rev acc_ps) @ (map (generator vTs) (concl_vs \\ vs))) 
+         else
+           if concl_vs subset vs then SOME (ts, rev acc_ps) else NONE
+    else NONE
+  end;
+
+fun check_modes_pred with_generator thy param_vs clauses modes gen_modes (p, ms) =
+  let val SOME rs = AList.lookup (op =) clauses p
+  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 => (Output.tracing ("Clause " ^ string_of_int (i + 1) ^ " of " ^
+      p ^ " violates mode " ^ string_of_mode m);
+        Output.tracing (commas (map (Syntax.string_of_term_global thy) (fst (nth 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 
+  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 =
+  let
+    val modes =
+      fixp (fn modes =>
+        map (check_modes_pred 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
+  end;
+
+fun remove_from rem [] = []
+  | remove_from rem ((k, vs) :: xs) =
+    (case AList.lookup (op =) rem k of
+      NONE => (k, vs)
+    | SOME vs' => (k, vs \\ vs'))
+    :: remove_from rem xs
+    
+fun infer_modes_with_generator thy extra_modes all_modes param_vs clauses =
+  let
+    val prednames = map fst clauses
+    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 modes =
+      fixp (fn modes =>
+        map (check_modes_pred 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
+  end;
+
+(* term construction *)
+
+fun mk_v (names, vs) s T = (case AList.lookup (op =) vs s of
+      NONE => (Free (s, T), (names, (s, [])::vs))
+    | SOME xs =>
+        let
+          val s' = Name.variant names s;
+          val v = Free (s', T)
+        in
+          (v, (s'::names, AList.update (op =) (s, v::xs) vs))
+        end);
+
+fun distinct_v (Free (s, T)) nvs = mk_v nvs s T
+  | distinct_v (t $ u) nvs =
+      let
+        val (t', nvs') = distinct_v t nvs;
+        val (u', nvs'') = distinct_v u nvs';
+      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 =
+	let
+    val Ts = binder_types T
+    (*val argnames = Name.variant_list names
+        (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
+    val args = map Free (argnames ~~ Ts)
+    val (inargs, outargs) = split_smode mode args*)
+		fun mk_split_lambda [] t = lambda (Free (Name.variant names "x", HOLogic.unitT)) t
+			| mk_split_lambda [x] t = lambda x t
+			| mk_split_lambda xs t =
+			let
+				fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t))
+					| mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t))
+			in
+				mk_split_lambda' xs t
+			end;
+  	fun mk_arg (i, T) =
+		  let
+	  	  val vname = Name.variant names ("x" ^ string_of_int i)
+		    val default = Free (vname, T)
+		  in 
+		    case AList.lookup (op =) mode i of
+		      NONE => (([], [default]), [default])
+			  | SOME NONE => (([default], []), [default])
+			  | SOME (SOME pis) =>
+				  case HOLogic.strip_tupleT T of
+						[] => error "pair mode but unit tuple" (*(([default], []), [default])*)
+					| [_] => error "pair mode but not a tuple" (*(([default], []), [default])*)
+					| Ts =>
+					  let
+							val vnames = Name.variant_list names
+								(map (fn j => "x" ^ string_of_int i ^ "p" ^ string_of_int j)
+									(1 upto length Ts))
+							val args = map Free (vnames ~~ Ts)
+							fun split_args (i, arg) (ins, outs) =
+							  if member (op =) pis i then
+							    (arg::ins, outs)
+								else
+								  (ins, arg::outs)
+							val (inargs, outargs) = fold_rev split_args ((1 upto length Ts) ~~ args) ([], [])
+							fun tuple args = if null args then [] else [HOLogic.mk_tuple args]
+						in ((tuple inargs, tuple outargs), args) end
+			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 t = fold_rev mk_split_lambda args r
+  in
+    (t, names)
+  end;
+
+fun compile_arg size 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
+        | 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) =
+  let
+    fun check_constrt t (names, eqs) =
+      if is_constrt thy t then (t, (names, eqs)) else
+        let
+          val s = Name.variant names "x";
+          val v = Free (s, fastype_of t)
+        in (v, (s::names, HOLogic.mk_eq (v, t)::eqs)) end;
+
+    val (in_ts, out_ts) = split_smode is ts;
+    val (in_ts', (all_vs', eqs)) =
+      fold_map check_constrt in_ts (all_vs, []);
+
+    fun compile_prems out_ts' vs names [] =
+          let
+            val (out_ts'', (names', eqs')) =
+              fold_map check_constrt out_ts' (names, []);
+            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)
+          end
+      | compile_prems out_ts vs names ((p, mode as Mode ((_, is), _, _)) :: ps) =
+          let
+            val vs' = distinct (op =) (flat (vs :: map term_vs out_ts));
+            val (out_ts', (names', eqs)) =
+              fold_map check_constrt out_ts (names, [])
+            val (out_ts'', (names'', constr_vs')) = fold_map distinct_v
+              out_ts' ((names', map (rpair []) vs))
+            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 rest = compile_prems out_ts''' vs' names'' ps
+                 in
+                   (u, rest)
+                 end
+             | 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 rest = compile_prems out_ts''' vs' names'' ps
+                 in
+                   (u, rest)
+                 end
+             | Sidecond t =>
+                 let
+                   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 rest = compile_prems [Free (v, T)]  vs' names'' ps;
+                 in
+                   (u, rest)
+                 end
+          in
+            compile_match thy compfuns constr_vs' eqs out_ts'' 
+              (mk_bind compfuns (compiled_clause, rest))
+          end
+    val prem_t = compile_prems in_ts' param_vs all_vs' moded_ps;
+  in
+    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 =
+  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) =
+		    [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"
+					   | [T] => [Free (Name.variant (all_vs @ param_vs) ("x" ^ string_of_int i), nth Ts2 (i - 1))]
+						 | Ts => let
+							 val vnames = Name.variant_list (all_vs @ param_vs)
+								(map (fn j => "x" ^ string_of_int i ^ "p" ^ string_of_int j)
+									pis)
+						 in if null pis then []
+						   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 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)
+  in
+    HOLogic.mk_Trueprop (HOLogic.mk_eq eq)
+  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))
+	setSolver (mk_solver "True_solver" (fn _ => rtac @{thm TrueI}))
+
+(* Definition of executable functions and their intro and elim rules *)
+
+fun print_arities arities = tracing ("Arities:\n" ^
+  cat_lines (map (fn (s, (ks, k)) => s ^ ": " ^
+    space_implode " -> " (map
+      (fn NONE => "X" | SOME k' => string_of_int k')
+        (ks @ [SOME k]))) arities));
+
+fun create_intro_elim_rule (mode as (iss, is)) defthm mode_id funT pred thy =
+let
+  val Ts = binder_types (fastype_of pred)
+  val funtrm = Const (mode_id, funT)
+  val (Ts1, Ts2) = chop (length iss) Ts;
+  val Ts1' = map2 (fn NONE => I | SOME is => funT_of (PredicateCompFuns.compfuns) ([], is)) iss Ts1
+	val param_names = Name.variant_list []
+    (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts1)));
+  val params = map Free (param_names ~~ Ts1')
+	fun mk_args (i, T) argnames =
+    let
+		  val vname = Name.variant (param_names @ argnames) ("x" ^ string_of_int (length Ts1' + i))
+		  val default = (Free (vname, T), vname :: argnames)
+	  in
+  	  case AList.lookup (op =) is i of
+						 NONE => default
+					 | SOME NONE => default
+        	 | SOME (SOME pis) =>
+					   case HOLogic.strip_tupleT T of
+						   [] => default
+					   | [_] => default
+						 | Ts => 
+						let
+							val vnames = Name.variant_list (param_names @ argnames)
+								(map (fn j => "x" ^ string_of_int (length Ts1' + i) ^ "p" ^ string_of_int j)
+									(1 upto (length Ts)))
+						 in (HOLogic.mk_tuple (map Free (vnames ~~ Ts)), vnames  @ argnames) end
+		end
+	val (args, argnames) = fold_map mk_args (1 upto (length Ts2) ~~ Ts2) []
+  val (inargs, outargs) = split_smode is args
+  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 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))
+  val funpropI = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, funargs),
+                   mk_tuple outargs))
+  val introtrm = Logic.list_implies (predpropI :: param_eqs, funpropI)
+  val simprules = [defthm, @{thm eval_pred},
+	  @{thm "split_beta"}, @{thm "fst_conv"}, @{thm "snd_conv"}, @{thm pair_collapse}]
+  val unfolddef_tac = Simplifier.asm_full_simp_tac (HOL_basic_ss addsimps simprules) 1
+  val introthm = Goal.prove (ProofContext.init thy) (argnames @ param_names @ param_names' @ ["y"]) [] introtrm (fn {...} => unfolddef_tac)
+  val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT));
+  val elimtrm = Logic.list_implies ([funpropE, Logic.mk_implies (predpropE, P)], P)
+  val elimthm = Goal.prove (ProofContext.init thy) (argnames @ param_names @ param_names' @ ["y", "P"]) [] elimtrm (fn {...} => unfolddef_tac)
+in
+  (introthm, elimthm)
+end;
+
+fun create_constname_of_mode thy prefix name mode = 
+  let
+    fun string_of_mode mode = if null mode then "0"
+      else space_implode "_" (map (fn (i, NONE) => string_of_int i | (i, SOME pis) => string_of_int i ^ "p"
+        ^ space_implode "p" (map string_of_int pis)) mode)
+    val HOmode = space_implode "_and_"
+      (fold (fn NONE => I | SOME mode => cons (string_of_mode mode)) (fst mode) [])
+  in
+    (Sign.full_bname thy (prefix ^ (Long_Name.base_name name))) ^
+      (if HOmode = "" then "_" else "_for_" ^ HOmode ^ "_yields_") ^ (string_of_mode (snd mode))
+  end;
+
+fun split_tupleT is T =
+	let
+		fun split_tuple' _ _ [] = ([], [])
+			| split_tuple' is i (T::Ts) =
+			(if i mem is then apfst else apsnd) (cons T)
+				(split_tuple' is (i+1) Ts)
+	in
+	  split_tuple' is 1 (HOLogic.strip_tupleT T)
+  end
+	
+fun mk_arg xin xout pis T =
+  let
+	  val n = length (HOLogic.strip_tupleT T)
+		val ni = length pis
+	  fun mk_proj i j t =
+		  (if i = j then I else HOLogic.mk_fst)
+			  (funpow (i - 1) HOLogic.mk_snd t)
+	  fun mk_arg' i (si, so) = if i mem pis then
+		    (mk_proj si ni xin, (si+1, so))
+		  else
+			  (mk_proj so (n - ni) xout, (si, so+1))
+	  val (args, _) = fold_map mk_arg' (1 upto n) (1, 1)
+	in
+	  HOLogic.mk_tuple args
+	end
+
+fun create_definitions preds (name, modes) thy =
+  let
+    val compfuns = PredicateCompFuns.compfuns
+    val T = AList.lookup (op =) preds name |> the
+    fun create_definition (mode as (iss, is)) thy = let
+      val mode_cname = create_constname_of_mode thy "" name mode
+      val mode_cbasename = Long_Name.base_name mode_cname
+      val Ts = binder_types T
+      val (Ts1, Ts2) = chop (length iss) Ts
+      val (Us1, Us2) =  split_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 names = Name.variant_list []
+        (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
+			(* old *)
+			(*
+		  val xs = map Free (names ~~ (Ts1' @ Ts2))
+      val (xparams, xargs) = chop (length iss) xs
+      val (xins, xouts) = split_smode is xargs
+			*)
+			(* new *)
+			val param_names = Name.variant_list []
+			  (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts1')))
+		  val xparams = map Free (param_names ~~ Ts1')
+      fun mk_vars (i, T) names =
+			  let
+				  val vname = Name.variant names ("x" ^ string_of_int (length Ts1' + i))
+				in
+					case AList.lookup (op =) is i of
+						 NONE => ((([], [Free (vname, T)]), Free (vname, T)), vname :: names)
+					 | SOME NONE => ((([Free (vname, T)], []), Free (vname, T)), vname :: names)
+        	 | SOME (SOME pis) =>
+					   let
+						   val (Tins, Touts) = split_tupleT pis T
+							 val name_in = Name.variant names ("x" ^ string_of_int (length Ts1' + i) ^ "in")
+							 val name_out = Name.variant names ("x" ^ string_of_int (length Ts1' + i) ^ "out")
+						   val xin = Free (name_in, HOLogic.mk_tupleT Tins)
+							 val xout = Free (name_out, HOLogic.mk_tupleT Touts)
+							 val xarg = mk_arg xin xout pis T
+						 in (((if null Tins then [] else [xin], if null Touts then [] else [xout]), xarg), name_in :: name_out :: names) end
+						 end
+   	  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
+      fun mk_split_lambda [] t = lambda (Free (Name.variant names' "x", HOLogic.unitT)) t
+        | mk_split_lambda [x] t = lambda x t
+        | mk_split_lambda xs t =
+        let
+          fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t))
+            | mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t))
+        in
+          mk_split_lambda' xs t
+        end;
+      val predterm = PredicateCompFuns.mk_Enum (mk_split_lambda xouts
+        (list_comb (Const (name, T), xparams' @ xargs)))
+      val lhs = list_comb (Const (mode_cname, funT), xparams @ xins)
+      val def = Logic.mk_equals (lhs, predterm)
+      val ([definition], thy') = thy |>
+        Sign.add_consts_i [(Binding.name mode_cbasename, funT, NoSyn)] |>
+        PureThy.add_defs false [((Binding.name (mode_cbasename ^ "_def"), def), [])]
+      val (intro, elim) =
+        create_intro_elim_rule mode definition mode_cname funT (Const (name, T)) thy'
+      in thy'
+			  |> add_predfun name mode (mode_cname, definition, intro, elim)
+        |> PureThy.store_thm (Binding.name (mode_cbasename ^ "I"), intro) |> snd
+        |> PureThy.store_thm (Binding.name (mode_cbasename ^ "E"), elim)  |> snd
+        |> Theory.checkpoint
+      end;
+  in
+    fold create_definition modes thy
+  end;
+
+fun sizelim_create_definitions preds (name, modes) thy =
+  let
+    val T = AList.lookup (op =) preds name |> the
+    fun create_definition mode thy =
+      let
+        val mode_cname = create_constname_of_mode thy "sizelim_" name mode
+        val funT = sizelim_funT_of PredicateCompFuns.compfuns mode T
+      in
+        thy |> Sign.add_consts_i [(Binding.name (Long_Name.base_name mode_cname), funT, NoSyn)]
+        |> set_sizelim_function_name name mode mode_cname 
+      end;
+  in
+    fold create_definition modes thy
+  end;
+
+fun generator_funT_of (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 
+  in
+    (paramTs' @ inargTs @ [@{typ "code_numeral"}]) ---> (mk_predT RPredCompFuns.compfuns (mk_tupleT outargTs))
+  end
+
+fun rpred_create_definitions preds (name, modes) thy =
+  let
+    val T = AList.lookup (op =) preds name |> the
+    fun create_definition mode thy =
+      let
+        val mode_cname = create_constname_of_mode thy "gen_" name mode
+        val funT = 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 
+      end;
+  in
+    fold create_definition modes thy
+  end;
+  
+(* Proving equivalence of term *)
+
+fun is_Type (Type _) = true
+  | is_Type _ = false
+
+(* returns true if t is an application of an datatype constructor *)
+(* which then consequently would be splitted *)
+(* else false *)
+fun is_constructor thy t =
+  if (is_Type (fastype_of t)) then
+    (case Datatype.get_info thy ((fst o dest_Type o fastype_of) t) of
+      NONE => false
+    | SOME info => (let
+      val constr_consts = maps (fn (_, (_, _, constrs)) => map fst constrs) (#descr info)
+      val (c, _) = strip_comb t
+      in (case c of
+        Const (name, _) => name mem_string constr_consts
+        | _ => false) end))
+  else false
+
+(* MAJOR FIXME:  prove_params should be simple
+ - different form of introrule for parameters ? *)
+fun prove_param thy (NONE, t) = TRY (rtac @{thm refl} 1)
+  | prove_param thy (m as SOME (Mode (mode, is, ms)), t) =
+  let
+    val  (f, args) = strip_comb (Envir.eta_contract t)
+    val (params, _) = chop (length ms) args
+    val f_tac = case f of
+      Const (name, T) => simp_tac (HOL_basic_ss addsimps 
+         ([@{thm eval_pred}, (predfun_definition_of thy name mode),
+         @{thm "split_eta"}, @{thm "split_beta"}, @{thm "fst_conv"},
+				 @{thm "snd_conv"}, @{thm pair_collapse}, @{thm "Product_Type.split_conv"}])) 1
+    | Free _ => TRY (rtac @{thm refl} 1)
+    | Abs _ => error "prove_param: No valid parameter term"
+  in
+    REPEAT_DETERM (etac @{thm thin_rl} 1)
+    THEN REPEAT_DETERM (rtac @{thm ext} 1)
+    THEN print_tac "prove_param"
+    THEN f_tac
+    THEN print_tac "after simplification in prove_args"
+    THEN (EVERY (map (prove_param thy) (ms ~~ params)))
+    THEN (REPEAT_DETERM (atac 1))
+  end
+
+fun prove_expr thy (Mode (mode, is, ms), t, us) (premposition : int) =
+  case strip_comb t of
+    (Const (name, T), args) =>  
+      let
+        val introrule = predfun_intro_of thy name mode
+        val (args1, args2) = chop (length ms) args
+      in
+        rtac @{thm bindI} 1
+        THEN print_tac "before intro rule:"
+        (* for the right assumption in first position *)
+        THEN rotate_tac premposition 1
+        THEN debug_tac (Display.string_of_thm (ProofContext.init thy) introrule)
+        THEN rtac introrule 1
+        THEN print_tac "after intro rule"
+        (* work with parameter arguments *)
+        THEN (atac 1)
+        THEN (print_tac "parameter goal")
+        THEN (EVERY (map (prove_param thy) (ms ~~ args1)))
+        THEN (REPEAT_DETERM (atac 1))
+      end
+  | _ => rtac @{thm bindI} 1
+	  THEN asm_full_simp_tac
+		  (HOL_basic_ss' addsimps [@{thm "split_eta"}, @{thm "split_beta"}, @{thm "fst_conv"},
+				 @{thm "snd_conv"}, @{thm pair_collapse}]) 1
+	  THEN (atac 1)
+	  THEN print_tac "after prove parameter call"
+		
+
+fun SOLVED tac st = FILTER (fn st' => nprems_of st' = nprems_of st - 1) tac st; 
+
+fun SOLVEDALL tac st = FILTER (fn st' => nprems_of st' = 0) tac st
+
+fun prove_match thy (out_ts : term list) = let
+  fun get_case_rewrite t =
+    if (is_constructor thy t) then let
+      val case_rewrites = (#case_rewrites (Datatype.the_info thy
+        ((fst o dest_Type o fastype_of) t)))
+      in case_rewrites @ (flat (map get_case_rewrite (snd (strip_comb t)))) end
+    else []
+  val simprules = @{thm "unit.cases"} :: @{thm "prod.cases"} :: (flat (map get_case_rewrite out_ts))
+(* replace TRY by determining if it necessary - are there equations when calling compile match? *)
+in
+   (* make this simpset better! *)
+  asm_full_simp_tac (HOL_basic_ss' addsimps simprules) 1
+  THEN print_tac "after prove_match:"
+  THEN (DETERM (TRY (EqSubst.eqsubst_tac (ProofContext.init thy) [0] [@{thm "HOL.if_P"}] 1
+         THEN (REPEAT_DETERM (rtac @{thm conjI} 1 THEN (SOLVED (asm_simp_tac HOL_basic_ss 1))))
+         THEN (SOLVED (asm_simp_tac HOL_basic_ss 1)))))
+  THEN print_tac "after if simplification"
+end;
+
+(* corresponds to compile_fun -- maybe call that also compile_sidecond? *)
+
+fun prove_sidecond thy modes t =
+  let
+    fun preds_of t nameTs = case strip_comb t of 
+      (f as Const (name, T), args) =>
+        if AList.defined (op =) modes name then (name, T) :: nameTs
+          else fold preds_of args nameTs
+      | _ => nameTs
+    val preds = preds_of t []
+    val defs = map
+      (fn (pred, T) => predfun_definition_of thy pred
+        ([], map (rpair NONE) (1 upto (length (binder_types T)))))
+        preds
+  in 
+    (* remove not_False_eq_True when simpset in prove_match is better *)
+    simp_tac (HOL_basic_ss addsimps
+      (@{thms "HOL.simp_thms"} @ (@{thm not_False_eq_True} :: @{thm eval_pred} :: defs))) 1 
+    (* need better control here! *)
+  end
+
+fun prove_clause thy nargs modes (iss, is) (_, clauses) (ts, moded_ps) =
+  let
+    val (in_ts, clause_out_ts) = split_smode is ts;
+    fun prove_prems out_ts [] =
+      (prove_match thy out_ts)
+			THEN print_tac "before simplifying assumptions"
+      THEN asm_full_simp_tac HOL_basic_ss' 1
+			THEN print_tac "before single intro rule"
+      THEN (rtac (if null clause_out_ts then @{thm singleI_unit} else @{thm singleI}) 1)
+    | prove_prems out_ts ((p, mode as Mode ((iss, is), _, param_modes)) :: ps) =
+      let
+        val premposition = (find_index (equal p) clauses) + nargs
+        val rest_tac = (case p of Prem (us, t) =>
+            let
+              val (_, out_ts''') = split_smode is us
+              val rec_tac = prove_prems out_ts''' ps
+            in
+              print_tac "before clause:"
+              THEN asm_simp_tac HOL_basic_ss 1
+              THEN print_tac "before prove_expr:"
+              THEN prove_expr thy (mode, t, us) premposition
+              THEN print_tac "after prove_expr:"
+              THEN rec_tac
+            end
+          | Negprem (us, t) =>
+            let
+              val (_, out_ts''') = split_smode is us
+              val rec_tac = prove_prems out_ts''' ps
+              val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE)
+              val (_, params) = strip_comb t
+            in
+              rtac @{thm bindI} 1
+              THEN (if (is_some name) then
+                  simp_tac (HOL_basic_ss addsimps [predfun_definition_of thy (the name) (iss, is)]) 1
+                  THEN rtac @{thm not_predI} 1
+                  THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
+                  THEN (REPEAT_DETERM (atac 1))
+                  (* FIXME: work with parameter arguments *)
+                  THEN (EVERY (map (prove_param thy) (param_modes ~~ params)))
+                else
+                  rtac @{thm not_predI'} 1)
+                  THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
+              THEN rec_tac
+            end
+          | Sidecond t =>
+           rtac @{thm bindI} 1
+           THEN rtac @{thm if_predI} 1
+           THEN print_tac "before sidecond:"
+           THEN prove_sidecond thy modes t
+           THEN print_tac "after sidecond:"
+           THEN prove_prems [] ps)
+      in (prove_match thy out_ts)
+          THEN rest_tac
+      end;
+    val prems_tac = prove_prems in_ts moded_ps
+  in
+    rtac @{thm bindI} 1
+    THEN rtac @{thm singleI} 1
+    THEN prems_tac
+  end;
+
+fun select_sup 1 1 = []
+  | select_sup _ 1 = [rtac @{thm supI1}]
+  | select_sup n i = (rtac @{thm supI2})::(select_sup (n - 1) (i - 1));
+
+fun prove_one_direction thy 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 etac (predfun_elim_of thy pred mode) 1
+    THEN etac pred_case_rule 1
+    THEN (EVERY (map
+           (fn i => EVERY' (select_sup (length moded_clauses) i) i) 
+             (1 upto (length moded_clauses))))
+    THEN (EVERY (map2 (prove_clause thy nargs modes mode) clauses moded_clauses))
+    THEN print_tac "proved one direction"
+  end;
+
+(** Proof in the other direction **)
+
+fun prove_match2 thy out_ts = let
+  fun split_term_tac (Free _) = all_tac
+    | split_term_tac t =
+      if (is_constructor thy t) then let
+        val info = Datatype.the_info thy ((fst o dest_Type o fastype_of) t)
+        val num_of_constrs = length (#case_rewrites info)
+        (* special treatment of pairs -- because of fishing *)
+        val split_rules = case (fst o dest_Type o fastype_of) t of
+          "*" => [@{thm prod.split_asm}] 
+          | _ => 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))
+        THEN (EVERY (map split_term_tac ts))
+      end
+    else all_tac
+  in
+    split_term_tac (mk_tuple out_ts)
+    THEN (DETERM (TRY ((Splitter.split_asm_tac [@{thm "split_if_asm"}] 1) THEN (etac @{thm botE} 2))))
+  end
+
+(* VERY LARGE SIMILIRATIY to function prove_param 
+-- join both functions
+*)
+(* TODO: remove function *)
+
+fun prove_param2 thy (NONE, t) = all_tac 
+  | prove_param2 thy (m as SOME (Mode (mode, is, ms)), t) = let
+    val  (f, args) = strip_comb (Envir.eta_contract t)
+    val (params, _) = chop (length ms) args
+    val f_tac = case f of
+        Const (name, T) => full_simp_tac (HOL_basic_ss addsimps 
+           (@{thm eval_pred}::(predfun_definition_of thy name mode)
+           :: @{thm "Product_Type.split_conv"}::[])) 1
+      | Free _ => all_tac
+      | _ => error "prove_param2: illegal parameter term"
+  in  
+    print_tac "before simplification in prove_args:"
+    THEN f_tac
+    THEN print_tac "after simplification in prove_args"
+    THEN (EVERY (map (prove_param2 thy) (ms ~~ params)))
+  end
+
+
+fun prove_expr2 thy (Mode (mode, is, ms), t) = 
+  (case strip_comb t of
+    (Const (name, T), args) =>
+      etac @{thm bindE} 1
+      THEN (REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"})))
+      THEN print_tac "prove_expr2-before"
+      THEN (debug_tac (Syntax.string_of_term_global thy
+        (prop_of (predfun_elim_of thy name mode))))
+      THEN (etac (predfun_elim_of thy name mode) 1)
+      THEN print_tac "prove_expr2"
+      THEN (EVERY (map (prove_param2 thy) (ms ~~ args)))
+      THEN print_tac "finished prove_expr2"      
+    | _ => etac @{thm bindE} 1)
+    
+(* FIXME: what is this for? *)
+(* replace defined by has_mode thy pred *)
+(* TODO: rewrite function *)
+fun prove_sidecond2 thy modes t = let
+  fun preds_of t nameTs = case strip_comb t of 
+    (f as Const (name, T), args) =>
+      if AList.defined (op =) modes name then (name, T) :: nameTs
+        else fold preds_of args nameTs
+    | _ => nameTs
+  val preds = preds_of t []
+  val defs = map
+    (fn (pred, T) => predfun_definition_of thy pred 
+      ([], map (rpair NONE) (1 upto (length (binder_types T)))))
+      preds
+  in
+   (* only simplify the one assumption *)
+   full_simp_tac (HOL_basic_ss' addsimps @{thm eval_pred} :: defs) 1 
+   (* need better control here! *)
+   THEN print_tac "after sidecond2 simplification"
+   end
+  
+fun prove_clause2 thy modes pred (iss, is) (ts, ps) i =
+  let
+    val pred_intro_rule = nth (intros_of thy pred) (i - 1)
+    val (in_ts, clause_out_ts) = split_smode is ts;
+    fun prove_prems2 out_ts [] =
+      print_tac "before prove_match2 - last call:"
+      THEN prove_match2 thy out_ts
+      THEN print_tac "after prove_match2 - last call:"
+      THEN (etac @{thm singleE} 1)
+      THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
+      THEN (asm_full_simp_tac HOL_basic_ss' 1)
+      THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
+      THEN (asm_full_simp_tac HOL_basic_ss' 1)
+      THEN SOLVED (print_tac "state before applying intro rule:"
+      THEN (rtac pred_intro_rule 1)
+      (* How to handle equality correctly? *)
+      THEN (print_tac "state before assumption matching")
+      THEN (REPEAT (atac 1 ORELSE 
+         (CHANGED (asm_full_simp_tac (HOL_basic_ss' addsimps
+					 [@{thm split_eta}, @{thm "split_beta"}, @{thm "fst_conv"}, @{thm "snd_conv"}, @{thm pair_collapse}]) 1)
+          THEN print_tac "state after simp_tac:"))))
+    | prove_prems2 out_ts ((p, mode as Mode ((iss, is), _, param_modes)) :: ps) =
+      let
+        val rest_tac = (case p of
+          Prem (us, t) =>
+          let
+            val (_, out_ts''') = split_smode is us
+            val rec_tac = prove_prems2 out_ts''' ps
+          in
+            (prove_expr2 thy (mode, t)) THEN rec_tac
+          end
+        | Negprem (us, t) =>
+          let
+            val (_, out_ts''') = split_smode is us
+            val rec_tac = prove_prems2 out_ts''' ps
+            val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE)
+            val (_, params) = strip_comb t
+          in
+            print_tac "before neg prem 2"
+            THEN etac @{thm bindE} 1
+            THEN (if is_some name then
+                full_simp_tac (HOL_basic_ss addsimps [predfun_definition_of thy (the name) (iss, is)]) 1 
+                THEN etac @{thm not_predE} 1
+                THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
+                THEN (EVERY (map (prove_param2 thy) (param_modes ~~ params)))
+              else
+                etac @{thm not_predE'} 1)
+            THEN rec_tac
+          end 
+        | Sidecond t =>
+          etac @{thm bindE} 1
+          THEN etac @{thm if_predE} 1
+          THEN prove_sidecond2 thy modes t 
+          THEN prove_prems2 [] ps)
+      in print_tac "before prove_match2:"
+         THEN prove_match2 thy out_ts
+         THEN print_tac "after prove_match2:"
+         THEN rest_tac
+      end;
+    val prems_tac = prove_prems2 in_ts ps 
+  in
+    print_tac "starting prove_clause2"
+    THEN etac @{thm bindE} 1
+    THEN (etac @{thm singleE'} 1)
+    THEN (TRY (etac @{thm Pair_inject} 1))
+    THEN print_tac "after singleE':"
+    THEN prems_tac
+  end;
+ 
+fun prove_other_direction thy modes pred mode moded_clauses =
+  let
+    fun prove_clause clause i =
+      (if i < length moded_clauses then etac @{thm supE} 1 else all_tac)
+      THEN (prove_clause2 thy modes pred mode clause i)
+  in
+    (DETERM (TRY (rtac @{thm unit.induct} 1)))
+     THEN (REPEAT_DETERM (CHANGED (rewtac @{thm split_paired_all})))
+     THEN (rtac (predfun_intro_of thy pred mode) 1)
+     THEN (REPEAT_DETERM (rtac @{thm refl} 2))
+     THEN (EVERY (map2 prove_clause moded_clauses (1 upto (length moded_clauses))))
+  end;
+
+(** proof procedure **)
+
+fun prove_pred thy clauses preds modes pred mode (moded_clauses, compiled_term) =
+  let
+    val ctxt = ProofContext.init thy
+    val clauses = the (AList.lookup (op =) clauses pred)
+  in
+    Goal.prove ctxt (Term.add_free_names compiled_term []) [] compiled_term
+      (if !do_proofs then
+        (fn _ =>
+        rtac @{thm pred_iffI} 1
+				THEN 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 _ => setmp quick_and_dirty true SkipProof.cheat_tac thy))
+  end;
+
+(* composition of mode inference, definition, compilation and proof *)
+
+(** auxillary combinators for table of preds and modes **)
+
+fun map_preds_modes f preds_modes_table =
+  map (fn (pred, modes) =>
+    (pred, map (fn (mode, value) => (mode, f pred mode value)) modes)) preds_modes_table
+
+fun join_preds_modes table1 table2 =
+  map_preds_modes (fn pred => fn mode => fn value =>
+    (value, the (AList.lookup (op =) (the (AList.lookup (op =) table2 pred)) mode))) table1
+    
+fun maps_modes preds_modes_table =
+  map (fn (pred, modes) =>
+    (pred, map (fn (mode, value) => value) modes)) preds_modes_table  
+    
+fun compile_preds compfuns mk_fun_of use_size thy all_vs param_vs preds moded_clauses =
+  map_preds_modes (fn pred => compile_pred compfuns mk_fun_of use_size thy all_vs param_vs pred
+      (the (AList.lookup (op =) preds pred))) moded_clauses  
+  
+fun prove thy clauses preds modes moded_clauses compiled_terms =
+  map_preds_modes (prove_pred thy clauses preds modes)
+    (join_preds_modes moded_clauses compiled_terms)
+
+fun prove_by_skip thy _ _ _ _ compiled_terms =
+  map_preds_modes (fn pred => fn mode => fn t => Drule.standard (setmp quick_and_dirty true (SkipProof.make_thm thy) t))
+    compiled_terms
+    
+fun prepare_intrs thy prednames =
+  let
+    val intrs = maps (intros_of thy) prednames
+      |> map (Logic.unvarify o prop_of)
+    val nparams = nparams_of thy (hd prednames)
+    val extra_modes = all_modes_of thy |> filter_out (fn (name, _) => member (op =) prednames name)
+    val preds = distinct (op =) (map (dest_Const o fst o (strip_intro_concl nparams)) intrs)
+    val _ $ u = Logic.strip_imp_concl (hd intrs);
+    val params = List.take (snd (strip_comb u), nparams);
+    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 (Ts, Us) = chop nparams (binder_types T)
+    in
+      (AList.update op = (name, these (AList.lookup op = clauses name) @
+        [(ts2, prems)]) clauses,
+       AList.update op = (name, (map (fn U => (case strip_type U of
+                 (Rs as _ :: _, Type ("bool", [])) => SOME (length Rs)
+               | _ => NONE)) Ts,
+             length Us)) arities)
+    end;
+    val (clauses, arities) = fold add_clause intrs ([], []);
+    fun modes_of_arities arities =
+      (map (fn (s, (ks, k)) => (s, cprod (cprods (map
+            (fn NONE => [NONE]
+              | SOME k' => map SOME (map (map (rpair NONE)) (subsets 1 k'))) ks),
+       map (map (rpair NONE)) (subsets 1 k)))) arities)
+    fun modes_of_typ T =
+      let
+        val (Ts, Us) = chop nparams (binder_types T)
+        fun all_smodes_of_typs Ts = cprods_subset (
+          map_index (fn (i, U) =>
+            case HOLogic.strip_tupleT U of
+              [] => [(i + 1, NONE)]
+            | [U] => [(i + 1, NONE)]
+            | Us =>  (i + 1, NONE) ::
+              (map (pair (i + 1) o SOME) ((subsets 1 (length Us)) \\ [[], 1 upto (length Us)])))
+          Ts)
+      in
+        cprod (cprods (map (fn T => case strip_type T of
+          (Rs as _ :: _, Type ("bool", [])) => map SOME (all_smodes_of_typs Rs) | _ => [NONE]) Ts),
+           all_smodes_of_typs Us)
+      end
+    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;
+
+(** main function of predicate compiler **)
+
+fun add_equations_of steps prednames thy =
+  let
+    val _ = Output.tracing ("Starting predicate compiler for predicates " ^ commas prednames ^ "...")
+    val _ = Output.tracing (commas (map (Display.string_of_thm_global thy) (maps (intros_of thy) prednames)))
+    val (preds, nparams, all_vs, param_vs, extra_modes, clauses, all_modes) =
+      prepare_intrs thy prednames
+    val _ = Output.tracing "Infering modes..."
+    val moded_clauses = #infer_modes steps thy extra_modes all_modes param_vs clauses 
+    val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses
+    val _ = print_modes modes
+    val _ = print_moded_clauses thy moded_clauses
+    val _ = Output.tracing "Defining executable functions..."
+    val thy' = fold (#create_definitions steps preds) modes thy
+      |> Theory.checkpoint
+    val _ = Output.tracing "Compiling equations..."
+    val compiled_terms =
+      (#compile_preds steps) thy' all_vs param_vs preds moded_clauses
+    val _ = print_compiled_terms thy' compiled_terms
+    val _ = Output.tracing "Proving equations..."
+    val result_thms = #prove steps thy' clauses preds (extra_modes @ modes)
+      moded_clauses compiled_terms
+    val qname = #qname steps
+    (* val attrib = gn thy => Attrib.attribute_i thy Code.add_eqn_attrib *)
+    val attrib = fn thy => Attrib.attribute_i thy (Attrib.internal (K (Thm.declaration_attribute
+      (fn thm => Context.mapping (Code.add_eqn thm) I))))
+    val thy'' = fold (fn (name, result_thms) => fn thy => snd (PureThy.add_thmss
+      [((Binding.qualify true (Long_Name.base_name name) (Binding.name qname), result_thms),
+        [attrib thy ])] thy))
+      (maps_modes result_thms) thy'
+      |> Theory.checkpoint
+  in
+    thy''
+  end
+
+fun extend' value_of edges_of key (G, visited) =
+  let
+    val (G', v) = case try (Graph.get_node G) key of
+        SOME v => (G, v)
+      | NONE => (Graph.new_node (key, value_of key) G, value_of key)
+    val (G'', visited') = fold (extend' value_of edges_of) (edges_of (key, v) \\ visited)
+      (G', key :: visited) 
+  in
+    (fold (Graph.add_edge o (pair key)) (edges_of (key, v)) G'', visited')
+  end;
+
+fun extend value_of edges_of key G = fst (extend' value_of edges_of key (G, [])) 
+  
+fun gen_add_equations steps names thy =
+  let
+    val thy' = PredData.map (fold (extend (fetch_pred_data thy) (depending_preds_of thy)) names) thy
+      |> Theory.checkpoint;
+    fun strong_conn_of gr keys =
+      Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr)
+    val scc = strong_conn_of (PredData.get thy') names
+    val thy'' = fold_rev
+      (fn preds => fn thy =>
+        if #are_not_defined steps thy preds then add_equations_of steps preds thy else thy)
+      scc thy' |> Theory.checkpoint
+  in thy'' end
+
+(* different instantiantions of the predicate compiler *)
+
+val add_equations = gen_add_equations
+  {infer_modes = infer_modes,
+  create_definitions = create_definitions,
+  compile_preds = compile_preds PredicateCompFuns.compfuns mk_fun_of false,
+  prove = prove,
+  are_not_defined = fn thy => forall (null o modes_of thy),
+  qname = "equation"}
+
+val add_sizelim_equations = gen_add_equations
+  {infer_modes = infer_modes,
+  create_definitions = sizelim_create_definitions,
+  compile_preds = compile_preds PredicateCompFuns.compfuns mk_sizelim_fun_of true,
+  prove = prove_by_skip,
+  are_not_defined = fn thy => forall (null o sizelim_modes_of thy),
+  qname = "sizelim_equation"
+  }
+
+val add_quickcheck_equations = gen_add_equations
+  {infer_modes = infer_modes_with_generator,
+  create_definitions = rpred_create_definitions,
+  compile_preds = compile_preds RPredCompFuns.compfuns mk_generator_of true,
+  prove = prove_by_skip,
+  are_not_defined = fn thy => forall (null o rpred_modes_of thy),
+  qname = "rpred_equation"}
+
+(** user interface **)
+
+(* code_pred_intro attribute *)
+
+fun attrib f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I);
+
+val code_pred_intros_attrib = attrib add_intro;
+
+
+(*FIXME
+- Naming of auxiliary rules necessary?
+- add default code equations P x y z = P_i_i_i x y z
+*)
+
+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 =
+  let
+    val thy = ProofContext.theory_of lthy
+    val const = prep_const thy raw_const
+    val lthy' = LocalTheory.theory (PredData.map
+        (extend (fetch_pred_data thy) (depending_preds_of thy) const)) lthy
+      |> LocalTheory.checkpoint
+    val thy' = ProofContext.theory_of lthy'
+    val preds = Graph.all_preds (PredData.get thy') [const] |> filter_out (has_elim thy')
+    fun mk_cases const =
+      let
+        val nparams = nparams_of thy' const
+        val intros = intros_of thy' const
+      in mk_casesrule lthy' nparams intros end  
+    val cases_rules = map mk_cases preds
+    val cases =
+      map (fn case_rule => RuleCases.Case {fixes = [],
+        assumes = [("", Logic.strip_imp_prems case_rule)],
+        binds = [], cases = []}) cases_rules
+    val case_env = map2 (fn p => fn c => (Long_Name.base_name p, SOME c)) preds cases
+    val lthy'' = lthy'
+      |> fold Variable.auto_fixes cases_rules 
+      |> ProofContext.add_cases true case_env
+    fun after_qed thms goal_ctxt =
+      let
+        val global_thms = ProofContext.export goal_ctxt
+          (ProofContext.init (ProofContext.theory_of goal_ctxt)) (map the_single thms)
+      in
+        goal_ctxt |> LocalTheory.theory (fold set_elim global_thms #>
+          (if rpred then
+            (add_equations [const] #>
+             add_sizelim_equations [const] #> add_quickcheck_equations [const])
+        else add_equations [const]))
+      end  
+  in
+    Proof.theorem_i NONE after_qed (map (single o (rpair [])) cases_rules) lthy''
+  end;
+
+val code_pred = generic_code_pred (K I);
+val code_pred_cmd = generic_code_pred Code.read_const
+
+(* transformation for code generation *)
+
+val eval_ref = ref (NONE : (unit -> term Predicate.pred) option);
+
+(*FIXME turn this into an LCF-guarded preprocessor for comprehensions*)
+fun analyze_compr thy 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);
+    val (body, Ts, fp) = HOLogic.strip_psplits split;
+    val (pred as Const (name, T), all_args) = strip_comb body;
+    val (params, args) = chop (nparams_of thy name) all_args;
+    val user_mode = map_filter I (map_index
+      (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 modes = filter (fn Mode (_, is, _) => is = user_mode')
+      (modes_of_term (all_modes_of thy) (list_comb (pred, params)));
+    val m = case modes
+     of [] => error ("No mode possible for comprehension "
+                ^ Syntax.string_of_term_global thy t_compr)
+      | [m] => m
+      | 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 t_eval = if null outargs then t_pred else
+      let
+        val outargs_bounds = map (fn Bound i => i) outargs;
+        val outargsTs = map (nth Ts) outargs_bounds;
+        val T_pred = HOLogic.mk_tupleT outargsTs;
+        val T_compr = HOLogic.mk_ptupleT fp Ts;
+        val arrange_bounds = map_index I outargs_bounds
+          |> sort (prod_ord (K EQUAL) int_ord)
+          |> map fst;
+        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 t_eval end;
+
+fun eval thy 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;
+
+fun values ctxt k t_compr =
+  let
+    val thy = ProofContext.theory_of ctxt;
+    val (T, t) = eval thy t_compr;
+    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
+  end;
+  (*
+fun random_values ctxt k t = 
+  let
+    val thy = ProofContext.theory_of ctxt
+    val _ = 
+  in
+  end;
+  *)
+fun values_cmd modes 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 ty' = Term.type_of t';
+    val ctxt' = Variable.auto_fixes t' ctxt;
+    val p = PrintMode.with_modes modes (fn () =>
+      Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk,
+        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 _ = 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)));
+
+end;
+
+end;
--- a/src/HOL/Tools/hologic.ML	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Tools/hologic.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -613,17 +613,17 @@
   | mk_typerep (T as TFree _) = Const ("Typerep.typerep_class.typerep",
       Term.itselfT T --> typerepT) $ Logic.mk_type T;
 
-val termT = Type ("Code_Eval.term", []);
+val termT = Type ("Code_Evaluation.term", []);
 
-fun term_of_const T = Const ("Code_Eval.term_of_class.term_of", T --> termT);
+fun term_of_const T = Const ("Code_Evaluation.term_of_class.term_of", T --> termT);
 
 fun mk_term_of T t = term_of_const T $ t;
 
 fun reflect_term (Const (c, T)) =
-      Const ("Code_Eval.Const", literalT --> typerepT --> termT)
+      Const ("Code_Evaluation.Const", literalT --> typerepT --> termT)
         $ mk_literal c $ mk_typerep T
   | reflect_term (t1 $ t2) =
-      Const ("Code_Eval.App", termT --> termT --> termT)
+      Const ("Code_Evaluation.App", termT --> termT --> termT)
         $ reflect_term t1 $ reflect_term t2
   | reflect_term (Abs (v, _, t)) = Abs (v, termT, reflect_term t)
   | reflect_term t = t;
@@ -631,7 +631,7 @@
 fun mk_valtermify_app c vs T =
   let
     fun termifyT T = mk_prodT (T, unitT --> termT);
-    fun valapp T T' = Const ("Code_Eval.valapp",
+    fun valapp T T' = Const ("Code_Evaluation.valapp",
       termifyT (T --> T') --> termifyT T --> termifyT T');
     fun mk_fTs [] _ = []
       | mk_fTs (_ :: Ts) T = (Ts ---> T) :: mk_fTs Ts T;
--- a/src/HOL/Tools/inductive.ML	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Tools/inductive.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -103,7 +103,10 @@
       "(P & True) = P" "(True & P) = P"
     by (fact simp_thms)+};
 
-val simp_thms'' = inf_fun_eq :: inf_bool_eq :: simp_thms';
+val simp_thms'' = map mk_meta_eq [@{thm inf_fun_eq}, @{thm inf_bool_eq}] @ simp_thms';
+
+val simp_thms''' = map mk_meta_eq
+  [@{thm le_fun_def}, @{thm le_bool_def}, @{thm sup_fun_eq}, @{thm sup_bool_eq}];
 
 
 (** context data **)
@@ -171,15 +174,15 @@
       (case concl of
           (_ $ (_ $ (Const ("Not", _) $ _) $ _)) => []
         | _ => [thm' RS (thm' RS eq_to_mono2)]);
-    fun dest_less_concl thm = dest_less_concl (thm RS le_funD)
-      handle THM _ => thm RS le_boolD
+    fun dest_less_concl thm = dest_less_concl (thm RS @{thm le_funD})
+      handle THM _ => thm RS @{thm le_boolD}
   in
     case concl of
       Const ("==", _) $ _ $ _ => eq2mono (thm RS meta_eq_to_obj_eq)
     | _ $ (Const ("op =", _) $ _ $ _) => eq2mono thm
     | _ $ (Const (@{const_name HOL.less_eq}, _) $ _ $ _) =>
       [dest_less_concl (Seq.hd (REPEAT (FIRSTGOAL
-         (resolve_tac [le_funI, le_boolI'])) thm))]
+         (resolve_tac [@{thm le_funI}, @{thm le_boolI'}])) thm))]
     | _ => [thm]
   end handle THM _ =>
     error ("Bad monotonicity theorem:\n" ^ Display.string_of_thm_without_context thm);
@@ -323,11 +326,11 @@
     (HOLogic.mk_Trueprop
       (Const (@{const_name Orderings.mono}, (predT --> predT) --> HOLogic.boolT) $ fp_fun))
     (fn _ => EVERY [rtac @{thm monoI} 1,
-      REPEAT (resolve_tac [le_funI, le_boolI'] 1),
+      REPEAT (resolve_tac [@{thm le_funI}, @{thm le_boolI'}] 1),
       REPEAT (FIRST
         [atac 1,
          resolve_tac (List.concat (map mk_mono monos) @ get_monos ctxt) 1,
-         etac le_funE 1, dtac le_boolD 1])]));
+         etac @{thm le_funE} 1, dtac @{thm le_boolD} 1])]));
 
 
 (* prove introduction rules *)
@@ -338,7 +341,7 @@
 
     val unfold = funpow k (fn th => th RS fun_cong)
       (mono RS (fp_def RS
-        (if coind then def_gfp_unfold else def_lfp_unfold)));
+        (if coind then @{thm def_gfp_unfold} else @{thm def_lfp_unfold})));
 
     fun select_disj 1 1 = []
       | select_disj _ 1 = [rtac disjI1]
@@ -553,13 +556,13 @@
     val ind_concl = HOLogic.mk_Trueprop
       (HOLogic.mk_binrel "HOL.ord_class.less_eq" (rec_const, ind_pred));
 
-    val raw_fp_induct = (mono RS (fp_def RS def_lfp_induct));
+    val raw_fp_induct = (mono RS (fp_def RS @{thm def_lfp_induct}));
 
     val induct = SkipProof.prove ctxt'' [] ind_prems ind_concl
       (fn {prems, ...} => EVERY
         [rewrite_goals_tac [inductive_conj_def],
          DETERM (rtac raw_fp_induct 1),
-         REPEAT (resolve_tac [le_funI, le_boolI] 1),
+         REPEAT (resolve_tac [@{thm le_funI}, @{thm le_boolI}] 1),
          rewrite_goals_tac simp_thms'',
          (*This disjE separates out the introduction rules*)
          REPEAT (FIRSTGOAL (eresolve_tac [disjE, exE, FalseE])),
@@ -577,7 +580,7 @@
         [rewrite_goals_tac rec_preds_defs,
          REPEAT (EVERY
            [REPEAT (resolve_tac [conjI, impI] 1),
-            REPEAT (eresolve_tac [le_funE, le_boolE] 1),
+            REPEAT (eresolve_tac [@{thm le_funE}, @{thm le_boolE}] 1),
             atac 1,
             rewrite_goals_tac simp_thms',
             atac 1])])
@@ -763,8 +766,8 @@
            (snd (Variable.add_fixes (map (fst o dest_Free) params) ctxt1)) ctxt1)
            (rotate_prems ~1 (ObjectLogic.rulify
              (fold_rule rec_preds_defs
-               (rewrite_rule [le_fun_def, le_bool_def, sup_fun_eq, sup_bool_eq]
-                (mono RS (fp_def RS def_coinduct))))))
+               (rewrite_rule simp_thms'''
+                (mono RS (fp_def RS @{thm def_coinduct}))))))
        else
          prove_indrule quiet_mode cs argTs bs xs rec_const params intr_ts mono fp_def
            rec_preds_defs ctxt1);
--- a/src/HOL/Tools/quickcheck_generators.ML	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Tools/quickcheck_generators.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -76,7 +76,7 @@
     val lhs = HOLogic.mk_random T size;
     val rhs = HOLogic.mk_ST [((HOLogic.mk_random T' size, @{typ Random.seed}), SOME (v, Tm'))]
       (HOLogic.mk_return Tm @{typ Random.seed}
-      (mk_const "Code_Eval.valapp" [T', T]
+      (mk_const "Code_Evaluation.valapp" [T', T]
         $ HOLogic.mk_prod (t_constr, Abs ("u", @{typ unit}, HOLogic.reflect_term t_constr)) $ t_v))
       @{typ Random.seed} (SOME Tm, @{typ Random.seed});
     val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
--- a/src/HOL/UNITY/Simple/Common.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/UNITY/Simple/Common.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/UNITY/Common
-    ID:         $Id$
     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1998  University of Cambridge
 
@@ -10,7 +9,9 @@
 From Misra, "A Logic for Concurrent Programming" (1994), sections 5.1 and 13.1.
 *)
 
-theory Common imports "../UNITY_Main" begin
+theory Common
+imports "../UNITY_Main"
+begin
 
 consts
   ftime :: "nat=>nat"
@@ -65,7 +66,7 @@
 lemma "mk_total_program (UNIV, {range(%t.(t, max (ftime t) (gtime t)))}, UNIV)
        \<in> {m} co (maxfg m)"
 apply (simp add: mk_total_program_def) 
-apply (simp add: constrains_def maxfg_def gasc)
+apply (simp add: constrains_def maxfg_def gasc min_max.sup_absorb2)
 done
 
 (*This one is  t := t+1 if t <max (ftime t) (gtime t) *)
@@ -73,7 +74,7 @@
           (UNIV, { {(t, Suc t) | t. t < max (ftime t) (gtime t)} }, UNIV)   
        \<in> {m} co (maxfg m)"
 apply (simp add: mk_total_program_def) 
-apply (simp add: constrains_def maxfg_def gasc)
+apply (simp add: constrains_def maxfg_def gasc min_max.sup_absorb2)
 done
 
 
--- a/src/HOL/Word/BinBoolList.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Word/BinBoolList.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -919,7 +919,7 @@
   apply (drule spec)
   apply (erule trans)
   apply (drule_tac x = "bin_cat y n a" in spec)
-  apply (simp add : bin_cat_assoc_sym)
+  apply (simp add : bin_cat_assoc_sym min_max.inf_absorb2)
   done
 
 lemma bin_rcat_bl:
--- a/src/HOL/Word/BinGeneral.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Word/BinGeneral.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -508,7 +508,7 @@
 lemma sbintrunc_sbintrunc_min [simp]:
   "sbintrunc m (sbintrunc n w) = sbintrunc (min m n) w"
   apply (rule bin_eqI)
-  apply (auto simp: nth_sbintr)
+  apply (auto simp: nth_sbintr min_max.inf_absorb1 min_max.inf_absorb2)
   done
 
 lemmas bintrunc_Pls = 
--- a/src/HOL/Word/WordDefinition.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Word/WordDefinition.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -381,14 +381,14 @@
   apply (unfold word_size)
   apply (subst word_ubin.norm_Rep [symmetric]) 
   apply (simp only: bintrunc_bintrunc_min word_size)
-  apply simp
+  apply (simp add: min_max.inf_absorb2)
   done
 
 lemma wi_bintr': 
   "wb = word_of_int bin ==> n >= size wb ==> 
     word_of_int (bintrunc n bin) = wb"
   unfolding word_size
-  by (clarsimp simp add : word_ubin.norm_eq_iff [symmetric])
+  by (clarsimp simp add: word_ubin.norm_eq_iff [symmetric] min_max.inf_absorb1)
 
 lemmas bintr_uint = bintr_uint' [unfolded word_size]
 lemmas wi_bintr = wi_bintr' [unfolded word_size]
--- a/src/HOL/Word/WordShift.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/Word/WordShift.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -1017,8 +1017,8 @@
   (word_cat (word_of_int w :: 'b word) (b :: 'c word) :: 'a word) = 
   word_of_int (bin_cat w (size b) (uint b))"
   apply (unfold word_cat_def word_size) 
-  apply (clarsimp simp add : word_ubin.norm_eq_iff [symmetric]
-      word_ubin.eq_norm bintr_cat)
+  apply (clarsimp simp add: word_ubin.norm_eq_iff [symmetric]
+      word_ubin.eq_norm bintr_cat min_max.inf_absorb1)
   done
 
 lemma word_cat_split_alt:
--- a/src/HOL/ex/Predicate_Compile.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/ex/Predicate_Compile.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -1,8 +1,17 @@
 theory Predicate_Compile
 imports Complex_Main RPred
-uses "predicate_compile.ML"
+uses
+  "../Tools/Predicate_Compile/pred_compile_aux.ML"
+  "../Tools/Predicate_Compile/predicate_compile_core.ML"
+  "../Tools/Predicate_Compile/pred_compile_set.ML"
+  "../Tools/Predicate_Compile/pred_compile_data.ML"
+  "../Tools/Predicate_Compile/pred_compile_fun.ML"
+  "../Tools/Predicate_Compile/pred_compile_pred.ML"
+  "../Tools/Predicate_Compile/predicate_compile.ML"
+  "../Tools/Predicate_Compile/pred_compile_quickcheck.ML"
 begin
 
 setup {* Predicate_Compile.setup *}
+setup {* Quickcheck.add_generator ("pred_compile", Pred_Compile_Quickcheck.quickcheck) *}
 
 end
\ No newline at end of file
--- a/src/HOL/ex/Predicate_Compile_ex.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/ex/Predicate_Compile_ex.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -22,8 +22,10 @@
   | "append xs ys zs \<Longrightarrow> append (x # xs) ys (x # zs)"
 
 code_pred append .
+code_pred (inductify_all) (rpred) append .
 
 thm append.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}"
@@ -49,6 +51,22 @@
 
 thm 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"
@@ -70,15 +88,11 @@
   case tranclp
   from this converse_tranclpE[OF this(1)] show thesis by metis
 qed
-
+(*
+code_pred (inductify_all) (rpred) tranclp .
 thm tranclp.equation
-(*
-setup {* Predicate_Compile.add_sizelim_equations [@{const_name tranclp}] *}
-setup {* fn thy => exception_trace (fn () => Predicate_Compile.add_quickcheck_equations [@{const_name tranclp}] thy)  *}
-
 thm tranclp.rpred_equation
 *)
-
 inductive succ :: "nat \<Rightarrow> nat \<Rightarrow> bool" where
     "succ 0 1"
   | "succ m n \<Longrightarrow> succ (Suc m) (Suc n)"
@@ -157,6 +171,7 @@
 values 3 "{(a,q). step (par nil nil) a q}"
 *)
 
+subsection {* divmod *}
 
 inductive divmod_rel :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
     "k < l \<Longrightarrow> divmod_rel k l 0 k"
@@ -166,4 +181,261 @@
 
 value [code] "Predicate.singleton (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 .
+
+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"
+
+(*
+code_pred (inductify_all) (rpred) filterP .
+thm filterP.rpred_equation
+*)
+
+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 .
+thm lexn.equation
+
+code_pred (inductify_all) lenlex .
+thm lenlex.equation
+(*
+code_pred (inductify_all) (rpred) lenlex .
+thm lenlex.rpred_equation
+*)
+thm lists.intros
+code_pred (inductify_all) lists .
+
+thm lists.equation
+
+datatype 'a tree = ET | MKT 'a "'a tree" "'a tree" nat
+fun height :: "'a tree => nat" where
+"height ET = 0"
+| "height (MKT x l r h) = max (height l) (height r) + 1"
+
+consts avl :: "'a tree => bool"
+primrec
+  "avl ET = True"
+  "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 .
+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
+
+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
+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 .
+thm set_of.equation
+(* FIXME *)
+(*
+code_pred (inductify_all) is_ord .
+thm is_ord.equation
+*)
+section {* Definitions about Relations *}
+
+code_pred (inductify_all) converse .
+thm converse.equation
+
+code_pred (inductify_all) Domain .
+thm Domain.equation
+
+
+section {* Context Free Grammar *}
+
+datatype alphabet = a | b
+
+inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where
+  "[] \<in> S\<^isub>1"
+| "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
+| "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1"
+| "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1"
+| "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 .
+
+thm S\<^isub>1p.equation
+
+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]
+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"
+| "w \<in> B\<^isub>2 \<Longrightarrow> a # w \<in> S\<^isub>2"
+| "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"} *}
+*)
+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]
+oops
+
+inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
+  "[] \<in> S\<^isub>3"
+| "w \<in> A\<^isub>3 \<Longrightarrow> b # w \<in> S\<^isub>3"
+| "w \<in> B\<^isub>3 \<Longrightarrow> a # w \<in> S\<^isub>3"
+| "w \<in> S\<^isub>3 \<Longrightarrow> a # w \<in> A\<^isub>3"
+| "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_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]
+oops
+
+theorem S\<^isub>3_complete:
+"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<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"
+| "w \<in> A\<^isub>4 \<Longrightarrow> b # w \<in> S\<^isub>4"
+| "w \<in> B\<^isub>4 \<Longrightarrow> a # w \<in> S\<^isub>4"
+| "w \<in> S\<^isub>4 \<Longrightarrow> a # w \<in> A\<^isub>4"
+| "\<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]
+oops
+
+theorem S\<^isub>4_A\<^isub>4_B\<^isub>4_sound_and_complete:
+"w \<in> S\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+"w \<in> A\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] + 1"
+"w \<in> B\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = b] = length [x \<leftarrow> w. x = a] + 1"
+(*quickcheck[generator = pred_compile, size=5, iterations=1]*)
+oops
+
+
+section {* Lambda *}
+datatype type =
+    Atom nat
+  | Fun type type    (infixr "\<Rightarrow>" 200)
+
+datatype dB =
+    Var nat
+  | App dB dB (infixl "\<degree>" 200)
+  | Abs type dB
+
+primrec
+  nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>" [90, 0] 91)
+where
+  "[]\<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"
+  | 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"
+
+primrec
+  lift :: "[dB, nat] => dB"
+where
+    "lift (Var i) k = (if i < k then Var i else Var (i + 1))"
+  | "lift (s \<degree> t) k = lift s k \<degree> lift t k"
+  | "lift (Abs T s) k = Abs T (lift s (k + 1))"
+
+primrec
+  subst :: "[dB, dB, nat] => dB"  ("_[_'/_]" [300, 0, 0] 300)
+where
+    subst_Var: "(Var i)[s/k] =
+      (if k < i then Var (i - 1) else if i = k then s else Var i)"
+  | subst_App: "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"
+  | subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])"
+
+inductive beta :: "[dB, dB] => bool"  (infixl "\<rightarrow>\<^sub>\<beta>" 50)
+  where
+    beta [simp, intro!]: "Abs T s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
+  | appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
+  | appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
+  | 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]
+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>
+  \<exists>us vs. ys = ts @ x # vs \<and> (\<forall>u\<in>set us. \<not> P u) \<and> P x \<and> xs = filter P vs"
+(*quickcheck[generator = pred_compile]*)
+oops
+
+
 end
\ No newline at end of file
--- a/src/HOL/ex/RPred.thy	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/HOL/ex/RPred.thy	Thu Sep 24 17:26:05 2009 +0200
@@ -14,13 +14,15 @@
 definition return :: "'a => 'a rpred"
   where "return x = Pair (Predicate.single x)"
 
-definition bind :: "'a rpred \<Rightarrow> ('a \<Rightarrow> 'b rpred) \<Rightarrow> 'b rpred" (infixl "\<guillemotright>=" 60)
+definition bind :: "'a rpred \<Rightarrow> ('a \<Rightarrow> 'b rpred) \<Rightarrow> 'b rpred"
+(* (infixl "\<guillemotright>=" 60) *)
   where "bind RP f =
   (\<lambda>s. let (P, s') = RP s;
         (s1, s2) = Random.split_seed s'
     in (Predicate.bind P (%a. fst (f a s1)), s2))"
 
-definition supp :: "'a rpred \<Rightarrow> 'a rpred \<Rightarrow> 'a rpred" (infixl "\<squnion>" 80)
+definition supp :: "'a rpred \<Rightarrow> 'a rpred \<Rightarrow> 'a rpred"
+(* (infixl "\<squnion>" 80) *)
 where
   "supp RP1 RP2 = (\<lambda>s. let (P1, s') = RP1 s; (P2, s'') = RP2 s'
   in (upper_semilattice_class.sup P1 P2, s''))"
@@ -43,6 +45,8 @@
   where "lift_random g = scomp g (Pair o (Predicate.single o fst))"
 
 definition map_rpred :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a rpred \<Rightarrow> 'b rpred)"
-where "map_rpred f P = P \<guillemotright>= (return o f)"
+where "map_rpred f P = bind P (return o f)"
+
+hide (open) const return bind supp 
   
 end
\ No newline at end of file
--- a/src/HOL/ex/predicate_compile.ML	Thu Sep 24 17:25:42 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2182 +0,0 @@
-(* Author: Lukas Bulwahn, TU Muenchen
-
-(Prototype of) A compiler from predicates specified by intro/elim rules
-to equations.
-*)
-
-signature PREDICATE_COMPILE =
-sig
-  type mode = int list option list * int list
-  (*val add_equations_of: bool -> string list -> theory -> theory *)
-  val register_predicate : (thm list * thm * int) -> theory -> theory
-  val is_registered : theory -> string -> bool
- (* val fetch_pred_data : theory -> string -> (thm list * thm * int)  *)
-  val predfun_intro_of: theory -> string -> mode -> thm
-  val predfun_elim_of: theory -> string -> mode -> thm
-  val strip_intro_concl: int -> term -> term * (term list * term list)
-  val predfun_name_of: theory -> string -> mode -> string
-  val all_preds_of : theory -> string list
-  val modes_of: theory -> string -> mode list
-  val string_of_mode : mode -> string
-  val intros_of: theory -> string -> thm list
-  val nparams_of: theory -> string -> int
-  val add_intro: thm -> theory -> theory
-  val set_elim: thm -> theory -> theory
-  val setup: theory -> theory
-  val code_pred: string -> Proof.context -> Proof.state
-  val code_pred_cmd: string -> Proof.context -> Proof.state
-  val print_stored_rules: theory -> unit
-  val print_all_modes: theory -> unit
-  val do_proofs: bool ref
-  val mk_casesrule : Proof.context -> int -> thm list -> term
-  val analyze_compr: theory -> term -> term
-  val eval_ref: (unit -> term Predicate.pred) option ref
-  val add_equations : string list -> theory -> theory
-  val code_pred_intros_attrib : attribute
-  (* used by Quickcheck_Generator *) 
-  (*val funT_of : mode -> typ -> typ
-  val mk_if_pred : term -> term
-  val mk_Eval : term * term -> term*)
-  val mk_tupleT : typ list -> typ
-(*  val mk_predT :  typ -> typ *)
-  (* temporary for testing of the compilation *)
-  datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term |
-    GeneratorPrem of term list * term | Generator of (string * typ);
-  val prepare_intrs: theory -> string list ->
-    (string * typ) list * int * string list * string list * (string * mode list) list *
-    (string * (term list * indprem list) list) list * (string * (int option list * int)) list
-  datatype compilation_funs = CompilationFuns of {
-    mk_predT : typ -> typ,
-    dest_predT : typ -> typ,
-    mk_bot : typ -> term,
-    mk_single : term -> term,
-    mk_bind : term * term -> term,
-    mk_sup : term * term -> term,
-    mk_if : term -> term,
-    mk_not : term -> term,
-    mk_map : typ -> typ -> term -> term -> term,
-    lift_pred : term -> term
-  };  
-  datatype tmode = Mode of mode * int list * tmode option list;
-  type moded_clause = term list * (indprem * tmode) list
-  type 'a pred_mode_table = (string * (mode * 'a) list) list
-  val infer_modes : bool -> theory -> (string * (int list option list * int list) list) list
-    -> (string * (int option list * int)) list -> string list
-    -> (string * (term list * indprem list) list) list
-    -> (moded_clause list) pred_mode_table
-  val infer_modes_with_generator : theory -> (string * (int list option list * int list) list) list
-    -> (string * (int option list * int)) list -> string list
-    -> (string * (term list * indprem list) list) list
-    -> (moded_clause list) pred_mode_table  
-  (*val compile_preds : theory -> compilation_funs -> string list -> string list
-    -> (string * typ) list -> (moded_clause list) pred_mode_table -> term pred_mode_table
-  val rpred_create_definitions :(string * typ) list -> string * mode list
-    -> theory -> theory 
-  val split_smode : int list -> term list -> (term list * term list) *)
-  val print_moded_clauses :
-    theory -> (moded_clause list) pred_mode_table -> unit
-  val print_compiled_terms : theory -> term pred_mode_table -> unit
-  (*val rpred_prove_preds : theory -> term pred_mode_table -> thm pred_mode_table*)
-  val rpred_compfuns : compilation_funs
-  val dest_funT : typ -> typ * typ
- (* val depending_preds_of : theory -> thm list -> string list *)
-  val add_quickcheck_equations : string list -> theory -> theory
-  val add_sizelim_equations : string list -> theory -> theory
-  val is_inductive_predicate : theory -> string -> bool
-  val terms_vs : term list -> string list
-  val subsets : int -> int -> int list list
-  val check_mode_clause : bool -> theory -> string list ->
-    (string * mode list) list -> (string * mode list) list -> mode -> (term list * indprem list)
-      -> (term list * (indprem * tmode) list) option
-  val string_of_moded_prem : theory -> (indprem * tmode) -> string
-  val all_modes_of : theory -> (string * mode list) list
-  val all_generator_modes_of : theory -> (string * mode list) list
-  val compile_clause : compilation_funs -> term option -> (term list -> term) ->
-    theory -> string list -> string list -> mode -> term -> moded_clause -> term
-  val preprocess_intro : theory -> thm -> thm
-  val is_constrt : theory -> term -> bool
-  val is_predT : typ -> bool
-  val guess_nparams : typ -> int
-end;
-
-structure Predicate_Compile : PREDICATE_COMPILE =
-struct
-
-(** auxiliary **)
-
-(* debug stuff *)
-
-fun tracing s = (if ! Toplevel.debug then Output.tracing s else ());
-
-fun print_tac s = Seq.single; (* (if ! Toplevel.debug then Tactical.print_tac s else Seq.single); *)
-fun debug_tac msg = Seq.single; (* (fn st => (Output.tracing msg; Seq.single st)); *)
-
-val do_proofs = ref true;
-
-fun mycheat_tac thy i st =
-  (Tactic.rtac (SkipProof.make_thm thy (Var (("A", 0), propT))) i) st
-
-fun remove_last_goal thy st =
-  (Tactic.rtac (SkipProof.make_thm thy (Var (("A", 0), propT))) (nprems_of st)) st
-
-(* reference to preprocessing of InductiveSet package *)
-
-val ind_set_codegen_preproc = Inductive_Set.codegen_preproc;
-
-(** fundamentals **)
-
-(* syntactic operations *)
-
-fun mk_eq (x, xs) =
-  let fun mk_eqs _ [] = []
-        | mk_eqs a (b::cs) =
-            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
-    val U = fastype_of u
-    val [A] = binder_types T
-    val D = body_type U 
-  in 
-    Const (@{const_name "scomp"}, T --> U --> A --> D) $ t $ u
-  end;
-
-fun dest_funT (Type ("fun",[S, T])) = (S, T)
-  | dest_funT T = raise TYPE ("dest_funT", [T], [])
- 
-fun mk_fun_comp (t, u) =
-  let
-    val (_, B) = dest_funT (fastype_of t)
-    val (C, A) = dest_funT (fastype_of u)
-  in
-    Const(@{const_name "Fun.comp"}, (A --> B) --> (C --> A) --> C --> B) $ t $ u
-  end;
-
-fun dest_randomT (Type ("fun", [@{typ Random.seed},
-  Type ("*", [Type ("*", [T, @{typ "unit => Code_Eval.term"}]) ,@{typ Random.seed}])])) = T
-  | dest_randomT T = raise TYPE ("dest_randomT", [T], [])
-
-(* destruction of intro rules *)
-
-(* FIXME: look for other place where this functionality was used before *)
-fun strip_intro_concl nparams intro = let
-  val _ $ u = Logic.strip_imp_concl intro
-  val (pred, all_args) = strip_comb u
-  val (params, args) = chop nparams all_args
-in (pred, (params, args)) end
-
-(** data structures **)
-
-type smode = int list;
-type mode = smode option list * smode;
-datatype tmode = Mode of mode * int list * tmode option list;
-
-fun split_smode is ts =
-  let
-    fun split_smode' _ _ [] = ([], [])
-      | split_smode' is i (t::ts) = (if i mem is then apfst else apsnd) (cons t)
-          (split_smode' is (i+1) ts)
-  in split_smode' is 1 ts end
-
-fun split_mode (iss, is) ts =
-  let
-    val (t1, t2) = chop (length iss) ts 
-  in (t1, split_smode is t2) end
-
-fun string_of_mode (iss, is) = space_implode " -> " (map
-  (fn NONE => "X"
-    | SOME js => enclose "[" "]" (commas (map string_of_int js)))
-       (iss @ [SOME is]));
-
-fun string_of_tmode (Mode (predmode, termmode, param_modes)) =
-  "predmode: " ^ (string_of_mode predmode) ^ 
-  (if null param_modes then "" else
-    "; " ^ "params: " ^ commas (map (the_default "NONE" o Option.map string_of_tmode) param_modes))
-    
-datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term |
-  GeneratorPrem of term list * term | Generator of (string * typ);
-
-type moded_clause = term list * (indprem * tmode) list
-type 'a pred_mode_table = (string * (mode * 'a) list) list
-
-datatype predfun_data = PredfunData of {
-  name : string,
-  definition : thm,
-  intro : thm,
-  elim : thm
-};
-
-fun rep_predfun_data (PredfunData data) = data;
-fun mk_predfun_data (name, definition, intro, elim) =
-  PredfunData {name = name, definition = definition, intro = intro, elim = elim}
-
-datatype function_data = FunctionData of {
-  name : string,
-  equation : thm option (* is not used at all? *)
-};
-
-fun rep_function_data (FunctionData data) = data;
-fun mk_function_data (name, equation) =
-  FunctionData {name = name, equation = equation}
-
-datatype pred_data = PredData of {
-  intros : thm list,
-  elim : thm option,
-  nparams : int,
-  functions : (mode * predfun_data) list,
-  generators : (mode * function_data) list,
-  sizelim_functions : (mode * function_data) list 
-};
-
-fun rep_pred_data (PredData data) = data;
-fun mk_pred_data ((intros, elim, nparams), (functions, generators, sizelim_functions)) =
-  PredData {intros = intros, elim = elim, nparams = nparams,
-    functions = functions, generators = generators, sizelim_functions = sizelim_functions}
-fun map_pred_data f (PredData {intros, elim, nparams, functions, generators, sizelim_functions}) =
-  mk_pred_data (f ((intros, elim, nparams), (functions, generators, sizelim_functions)))
-  
-fun eq_option eq (NONE, NONE) = true
-  | eq_option eq (SOME x, SOME y) = eq (x, y)
-  | 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;
-  val empty = Graph.empty;
-  val copy = I;
-  val extend = I;
-  fun merge _ = Graph.merge eq_pred_data;
-);
-
-(* queries *)
-
-fun lookup_pred_data thy name =
-  Option.map rep_pred_data (try (Graph.get_node (PredData.get thy)) name)
-
-fun the_pred_data thy name = case lookup_pred_data thy name
- of NONE => error ("No such predicate " ^ quote name)  
-  | SOME data => data;
-
-val is_registered = is_some oo lookup_pred_data 
-
-val all_preds_of = Graph.keys o PredData.get
-
-val intros_of = #intros oo the_pred_data
-
-fun the_elim_of thy name = case #elim (the_pred_data thy name)
- of NONE => error ("No elimination rule for predicate " ^ quote name)
-  | SOME thm => thm 
-  
-val has_elim = is_some o #elim oo the_pred_data;
-
-val nparams_of = #nparams oo the_pred_data
-
-val modes_of = (map fst) o #functions oo the_pred_data
-
-fun all_modes_of thy = map (fn name => (name, modes_of thy name)) (all_preds_of thy) 
-
-val is_compiled = not o null o #functions oo the_pred_data
-
-fun lookup_predfun_data thy name mode =
-  Option.map rep_predfun_data (AList.lookup (op =)
-  (#functions (the_pred_data thy name)) mode)
-
-fun the_predfun_data thy name mode = case lookup_predfun_data thy name mode
-  of NONE => error ("No function defined for mode " ^ string_of_mode mode ^ " of predicate " ^ name)
-   | SOME data => data;
-
-val predfun_name_of = #name ooo the_predfun_data
-
-val predfun_definition_of = #definition ooo the_predfun_data
-
-val predfun_intro_of = #intro ooo the_predfun_data
-
-val predfun_elim_of = #elim ooo the_predfun_data
-
-fun lookup_generator_data thy name mode = 
-  Option.map rep_function_data (AList.lookup (op =)
-  (#generators (the_pred_data thy name)) mode)
-  
-fun the_generator_data thy name mode = case lookup_generator_data thy name mode
-  of NONE => error ("No generator defined for mode " ^ string_of_mode mode ^ " of predicate " ^ name)
-   | SOME data => data
-
-val generator_name_of = #name ooo the_generator_data
-
-val generator_modes_of = (map fst) o #generators oo the_pred_data
-
-fun all_generator_modes_of thy =
-  map (fn name => (name, generator_modes_of thy name)) (all_preds_of thy) 
-
-fun lookup_sizelim_function_data thy name mode =
-  Option.map rep_function_data (AList.lookup (op =)
-  (#sizelim_functions (the_pred_data thy name)) mode)
-
-fun the_sizelim_function_data thy name mode = case lookup_sizelim_function_data thy name mode
-  of NONE => error ("No size-limited function defined for mode " ^ string_of_mode mode
-    ^ " of predicate " ^ name)
-   | SOME data => data
-
-val sizelim_function_name_of = #name ooo the_sizelim_function_data
-
-(*val generator_modes_of = (map fst) o #generators oo the_pred_data*)
-     
-(* diagnostic display functions *)
-
-fun print_modes modes = Output.tracing ("Inferred modes:\n" ^
-  cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
-    string_of_mode ms)) modes));
-
-fun print_pred_mode_table string_of_entry thy pred_mode_table =
-  let
-    fun print_mode pred (mode, entry) =  "mode : " ^ (string_of_mode mode)
-      ^ (string_of_entry pred mode entry)  
-    fun print_pred (pred, modes) =
-      "predicate " ^ pred ^ ": " ^ cat_lines (map (print_mode pred) modes)
-    val _ = Output.tracing (cat_lines (map print_pred pred_mode_table))
-  in () end;
-
-fun string_of_moded_prem thy (Prem (ts, p), tmode) =
-    (Syntax.string_of_term_global thy (list_comb (p, ts))) ^
-    "(" ^ (string_of_tmode tmode) ^ ")"
-  | string_of_moded_prem thy (GeneratorPrem (ts, p), Mode (predmode, is, _)) =
-    (Syntax.string_of_term_global thy (list_comb (p, ts))) ^
-    "(generator_mode: " ^ (string_of_mode predmode) ^ ")"
-  | string_of_moded_prem thy (Generator (v, T), _) =
-    "Generator for " ^ v ^ " of Type " ^ (Syntax.string_of_typ_global thy T)
-  | string_of_moded_prem thy (Negprem (ts, p), Mode (_, is, _)) =
-    (Syntax.string_of_term_global thy (list_comb (p, ts))) ^
-    "(negative mode: " ^ (space_implode ", " (map string_of_int is)) ^ ")"
-  | string_of_moded_prem thy (Sidecond t, Mode (_, is, _)) =
-    (Syntax.string_of_term_global thy t) ^
-    "(sidecond mode: " ^ (space_implode ", " (map string_of_int is)) ^ ")"    
-  | string_of_moded_prem _ _ = error "string_of_moded_prem: unimplemented"
-     
-fun print_moded_clauses thy =
-  let        
-    fun string_of_clause pred mode clauses =
-      cat_lines (map (fn (ts, prems) => (space_implode " --> "
-        (map (string_of_moded_prem thy) prems)) ^ " --> " ^ pred ^ " "
-        ^ (space_implode " " (map (Syntax.string_of_term_global thy) ts))) clauses)
-  in print_pred_mode_table string_of_clause thy end;
-
-fun print_compiled_terms thy =
-  print_pred_mode_table (fn _ => fn _ => Syntax.string_of_term_global thy) thy
-    
-fun print_stored_rules thy =
-  let
-    val preds = (Graph.keys o PredData.get) thy
-    fun print pred () = let
-      val _ = writeln ("predicate: " ^ pred)
-      val _ = writeln ("number of parameters: " ^ string_of_int (nparams_of thy pred))
-      val _ = writeln ("introrules: ")
-      val _ = fold (fn thm => fn u => writeln (Display.string_of_thm_global thy thm))
-        (rev (intros_of thy pred)) ()
-    in
-      if (has_elim thy pred) then
-        writeln ("elimrule: " ^ Display.string_of_thm_global thy (the_elim_of thy pred))
-      else
-        writeln ("no elimrule defined")
-    end
-  in
-    fold print preds ()
-  end;
-
-fun print_all_modes thy =
-  let
-    val _ = writeln ("Inferred modes:")
-    fun print (pred, modes) u =
-      let
-        val _ = writeln ("predicate: " ^ pred)
-        val _ = writeln ("modes: " ^ (commas (map string_of_mode modes)))
-      in u end  
-  in
-    fold print (all_modes_of thy) ()
-  end
-  
-(** preprocessing rules **)  
-
-fun imp_prems_conv cv ct =
-  case Thm.term_of ct of
-    Const ("==>", _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv) (imp_prems_conv cv) ct
-  | _ => Conv.all_conv ct
-
-fun Trueprop_conv cv ct =
-  case Thm.term_of ct of
-    Const ("Trueprop", _) $ _ => Conv.arg_conv cv ct  
-  | _ => error "Trueprop_conv"
-
-fun preprocess_intro thy rule =
-  Conv.fconv_rule
-    (imp_prems_conv
-      (Trueprop_conv (Conv.try_conv (Conv.rewr_conv (Thm.symmetric @{thm Predicate.eq_is_eq})))))
-    (Thm.transfer thy rule)
-
-fun preprocess_elim thy nparams elimrule =
-  let
-    fun replace_eqs (Const ("Trueprop", _) $ (Const ("op =", T) $ lhs $ rhs)) =
-       HOLogic.mk_Trueprop (Const (@{const_name Predicate.eq}, T) $ lhs $ rhs)
-     | replace_eqs t = t
-    val prems = Thm.prems_of elimrule
-    val nargs = length (snd (strip_comb (HOLogic.dest_Trueprop (hd prems)))) - nparams
-    fun preprocess_case t =
-     let
-       val params = Logic.strip_params t
-       val (assums1, assums2) = chop nargs (Logic.strip_assums_hyp t)
-       val assums_hyp' = assums1 @ (map replace_eqs assums2)
-     in
-       list_all (params, Logic.list_implies (assums_hyp', Logic.strip_assums_concl t))
-     end 
-    val cases' = map preprocess_case (tl prems)
-    val elimrule' = Logic.list_implies ((hd prems) :: cases', Thm.concl_of elimrule)
-  in
-    Thm.equal_elim
-      (Thm.symmetric (Conv.implies_concl_conv (MetaSimplifier.rewrite true [@{thm eq_is_eq}])
-         (cterm_of thy elimrule')))
-      elimrule
-  end;
-
-(* special case: predicate with no introduction rule *)
-fun noclause thy predname elim = let
-  val T = (Logic.unvarifyT o Sign.the_const_type thy) predname
-  val Ts = binder_types T
-  val names = Name.variant_list []
-        (map (fn i => "x" ^ (string_of_int i)) (1 upto (length Ts)))
-  val vs = map2 (curry Free) names Ts
-  val clausehd = HOLogic.mk_Trueprop (list_comb (Const (predname, T), vs))
-  val intro_t = Logic.mk_implies (@{prop False}, clausehd)
-  val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT))
-  val elim_t = Logic.list_implies ([clausehd, Logic.mk_implies (@{prop False}, P)], P)
-  val intro = Goal.prove (ProofContext.init thy) names [] intro_t
-        (fn {...} => etac @{thm FalseE} 1)
-  val elim = Goal.prove (ProofContext.init thy) ("P" :: names) [] elim_t
-        (fn {...} => etac elim 1) 
-in
-  ([intro], elim)
-end
-
-fun fetch_pred_data thy name =
-  case try (Inductive.the_inductive (ProofContext.init thy)) name of
-    SOME (info as (_, result)) => 
-      let
-        fun is_intro_of intro =
-          let
-            val (const, _) = strip_comb (HOLogic.dest_Trueprop (concl_of intro))
-          in (fst (dest_Const const) = name) end;      
-        val intros = ind_set_codegen_preproc thy ((map (preprocess_intro thy))
-          (filter is_intro_of (#intrs result)))
-        val pre_elim = nth (#elims result) (find_index (fn s => s = name) (#names (fst info)))
-        val nparams = length (Inductive.params_of (#raw_induct result))
-        val elim = singleton (ind_set_codegen_preproc thy) (preprocess_elim thy nparams pre_elim)
-        val (intros, elim) = if null intros then noclause thy name elim else (intros, elim)
-      in
-        mk_pred_data ((intros, SOME elim, nparams), ([], [], []))
-      end                                                                    
-  | NONE => error ("No such predicate: " ^ quote name)
-  
-(* updaters *)
-
-fun apfst3 f (x, y, z) =  (f x, y, z)
-fun apsnd3 f (x, y, z) =  (x, f y, z)
-fun aptrd3 f (x, y, z) =  (x, y, f z)
-
-fun add_predfun name mode data =
-  let
-    val add = (apsnd o apfst3 o cons) (mode, mk_predfun_data data)
-  in PredData.map (Graph.map_node name (map_pred_data add)) end
-
-fun is_inductive_predicate thy name =
-  is_some (try (Inductive.the_inductive (ProofContext.init thy)) name)
-
-fun depending_preds_of thy (key, value) =
-  let
-    val intros = (#intros o rep_pred_data) value
-  in
-    fold Term.add_const_names (map Thm.prop_of intros) []
-      |> filter (fn c => (not (c = key)) andalso (is_inductive_predicate thy c orelse is_registered thy c))
-  end;
-    
-    
-(* code dependency graph *)    
-(*
-fun dependencies_of thy name =
-  let
-    val (intros, elim, nparams) = fetch_pred_data thy name 
-    val data = mk_pred_data ((intros, SOME elim, nparams), ([], [], []))
-    val keys = depending_preds_of thy intros
-  in
-    (data, keys)
-  end;
-*)
-(* guessing number of parameters *)
-fun find_indexes pred xs =
-  let
-    fun find is n [] = is
-      | find is n (x :: xs) = find (if pred x then (n :: is) else is) (n + 1) xs;
-  in rev (find [] 0 xs) end;
-
-fun is_predT (T as Type("fun", [_, _])) = (snd (strip_type T) = HOLogic.boolT)
-  | is_predT _ = false
-  
-fun guess_nparams T =
-  let
-    val argTs = binder_types T
-    val nparams = fold (curry Int.max)
-      (map (fn x => x + 1) (find_indexes is_predT argTs)) 0
-  in nparams end;
-
-fun add_intro thm thy = let
-   val (name, T) = dest_Const (fst (strip_intro_concl 0 (prop_of thm)))
-   fun cons_intro gr =
-     case try (Graph.get_node gr) name of
-       SOME pred_data => Graph.map_node name (map_pred_data
-         (apfst (fn (intro, elim, nparams) => (thm::intro, elim, nparams)))) gr
-     | NONE =>
-       let
-         val nparams = the_default (guess_nparams T)  (try (#nparams o rep_pred_data o (fetch_pred_data thy)) name)
-       in Graph.new_node (name, mk_pred_data (([thm], NONE, nparams), ([], [], []))) gr end;
-  in PredData.map cons_intro thy end
-
-fun set_elim thm = let
-    val (name, _) = dest_Const (fst 
-      (strip_comb (HOLogic.dest_Trueprop (hd (prems_of thm)))))
-    fun set (intros, _, nparams) = (intros, SOME thm, nparams)  
-  in PredData.map (Graph.map_node name (map_pred_data (apfst set))) end
-
-fun set_nparams name nparams = let
-    fun set (intros, elim, _ ) = (intros, elim, nparams) 
-  in PredData.map (Graph.map_node name (map_pred_data (apfst set))) end
-    
-fun register_predicate (pre_intros, pre_elim, nparams) thy = let
-    val (name, _) = dest_Const (fst (strip_intro_concl nparams (prop_of (hd pre_intros))))
-    (* preprocessing *)
-    val intros = ind_set_codegen_preproc thy (map (preprocess_intro thy) pre_intros)
-    val elim = singleton (ind_set_codegen_preproc thy) (preprocess_elim thy nparams pre_elim)
-  in
-    PredData.map
-      (Graph.new_node (name, mk_pred_data ((intros, SOME elim, nparams), ([], [], [])))) thy
-  end
-
-fun set_generator_name pred mode name = 
-  let
-    val set = (apsnd o apsnd3 o cons) (mode, mk_function_data (name, NONE))
-  in
-    PredData.map (Graph.map_node pred (map_pred_data set))
-  end
-
-fun set_sizelim_function_name pred mode name = 
-  let
-    val set = (apsnd o aptrd3 o cons) (mode, mk_function_data (name, NONE))
-  in
-    PredData.map (Graph.map_node pred (map_pred_data set))
-  end
-
-(** data structures for generic compilation for different monads **)
-
-(* maybe rename functions more generic:
-  mk_predT -> mk_monadT; dest_predT -> dest_monadT
-  mk_single -> mk_return (?)
-*)
-datatype compilation_funs = CompilationFuns of {
-  mk_predT : typ -> typ,
-  dest_predT : typ -> typ,
-  mk_bot : typ -> term,
-  mk_single : term -> term,
-  mk_bind : term * term -> term,
-  mk_sup : term * term -> term,
-  mk_if : term -> term,
-  mk_not : term -> term,
-(*  funT_of : mode -> typ -> typ, *)
-(*  mk_fun_of : theory -> (string * typ) -> mode -> term, *) 
-  mk_map : typ -> typ -> term -> term -> term,
-  lift_pred : term -> term
-};
-
-fun mk_predT (CompilationFuns funs) = #mk_predT funs
-fun dest_predT (CompilationFuns funs) = #dest_predT funs
-fun mk_bot (CompilationFuns funs) = #mk_bot funs
-fun mk_single (CompilationFuns funs) = #mk_single funs
-fun mk_bind (CompilationFuns funs) = #mk_bind funs
-fun mk_sup (CompilationFuns funs) = #mk_sup funs
-fun mk_if (CompilationFuns funs) = #mk_if funs
-fun mk_not (CompilationFuns funs) = #mk_not funs
-(*fun funT_of (CompilationFuns funs) = #funT_of funs*)
-(*fun mk_fun_of (CompilationFuns funs) = #mk_fun_of funs*)
-fun mk_map (CompilationFuns funs) = #mk_map funs
-fun lift_pred (CompilationFuns funs) = #lift_pred funs
-
-fun funT_of compfuns (iss, is) T =
-  let
-    val Ts = binder_types T
-    val (paramTs, (inargTs, outargTs)) = split_mode (iss, is) Ts
-    val paramTs' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) iss paramTs 
-  in
-    (paramTs' @ inargTs) ---> (mk_predT compfuns (mk_tupleT outargTs))
-  end;
-
-fun sizelim_funT_of compfuns (iss, is) T =
-  let
-    val Ts = binder_types T
-    val (paramTs, (inargTs, outargTs)) = split_mode (iss, is) Ts
-    val paramTs' = map2 (fn SOME is => sizelim_funT_of compfuns ([], is) | NONE => I) iss paramTs 
-  in
-    (paramTs' @ inargTs @ [@{typ "code_numeral"}]) ---> (mk_predT compfuns (mk_tupleT outargTs))
-  end;  
-
-fun mk_fun_of compfuns thy (name, T) mode = 
-  Const (predfun_name_of thy name mode, funT_of compfuns mode T)
-
-fun mk_sizelim_fun_of compfuns thy (name, T) mode =
-  Const (sizelim_function_name_of thy name mode, sizelim_funT_of compfuns mode T)
-  
-fun mk_generator_of compfuns thy (name, T) mode = 
-  Const (generator_name_of thy name mode, sizelim_funT_of compfuns mode T)
-
-
-structure PredicateCompFuns =
-struct
-
-fun mk_predT T = Type (@{type_name "Predicate.pred"}, [T])
-
-fun dest_predT (Type (@{type_name "Predicate.pred"}, [T])) = T
-  | dest_predT T = raise TYPE ("dest_predT", [T], []);
-
-fun mk_bot T = Const (@{const_name Orderings.bot}, mk_predT T);
-
-fun mk_single t =
-  let val T = fastype_of t
-  in Const(@{const_name Predicate.single}, T --> mk_predT T) $ t end;
-
-fun mk_bind (x, f) =
-  let val T as Type ("fun", [_, U]) = fastype_of f
-  in
-    Const (@{const_name Predicate.bind}, fastype_of x --> T --> U) $ x $ f
-  end;
-
-val mk_sup = HOLogic.mk_binop @{const_name sup};
-
-fun mk_if cond = Const (@{const_name Predicate.if_pred},
-  HOLogic.boolT --> mk_predT HOLogic.unitT) $ cond;
-
-fun mk_not t = let val T = mk_predT HOLogic.unitT
-  in Const (@{const_name Predicate.not_pred}, T --> T) $ t end
-
-fun mk_Enum f =
-  let val T as Type ("fun", [T', _]) = fastype_of f
-  in
-    Const (@{const_name Predicate.Pred}, T --> mk_predT T') $ f    
-  end;
-
-fun mk_Eval (f, x) =
-  let
-    val T = fastype_of x
-  in
-    Const (@{const_name Predicate.eval}, mk_predT T --> T --> HOLogic.boolT) $ f $ x
-  end;
-
-fun mk_map T1 T2 tf tp = Const (@{const_name Predicate.map},
-  (T1 --> T2) --> mk_predT T1 --> mk_predT T2) $ tf $ tp;
-
-val lift_pred = I
-
-val compfuns = CompilationFuns {mk_predT = mk_predT, dest_predT = dest_predT, mk_bot = mk_bot,
-  mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if, mk_not = mk_not,
-  mk_map = mk_map, lift_pred = lift_pred};
-
-end;
-
-(* termify_code:
-val termT = Type ("Code_Eval.term", []);
-fun termifyT T = HOLogic.mk_prodT (T, HOLogic.unitT --> termT)
-*)
-(*
-fun lift_random random =
-  let
-    val T = dest_randomT (fastype_of random)
-  in
-    mk_scomp (random,
-      mk_fun_comp (HOLogic.pair_const (PredicateCompFuns.mk_predT T) @{typ Random.seed},
-        mk_fun_comp (Const (@{const_name Predicate.single}, T --> (PredicateCompFuns.mk_predT T)),
-          Const (@{const_name "fst"}, HOLogic.mk_prodT (T, @{typ "unit => term"}) --> T)))) 
-  end;
-*)
- 
-structure RPredCompFuns =
-struct
-
-fun mk_rpredT T =
-  @{typ "Random.seed"} --> HOLogic.mk_prodT (PredicateCompFuns.mk_predT T, @{typ "Random.seed"})
-
-fun dest_rpredT (Type ("fun", [_,
-  Type (@{type_name "*"}, [Type (@{type_name "Predicate.pred"}, [T]), _])])) = T
-  | dest_rpredT T = raise TYPE ("dest_rpredT", [T], []); 
-
-fun mk_bot T = Const(@{const_name RPred.bot}, mk_rpredT T)
-
-fun mk_single t =
-  let
-    val T = fastype_of t
-  in
-    Const (@{const_name RPred.return}, T --> mk_rpredT T) $ t
-  end;
-
-fun mk_bind (x, f) =
-  let
-    val T as (Type ("fun", [_, U])) = fastype_of f
-  in
-    Const (@{const_name RPred.bind}, fastype_of x --> T --> U) $ x $ f
-  end
-
-val mk_sup = HOLogic.mk_binop @{const_name RPred.supp}
-
-fun mk_if cond = Const (@{const_name RPred.if_rpred},
-  HOLogic.boolT --> mk_rpredT HOLogic.unitT) $ cond;
-
-fun mk_not t = error "Negation is not defined for RPred"
-
-fun mk_map t = error "FIXME" (*FIXME*)
-
-fun lift_pred t =
-  let
-    val T = PredicateCompFuns.dest_predT (fastype_of t)
-    val lift_predT = PredicateCompFuns.mk_predT T --> mk_rpredT T 
-  in
-    Const (@{const_name "RPred.lift_pred"}, lift_predT) $ t  
-  end;
-
-val compfuns = CompilationFuns {mk_predT = mk_rpredT, dest_predT = dest_rpredT, mk_bot = mk_bot,
-    mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if, mk_not = mk_not,
-    mk_map = mk_map, lift_pred = lift_pred};
-
-end;
-(* for external use with interactive mode *)
-val rpred_compfuns = RPredCompFuns.compfuns;
-
-fun lift_random random =
-  let
-    val T = dest_randomT (fastype_of random)
-  in
-    Const (@{const_name lift_random}, (@{typ Random.seed} -->
-      HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed})) --> 
-      RPredCompFuns.mk_rpredT T) $ random
-  end;
- 
-(* Mode analysis *)
-
-(*** check if a term contains only constructor functions ***)
-fun is_constrt thy =
-  let
-    val cnstrs = flat (maps
-      (map (fn (_, (Tname, _, cs)) => map (apsnd (rpair Tname o length)) cs) o #descr o snd)
-      (Symtab.dest (Datatype.get_all thy)));
-    fun check t = (case strip_comb t of
-        (Free _, []) => true
-      | (Const (s, T), ts) => (case (AList.lookup (op =) cnstrs s, body_type T) of
-            (SOME (i, Tname), Type (Tname', _)) => length ts = i andalso Tname = Tname' andalso forall check ts
-          | _ => false)
-      | _ => false)
-  in check end;
-
-(*** check if a type is an equality type (i.e. doesn't contain fun)
-  FIXME this is only an approximation ***)
-fun is_eqT (Type (s, Ts)) = s <> "fun" andalso forall is_eqT Ts
-  | is_eqT _ = true;
-
-fun term_vs tm = fold_aterms (fn Free (x, T) => cons x | _ => I) tm [];
-val terms_vs = distinct (op =) o maps term_vs;
-
-(** collect all Frees in a term (with duplicates!) **)
-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 [[]];
-     
-(* FIXME: should be in library - map_prod *)
-fun cprod ([], ys) = []
-  | cprod (x :: xs, ys) = map (pair x) ys @ cprod (xs, ys);
-
-fun cprods xss = foldr (map op :: o cprod) [[]] xss;
-
-
-  
-(*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 = 1 upto length (binder_types (fastype_of t));
-    val default = [Mode (([], ks), ks, [])];
-    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 prfx = 1 upto k
-        in
-          if not (is_prefix op = prfx is) then [] else
-          let val is' = map (fn i => i - k) (List.drop (is, k))
-          in 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)))
-          end
-        end)) (AList.lookup op = modes name)
-
-  in
-    case strip_comb (Envir.eta_contract t) of
-      (Const (name, _), args) => the_default default (mk_modes name args)
-    | (Var ((name, _), _), args) => the (mk_modes name args)
-    | (Free (name, _), args) => the (mk_modes name args)
-    | (Abs _, []) => error "Abs at param position" (* modes_of_param default modes t *)
-    | _ => default
-  end
-  
-fun select_mode_prem thy modes vs ps =
-  find_first (is_some o snd) (ps ~~ map
-    (fn Prem (us, t) => find_first (fn Mode (_, is, _) =>
-          let
-            val (in_ts, out_ts) = split_smode is us;
-            val (out_ts', in_ts') = List.partition (is_constrt thy) out_ts;
-            val vTs = maps term_vTs out_ts';
-            val dupTs = map snd (duplicates (op =) vTs) @
-              List.mapPartial (AList.lookup (op =) vTs) vs;
-          in
-            terms_vs (in_ts @ in_ts') subset vs andalso
-            forall (is_eqT o fastype_of) in_ts' andalso
-            term_vs t subset vs andalso
-            forall is_eqT dupTs
-          end)
-            (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
-            terms_vs us subset vs andalso
-            term_vs t subset vs)
-            (modes_of_term modes t handle Option =>
-               error ("Bad predicate: " ^ Syntax.string_of_term_global thy t))
-      | Sidecond t => if term_vs t subset vs then SOME (Mode (([], []), [], []))
-          else NONE
-      ) ps);
-
-fun fold_prem f (Prem (args, _)) = fold f args
-  | fold_prem f (Negprem (args, _)) = fold f args
-  | fold_prem f (Sidecond t) = f t
-
-fun all_subsets [] = [[]]
-  | all_subsets (x::xs) = let val xss' = all_subsets xs in xss' @ (map (cons x) xss') end
-
-fun generator vTs v = 
-  let
-    val T = the (AList.lookup (op =) vTs v)
-  in
-    (Generator (v, T), Mode (([], []), [], []))
-  end;
-
-fun gen_prem (Prem (us, t)) = GeneratorPrem (us, t) 
-  | gen_prem _ = error "gen_prem : invalid input for gen_prem"
-
-fun param_gen_prem param_vs (p as Prem (us, t as Free (v, _))) =
-  if member (op =) param_vs v then
-    GeneratorPrem (us, t)
-  else p  
-  | param_gen_prem param_vs p = p
-  
-fun check_mode_clause with_generator thy param_vs modes gen_modes (iss, is) (ts, ps) =
-  let
-    val modes' = modes @ List.mapPartial
-      (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
-        (param_vs ~~ iss);
-    val gen_modes' = gen_modes @ List.mapPartial
-      (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
-        (param_vs ~~ iss);  
-    val vTs = distinct (op =) ((fold o fold_prem) Term.add_frees ps (fold Term.add_frees ts []))
-    val prem_vs = distinct (op =) ((fold o fold_prem) Term.add_free_names ps [])
-    fun check_mode_prems acc_ps vs [] = SOME (acc_ps, vs)
-      | check_mode_prems acc_ps vs ps = (case select_mode_prem thy modes' vs ps of
-          NONE =>
-            (if with_generator then
-              (case select_mode_prem thy gen_modes' vs ps of
-                  SOME (p, SOME mode) => check_mode_prems ((gen_prem p, mode) :: acc_ps) 
-                  (case p of Prem (us, _) => vs union terms_vs us | _ => vs)
-                  (filter_out (equal p) ps)
-                | NONE =>
-                  let 
-                    val all_generator_vs = all_subsets (prem_vs \\ vs) |> sort (int_ord o (pairself length))
-                  in
-                    case (find_first (fn generator_vs => is_some
-                      (select_mode_prem thy modes' (vs union generator_vs) ps)) all_generator_vs) of
-                      SOME generator_vs => check_mode_prems ((map (generator vTs) generator_vs) @ acc_ps)
-                        (vs union generator_vs) ps
-                    | NONE => NONE
-                  end)
-            else
-              NONE)
-        | SOME (p, SOME mode) => check_mode_prems ((if with_generator then param_gen_prem param_vs p else p, mode) :: acc_ps) 
-            (case p of Prem (us, _) => vs union terms_vs us | _ => vs)
-            (filter_out (equal p) ps))
-    val (in_ts, in_ts') = List.partition (is_constrt thy) (fst (split_smode is ts));
-    val in_vs = terms_vs in_ts;
-    val concl_vs = terms_vs ts
-  in
-    if forall is_eqT (map snd (duplicates (op =) (maps term_vTs in_ts))) andalso
-    forall (is_eqT o fastype_of) in_ts' then
-      case check_mode_prems [] (param_vs union in_vs) ps of
-         NONE => NONE
-       | SOME (acc_ps, vs) =>
-         if with_generator then
-           SOME (ts, (rev acc_ps) @ (map (generator vTs) (concl_vs \\ vs))) 
-         else
-           if concl_vs subset vs then SOME (ts, rev acc_ps) else NONE
-    else NONE
-  end;
-
-fun check_modes_pred with_generator thy param_vs preds modes gen_modes (p, ms) =
-  let val SOME rs = AList.lookup (op =) preds p
-  in (p, List.filter (fn m => case find_index
-    (is_none o check_mode_clause with_generator thy param_vs modes gen_modes m) rs of
-      ~1 => true
-    | i => (Output.tracing ("Clause " ^ string_of_int (i + 1) ^ " of " ^
-      p ^ " violates mode " ^ string_of_mode m); false)) ms)
-  end;
-
-fun get_modes_pred with_generator thy param_vs preds modes gen_modes (p, ms) =
-  let
-    val SOME rs = AList.lookup (op =) preds p 
-  in
-    (p, map (fn m =>
-      (m, map (the o check_mode_clause with_generator thy param_vs modes gen_modes m) rs)) ms)
-  end;
-  
-fun fixp f (x : (string * mode list) list) =
-  let val y = f x
-  in if x = y then x else fixp f y end;
-
-fun modes_of_arities arities =
-  (map (fn (s, (ks, k)) => (s, cprod (cprods (map
-            (fn NONE => [NONE]
-              | SOME k' => map SOME (subsets 1 k')) ks),
-            subsets 1 k))) arities)
-  
-fun infer_modes with_generator thy extra_modes arities param_vs preds =
-  let
-    val modes =
-      fixp (fn modes =>
-        map (check_modes_pred with_generator thy param_vs preds (modes @ extra_modes) []) modes)
-          (modes_of_arities arities)
-  in
-    map (get_modes_pred with_generator thy param_vs preds (modes @ extra_modes) []) modes
-  end;
-
-fun remove_from rem [] = []
-  | remove_from rem ((k, vs) :: xs) =
-    (case AList.lookup (op =) rem k of
-      NONE => (k, vs)
-    | SOME vs' => (k, vs \\ vs'))
-    :: remove_from rem xs
-    
-fun infer_modes_with_generator thy extra_modes arities param_vs preds =
-  let
-    val prednames = map fst preds
-    val extra_modes = all_modes_of thy
-    val gen_modes = all_generator_modes_of thy
-      |> filter_out (fn (name, _) => member (op =) prednames name)
-    val starting_modes = remove_from extra_modes (modes_of_arities arities) 
-    val modes =
-      fixp (fn modes =>
-        map (check_modes_pred true thy param_vs preds extra_modes (gen_modes @ modes)) modes)
-         starting_modes 
-  in
-    map (get_modes_pred true thy param_vs preds extra_modes (gen_modes @ modes)) modes
-  end;
-
-(* term construction *)
-
-fun mk_v (names, vs) s T = (case AList.lookup (op =) vs s of
-      NONE => (Free (s, T), (names, (s, [])::vs))
-    | SOME xs =>
-        let
-          val s' = Name.variant names s;
-          val v = Free (s', T)
-        in
-          (v, (s'::names, AList.update (op =) (s, v::xs) vs))
-        end);
-
-fun distinct_v (Free (s, T)) nvs = mk_v nvs s T
-  | distinct_v (t $ u) nvs =
-      let
-        val (t', nvs') = distinct_v t nvs;
-        val (u', nvs'') = distinct_v u nvs';
-      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) false [] 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 size thy compfuns (NONE, t) = t
-  | compile_param size thy compfuns (m as SOME (Mode ((iss, is'), is, ms)), t) =
-   let
-     val (f, args) = strip_comb (Envir.eta_contract t)
-     val (params, args') = chop (length ms) args
-     val params' = map (compile_param size thy compfuns) (ms ~~ params)
-     val mk_fun_of = case size of NONE => mk_fun_of | SOME _ => mk_sizelim_fun_of
-     val funT_of = case size of NONE => funT_of | SOME _ => sizelim_funT_of
-     val f' =
-       case f of
-         Const (name, T) =>
-           mk_fun_of compfuns thy (name, T) (iss, is')
-       | Free (name, T) => Free (name, funT_of compfuns (iss, is') T)
-       | _ => error ("PredicateCompiler: illegal parameter term")
-   in list_comb (f', params' @ args') end
-   
-fun compile_expr size thy ((Mode (mode, is, ms)), t) =
-  case strip_comb t of
-    (Const (name, T), params) =>
-       let
-         val params' = map (compile_param size thy PredicateCompFuns.compfuns) (ms ~~ params)
-         val mk_fun_of = case size of NONE => mk_fun_of | SOME _ => mk_sizelim_fun_of
-       in
-         list_comb (mk_fun_of PredicateCompFuns.compfuns thy (name, T) mode, params')
-       end
-  | (Free (name, T), args) =>
-       let 
-         val funT_of = case size of NONE => funT_of | SOME _ => sizelim_funT_of 
-       in
-         list_comb (Free (name, funT_of PredicateCompFuns.compfuns ([], is) T), args)
-       end;
-       
-fun compile_gen_expr size thy compfuns ((Mode (mode, is, ms)), t) =
-  case strip_comb t of
-    (Const (name, T), params) =>
-      let
-        val params' = map (compile_param size thy compfuns) (ms ~~ params)
-      in
-        list_comb (mk_generator_of compfuns thy (name, T) mode, params')
-      end
-    | (Free (name, T), args) =>
-      list_comb (Free (name, sizelim_funT_of RPredCompFuns.compfuns ([], is) T), args)
-          
-(** specific rpred functions -- move them to the correct place in this file *)
-
-(* uncommented termify code; causes more trouble than expected at first *) 
-(*
-fun mk_valtermify_term (t as Const (c, T)) = HOLogic.mk_prod (t, Abs ("u", HOLogic.unitT, HOLogic.reflect_term t))
-  | mk_valtermify_term (Free (x, T)) = Free (x, termifyT T) 
-  | mk_valtermify_term (t1 $ t2) =
-    let
-      val T = fastype_of t1
-      val (T1, T2) = dest_funT T
-      val t1' = mk_valtermify_term t1
-      val t2' = mk_valtermify_term t2
-    in
-      Const ("Code_Eval.valapp", termifyT T --> termifyT T1 --> termifyT T2) $ t1' $ t2'
-    end
-  | mk_valtermify_term _ = error "Not a valid term for mk_valtermify_term"
-*)
-
-fun compile_clause compfuns size final_term thy all_vs param_vs (iss, is) inp (ts, moded_ps) =
-  let
-    fun check_constrt t (names, eqs) =
-      if is_constrt thy t then (t, (names, eqs)) else
-        let
-          val s = Name.variant names "x";
-          val v = Free (s, fastype_of t)
-        in (v, (s::names, HOLogic.mk_eq (v, t)::eqs)) end;
-
-    val (in_ts, out_ts) = split_smode is ts;
-    val (in_ts', (all_vs', eqs)) =
-      fold_map check_constrt in_ts (all_vs, []);
-
-    fun compile_prems out_ts' vs names [] =
-          let
-            val (out_ts'', (names', eqs')) =
-              fold_map check_constrt out_ts' (names, []);
-            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)
-          end
-      | compile_prems out_ts vs names ((p, mode as Mode ((_, is), _, _)) :: ps) =
-          let
-            val vs' = distinct (op =) (flat (vs :: map term_vs out_ts));
-            val (out_ts', (names', eqs)) =
-              fold_map check_constrt out_ts (names, [])
-            val (out_ts'', (names'', constr_vs')) = fold_map distinct_v
-              out_ts' ((names', map (rpair []) vs))
-            val (compiled_clause, rest) = case p of
-               Prem (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 = lift_pred compfuns
-                     (list_comb (compile_expr size thy (mode, t), args))                     
-                   val rest = compile_prems out_ts''' vs' names'' ps
-                 in
-                   (u, rest)
-                 end
-             | Negprem (us, t) =>
-                 let
-                   val (in_ts, out_ts''') = split_smode is us
-                   val u = lift_pred compfuns
-                     (mk_not PredicateCompFuns.compfuns (list_comb (compile_expr NONE thy (mode, t), in_ts)))
-                   val rest = compile_prems out_ts''' vs' names'' ps
-                 in
-                   (u, rest)
-                 end
-             | Sidecond t =>
-                 let
-                   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 = list_comb (compile_gen_expr size thy compfuns (mode, t), args)
-                   val rest = compile_prems out_ts''' vs' names'' ps
-                 in
-                   (u, rest)
-                 end
-             | Generator (v, T) =>
-                 let
-                   val u = lift_random (HOLogic.mk_random T @{term "1::code_numeral"})
-                   val rest = compile_prems [Free (v, T)]  vs' names'' ps;
-                 in
-                   (u, rest)
-                 end
-          in
-            compile_match thy compfuns 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;
-  in
-    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 =
-  let
-    val (Ts1, (Us1, Us2)) = split_mode mode (binder_types T)
-    val funT_of = if use_size then sizelim_funT_of else funT_of 
-    val Ts1' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) (fst mode) Ts1
-    val xnames = Name.variant_list (all_vs @ param_vs)
-      (map (fn i => "x" ^ string_of_int i) (snd mode));
-    val size_name = Name.variant (all_vs @ param_vs @ xnames) "size"
-    (* termify code: val xs = map2 (fn s => fn T => Free (s, termifyT T)) xnames Us1; *)
-    val xs = map2 (fn s => fn T => Free (s, T)) xnames Us1;
-    val params = map2 (fn s => fn T => Free (s, T)) param_vs Ts1'
-    val size = Free (size_name, @{typ "code_numeral"})
-    val decr_size =
-      if use_size then
-        SOME (Const ("HOL.minus_class.minus", @{typ "code_numeral => code_numeral => code_numeral"})
-          $ size $ Const ("HOL.one_class.one", @{typ "Code_Numeral.code_numeral"}))
-      else
-        NONE
-    val cl_ts =
-      map (compile_clause compfuns decr_size (fn out_ts => mk_single compfuns (mk_tuple out_ts))
-        thy all_vs param_vs mode (mk_tuple xs)) moded_cls;
-    val t = foldr1 (mk_sup compfuns) cl_ts
-    val T' = mk_predT compfuns (mk_tupleT Us2)
-    val size_t = Const (@{const_name "If"}, @{typ bool} --> T' --> T' --> T')
-      $ HOLogic.mk_eq (size, @{term "0 :: code_numeral"})
-      $ mk_bot compfuns (dest_predT compfuns T') $ t
-    val fun_const = mk_fun_of compfuns thy (s, T) mode
-    val eq = if use_size then
-      (list_comb (fun_const, params @ xs @ [size]), size_t)
-    else
-      (list_comb (fun_const, params @ xs), t)
-  in
-    HOLogic.mk_Trueprop (HOLogic.mk_eq eq)
-  end;
-  
-(* special setup for simpset *)                  
-val HOL_basic_ss' = HOL_basic_ss setSolver 
-  (mk_solver "all_tac_solver" (fn _ => fn _ => all_tac))
-
-(* Definition of executable functions and their intro and elim rules *)
-
-fun print_arities arities = tracing ("Arities:\n" ^
-  cat_lines (map (fn (s, (ks, k)) => s ^ ": " ^
-    space_implode " -> " (map
-      (fn NONE => "X" | SOME k' => string_of_int k')
-        (ks @ [SOME k]))) arities));
-
-fun mk_Eval_of ((x, T), NONE) names = (x, names)
-  | mk_Eval_of ((x, T), SOME mode) names = let
-  val Ts = binder_types T
-  val argnames = Name.variant_list names
-        (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
-  val args = map Free (argnames ~~ Ts)
-  val (inargs, outargs) = split_smode mode args
-  val r = PredicateCompFuns.mk_Eval (list_comb (x, inargs), mk_tuple outargs)
-  val t = fold_rev lambda args r 
-in
-  (t, argnames @ names)
-end;
-
-fun create_intro_elim_rule (mode as (iss, is)) defthm mode_id funT pred thy =
-let
-  val Ts = binder_types (fastype_of pred)
-  val funtrm = Const (mode_id, funT)
-  val argnames = Name.variant_list []
-        (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
-  val (Ts1, Ts2) = chop (length iss) Ts;
-  val Ts1' = map2 (fn NONE => I | SOME is => funT_of (PredicateCompFuns.compfuns) ([], is)) iss Ts1
-  val args = map Free (argnames ~~ (Ts1' @ Ts2))
-  val (params, ioargs) = chop (length iss) args
-  val (inargs, outargs) = split_smode is ioargs
-  val param_names = Name.variant_list argnames
-    (map (fn i => "p" ^ string_of_int i) (1 upto (length iss)))
-  val param_vs = map Free (param_names ~~ Ts1)
-  val (params', names) = fold_map mk_Eval_of ((params ~~ Ts1) ~~ iss) []
-  val predpropI = HOLogic.mk_Trueprop (list_comb (pred, param_vs @ ioargs))
-  val predpropE = HOLogic.mk_Trueprop (list_comb (pred, params' @ ioargs))
-  val param_eqs = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (param_vs ~~ params')
-  val funargs = params @ inargs
-  val funpropE = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, funargs),
-                  if null outargs then Free("y", HOLogic.unitT) else mk_tuple outargs))
-  val funpropI = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, funargs),
-                   mk_tuple outargs))
-  val introtrm = Logic.list_implies (predpropI :: param_eqs, funpropI)
-  val simprules = [defthm, @{thm eval_pred},
-                   @{thm "split_beta"}, @{thm "fst_conv"}, @{thm "snd_conv"}]
-  val unfolddef_tac = Simplifier.asm_full_simp_tac (HOL_basic_ss addsimps simprules) 1
-  val introthm = Goal.prove (ProofContext.init thy) (argnames @ param_names @ ["y"]) [] introtrm (fn {...} => unfolddef_tac)
-  val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT));
-  val elimtrm = Logic.list_implies ([funpropE, Logic.mk_implies (predpropE, P)], P)
-  val elimthm = Goal.prove (ProofContext.init thy) (argnames @ param_names @ ["y", "P"]) [] elimtrm (fn {...} => unfolddef_tac)
-in 
-  (introthm, elimthm)
-end;
-
-fun create_constname_of_mode thy prefix name mode = 
-  let
-    fun string_of_mode mode = if null mode then "0"
-      else space_implode "_" (map string_of_int mode)
-    val HOmode = space_implode "_and_"
-      (fold (fn NONE => I | SOME mode => cons (string_of_mode mode)) (fst mode) [])
-  in
-    (Sign.full_bname thy (prefix ^ (Long_Name.base_name name))) ^
-      (if HOmode = "" then "_" else "_for_" ^ HOmode ^ "_yields_") ^ (string_of_mode (snd mode))
-  end;
-  
-fun create_definitions preds (name, modes) thy =
-  let
-    val compfuns = PredicateCompFuns.compfuns
-    val T = AList.lookup (op =) preds name |> the
-    fun create_definition (mode as (iss, is)) thy = let
-      val mode_cname = create_constname_of_mode thy "" name mode
-      val mode_cbasename = Long_Name.base_name mode_cname
-      val Ts = binder_types T
-      val (Ts1, Ts2) = chop (length iss) Ts
-      val (Us1, Us2) =  split_smode is Ts2
-      val Ts1' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) iss Ts1
-      val funT = (Ts1' @ Us1) ---> (mk_predT compfuns (mk_tupleT Us2))
-      val names = Name.variant_list []
-        (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
-      val xs = map Free (names ~~ (Ts1' @ Ts2));                   
-      val (xparams, xargs) = chop (length iss) xs;
-      val (xins, xouts) = split_smode is xargs 
-      val (xparams', names') = fold_map mk_Eval_of ((xparams ~~ Ts1) ~~ iss) names
-      fun mk_split_lambda [] t = lambda (Free (Name.variant names' "x", HOLogic.unitT)) t
-        | mk_split_lambda [x] t = lambda x t
-        | mk_split_lambda xs t =
-        let
-          fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t))
-            | mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t))
-        in
-          mk_split_lambda' xs t
-        end;
-      val predterm = PredicateCompFuns.mk_Enum (mk_split_lambda xouts
-        (list_comb (Const (name, T), xparams' @ xargs)))
-      val lhs = list_comb (Const (mode_cname, funT), xparams @ xins)
-      val def = Logic.mk_equals (lhs, predterm)
-      val ([definition], thy') = thy |>
-        Sign.add_consts_i [(Binding.name mode_cbasename, funT, NoSyn)] |>
-        PureThy.add_defs false [((Binding.name (mode_cbasename ^ "_def"), def), [])]
-      val (intro, elim) =
-        create_intro_elim_rule mode definition mode_cname funT (Const (name, T)) thy'
-      in thy' |> add_predfun name mode (mode_cname, definition, intro, elim)
-        |> PureThy.store_thm (Binding.name (mode_cbasename ^ "I"), intro) |> snd
-        |> PureThy.store_thm (Binding.name (mode_cbasename ^ "E"), elim)  |> snd
-        |> Theory.checkpoint
-      end;
-  in
-    fold create_definition modes thy
-  end;
-
-fun sizelim_create_definitions preds (name, modes) thy =
-  let
-    val T = AList.lookup (op =) preds name |> the
-    fun create_definition mode thy =
-      let
-        val mode_cname = create_constname_of_mode thy "sizelim_" name mode
-        val funT = sizelim_funT_of PredicateCompFuns.compfuns mode T
-      in
-        thy |> Sign.add_consts_i [(Binding.name (Long_Name.base_name mode_cname), funT, NoSyn)]
-        |> set_sizelim_function_name name mode mode_cname 
-      end;
-  in
-    fold create_definition modes thy
-  end;
-    
-fun rpred_create_definitions preds (name, modes) thy =
-  let
-    val T = AList.lookup (op =) preds name |> the
-    fun create_definition mode thy =
-      let
-        val mode_cname = create_constname_of_mode thy "gen_" name mode
-        val funT = sizelim_funT_of RPredCompFuns.compfuns mode T
-      in
-        thy |> Sign.add_consts_i [(Binding.name (Long_Name.base_name mode_cname), funT, NoSyn)]
-        |> set_generator_name name mode mode_cname 
-      end;
-  in
-    fold create_definition modes thy
-  end;
-  
-(* Proving equivalence of term *)
-
-fun is_Type (Type _) = true
-  | is_Type _ = false
-
-(* returns true if t is an application of an datatype constructor *)
-(* which then consequently would be splitted *)
-(* else false *)
-fun is_constructor thy t =
-  if (is_Type (fastype_of t)) then
-    (case Datatype.get_info thy ((fst o dest_Type o fastype_of) t) of
-      NONE => false
-    | SOME info => (let
-      val constr_consts = maps (fn (_, (_, _, constrs)) => map fst constrs) (#descr info)
-      val (c, _) = strip_comb t
-      in (case c of
-        Const (name, _) => name mem_string constr_consts
-        | _ => false) end))
-  else false
-
-(* MAJOR FIXME:  prove_params should be simple
- - different form of introrule for parameters ? *)
-fun prove_param thy (NONE, t) = TRY (rtac @{thm refl} 1)
-  | prove_param thy (m as SOME (Mode (mode, is, ms)), t) =
-  let
-    val  (f, args) = strip_comb (Envir.eta_contract t)
-    val (params, _) = chop (length ms) args
-    val f_tac = case f of
-      Const (name, T) => simp_tac (HOL_basic_ss addsimps 
-         (@{thm eval_pred}::(predfun_definition_of thy name mode)::
-         @{thm "Product_Type.split_conv"}::[])) 1
-    | Free _ => TRY (rtac @{thm refl} 1)
-    | Abs _ => error "prove_param: No valid parameter term"
-  in
-    REPEAT_DETERM (etac @{thm thin_rl} 1)
-    THEN REPEAT_DETERM (rtac @{thm ext} 1)
-    THEN print_tac "prove_param"
-    THEN f_tac
-    THEN print_tac "after simplification in prove_args"
-    THEN (EVERY (map (prove_param thy) (ms ~~ params)))
-    THEN (REPEAT_DETERM (atac 1))
-  end
-
-fun prove_expr thy (Mode (mode, is, ms), t, us) (premposition : int) =
-  case strip_comb t of
-    (Const (name, T), args) =>  
-      let
-        val introrule = predfun_intro_of thy name mode
-        val (args1, args2) = chop (length ms) args
-      in
-        rtac @{thm bindI} 1
-        THEN print_tac "before intro rule:"
-        (* for the right assumption in first position *)
-        THEN rotate_tac premposition 1
-        THEN debug_tac (Display.string_of_thm (ProofContext.init thy) introrule)
-        THEN rtac introrule 1
-        THEN print_tac "after intro rule"
-        (* work with parameter arguments *)
-        THEN (atac 1)
-        THEN (print_tac "parameter goal")
-        THEN (EVERY (map (prove_param thy) (ms ~~ args1)))
-        THEN (REPEAT_DETERM (atac 1))
-      end
-  | _ => rtac @{thm bindI} 1 THEN atac 1
-
-fun SOLVED tac st = FILTER (fn st' => nprems_of st' = nprems_of st - 1) tac st; 
-
-fun SOLVEDALL tac st = FILTER (fn st' => nprems_of st' = 0) tac st
-
-fun prove_match thy (out_ts : term list) = let
-  fun get_case_rewrite t =
-    if (is_constructor thy t) then let
-      val case_rewrites = (#case_rewrites (Datatype.the_info thy
-        ((fst o dest_Type o fastype_of) t)))
-      in case_rewrites @ (flat (map get_case_rewrite (snd (strip_comb t)))) end
-    else []
-  val simprules = @{thm "unit.cases"} :: @{thm "prod.cases"} :: (flat (map get_case_rewrite out_ts))
-(* replace TRY by determining if it necessary - are there equations when calling compile match? *)
-in
-   (* make this simpset better! *)
-  asm_simp_tac (HOL_basic_ss' addsimps simprules) 1
-  THEN print_tac "after prove_match:"
-  THEN (DETERM (TRY (EqSubst.eqsubst_tac (ProofContext.init thy) [0] [@{thm "HOL.if_P"}] 1
-         THEN (REPEAT_DETERM (rtac @{thm conjI} 1 THEN (SOLVED (asm_simp_tac HOL_basic_ss 1))))
-         THEN (SOLVED (asm_simp_tac HOL_basic_ss 1)))))
-  THEN print_tac "after if simplification"
-end;
-
-(* corresponds to compile_fun -- maybe call that also compile_sidecond? *)
-
-fun prove_sidecond thy modes t =
-  let
-    fun preds_of t nameTs = case strip_comb t of 
-      (f as Const (name, T), args) =>
-        if AList.defined (op =) modes name then (name, T) :: nameTs
-          else fold preds_of args nameTs
-      | _ => nameTs
-    val preds = preds_of t []
-    val defs = map
-      (fn (pred, T) => predfun_definition_of thy pred ([], (1 upto (length (binder_types T)))))
-        preds
-  in 
-    (* remove not_False_eq_True when simpset in prove_match is better *)
-    simp_tac (HOL_basic_ss addsimps @{thm not_False_eq_True} :: @{thm eval_pred} :: defs) 1 
-    (* need better control here! *)
-  end
-
-fun prove_clause thy nargs modes (iss, is) (_, clauses) (ts, moded_ps) =
-  let
-    val (in_ts, clause_out_ts) = split_smode is ts;
-    fun prove_prems out_ts [] =
-      (prove_match thy out_ts)
-      THEN asm_simp_tac HOL_basic_ss' 1
-      THEN (rtac (if null clause_out_ts then @{thm singleI_unit} else @{thm singleI}) 1)
-    | prove_prems out_ts ((p, mode as Mode ((iss, is), _, param_modes)) :: ps) =
-      let
-        val premposition = (find_index (equal p) clauses) + nargs
-        val rest_tac = (case p of Prem (us, t) =>
-            let
-              val (_, out_ts''') = split_smode is us
-              val rec_tac = prove_prems out_ts''' ps
-            in
-              print_tac "before clause:"
-              THEN asm_simp_tac HOL_basic_ss 1
-              THEN print_tac "before prove_expr:"
-              THEN prove_expr thy (mode, t, us) premposition
-              THEN print_tac "after prove_expr:"
-              THEN rec_tac
-            end
-          | Negprem (us, t) =>
-            let
-              val (_, out_ts''') = split_smode is us
-              val rec_tac = prove_prems out_ts''' ps
-              val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE)
-              val (_, params) = strip_comb t
-            in
-              rtac @{thm bindI} 1
-              THEN (if (is_some name) then
-                  simp_tac (HOL_basic_ss addsimps [predfun_definition_of thy (the name) (iss, is)]) 1
-                  THEN rtac @{thm not_predI} 1
-                  THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
-                  THEN (REPEAT_DETERM (atac 1))
-                  (* FIXME: work with parameter arguments *)
-                  THEN (EVERY (map (prove_param thy) (param_modes ~~ params)))
-                else
-                  rtac @{thm not_predI'} 1)
-                  THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
-              THEN rec_tac
-            end
-          | Sidecond t =>
-           rtac @{thm bindI} 1
-           THEN rtac @{thm if_predI} 1
-           THEN print_tac "before sidecond:"
-           THEN prove_sidecond thy modes t
-           THEN print_tac "after sidecond:"
-           THEN prove_prems [] ps)
-      in (prove_match thy out_ts)
-          THEN rest_tac
-      end;
-    val prems_tac = prove_prems in_ts moded_ps
-  in
-    rtac @{thm bindI} 1
-    THEN rtac @{thm singleI} 1
-    THEN prems_tac
-  end;
-
-fun select_sup 1 1 = []
-  | select_sup _ 1 = [rtac @{thm supI1}]
-  | select_sup n i = (rtac @{thm supI2})::(select_sup (n - 1) (i - 1));
-
-fun prove_one_direction thy clauses preds modes pred mode moded_clauses =
-  let
-    val T = the (AList.lookup (op =) preds pred)
-    val nargs = length (binder_types T) - nparams_of thy pred
-    val pred_case_rule = the_elim_of thy pred
-  in
-    REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"}))
-    THEN etac (predfun_elim_of thy pred mode) 1
-    THEN etac pred_case_rule 1
-    THEN (EVERY (map
-           (fn i => EVERY' (select_sup (length moded_clauses) i) i) 
-             (1 upto (length moded_clauses))))
-    THEN (EVERY (map2 (prove_clause thy nargs modes mode) clauses moded_clauses))
-    THEN print_tac "proved one direction"
-  end;
-
-(** Proof in the other direction **)
-
-fun prove_match2 thy out_ts = let
-  fun split_term_tac (Free _) = all_tac
-    | split_term_tac t =
-      if (is_constructor thy t) then let
-        val info = Datatype.the_info thy ((fst o dest_Type o fastype_of) t)
-        val num_of_constrs = length (#case_rewrites info)
-        (* special treatment of pairs -- because of fishing *)
-        val split_rules = case (fst o dest_Type o fastype_of) t of
-          "*" => [@{thm prod.split_asm}] 
-          | _ => 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))
-        THEN (EVERY (map split_term_tac ts))
-      end
-    else all_tac
-  in
-    split_term_tac (mk_tuple out_ts)
-    THEN (DETERM (TRY ((Splitter.split_asm_tac [@{thm "split_if_asm"}] 1) THEN (etac @{thm botE} 2))))
-  end
-
-(* VERY LARGE SIMILIRATIY to function prove_param 
--- join both functions
-*)
-(* TODO: remove function *)
-
-fun prove_param2 thy (NONE, t) = all_tac 
-  | prove_param2 thy (m as SOME (Mode (mode, is, ms)), t) = let
-    val  (f, args) = strip_comb (Envir.eta_contract t)
-    val (params, _) = chop (length ms) args
-    val f_tac = case f of
-        Const (name, T) => full_simp_tac (HOL_basic_ss addsimps 
-           (@{thm eval_pred}::(predfun_definition_of thy name mode)
-           :: @{thm "Product_Type.split_conv"}::[])) 1
-      | Free _ => all_tac
-      | _ => error "prove_param2: illegal parameter term"
-  in  
-    print_tac "before simplification in prove_args:"
-    THEN f_tac
-    THEN print_tac "after simplification in prove_args"
-    THEN (EVERY (map (prove_param2 thy) (ms ~~ params)))
-  end
-
-
-fun prove_expr2 thy (Mode (mode, is, ms), t) = 
-  (case strip_comb t of
-    (Const (name, T), args) =>
-      etac @{thm bindE} 1
-      THEN (REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"})))
-      THEN print_tac "prove_expr2-before"
-      THEN (debug_tac (Syntax.string_of_term_global thy
-        (prop_of (predfun_elim_of thy name mode))))
-      THEN (etac (predfun_elim_of thy name mode) 1)
-      THEN print_tac "prove_expr2"
-      THEN (EVERY (map (prove_param2 thy) (ms ~~ args)))
-      THEN print_tac "finished prove_expr2"      
-    | _ => etac @{thm bindE} 1)
-    
-(* FIXME: what is this for? *)
-(* replace defined by has_mode thy pred *)
-(* TODO: rewrite function *)
-fun prove_sidecond2 thy modes t = let
-  fun preds_of t nameTs = case strip_comb t of 
-    (f as Const (name, T), args) =>
-      if AList.defined (op =) modes name then (name, T) :: nameTs
-        else fold preds_of args nameTs
-    | _ => nameTs
-  val preds = preds_of t []
-  val defs = map
-    (fn (pred, T) => predfun_definition_of thy pred ([], (1 upto (length (binder_types T)))))
-      preds
-  in
-   (* only simplify the one assumption *)
-   full_simp_tac (HOL_basic_ss' addsimps @{thm eval_pred} :: defs) 1 
-   (* need better control here! *)
-   THEN print_tac "after sidecond2 simplification"
-   end
-  
-fun prove_clause2 thy modes pred (iss, is) (ts, ps) i =
-  let
-    val pred_intro_rule = nth (intros_of thy pred) (i - 1)
-    val (in_ts, clause_out_ts) = split_smode is ts;
-    fun prove_prems2 out_ts [] =
-      print_tac "before prove_match2 - last call:"
-      THEN prove_match2 thy out_ts
-      THEN print_tac "after prove_match2 - last call:"
-      THEN (etac @{thm singleE} 1)
-      THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
-      THEN (asm_full_simp_tac HOL_basic_ss' 1)
-      THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
-      THEN (asm_full_simp_tac HOL_basic_ss' 1)
-      THEN SOLVED (print_tac "state before applying intro rule:"
-      THEN (rtac pred_intro_rule 1)
-      (* How to handle equality correctly? *)
-      THEN (print_tac "state before assumption matching")
-      THEN (REPEAT (atac 1 ORELSE 
-         (CHANGED (asm_full_simp_tac HOL_basic_ss' 1)
-          THEN print_tac "state after simp_tac:"))))
-    | prove_prems2 out_ts ((p, mode as Mode ((iss, is), _, param_modes)) :: ps) =
-      let
-        val rest_tac = (case p of
-          Prem (us, t) =>
-          let
-            val (_, out_ts''') = split_smode is us
-            val rec_tac = prove_prems2 out_ts''' ps
-          in
-            (prove_expr2 thy (mode, t)) THEN rec_tac
-          end
-        | Negprem (us, t) =>
-          let
-            val (_, out_ts''') = split_smode is us
-            val rec_tac = prove_prems2 out_ts''' ps
-            val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE)
-            val (_, params) = strip_comb t
-          in
-            print_tac "before neg prem 2"
-            THEN etac @{thm bindE} 1
-            THEN (if is_some name then
-                full_simp_tac (HOL_basic_ss addsimps [predfun_definition_of thy (the name) (iss, is)]) 1 
-                THEN etac @{thm not_predE} 1
-                THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
-                THEN (EVERY (map (prove_param2 thy) (param_modes ~~ params)))
-              else
-                etac @{thm not_predE'} 1)
-            THEN rec_tac
-          end 
-        | Sidecond t =>
-          etac @{thm bindE} 1
-          THEN etac @{thm if_predE} 1
-          THEN prove_sidecond2 thy modes t 
-          THEN prove_prems2 [] ps)
-      in print_tac "before prove_match2:"
-         THEN prove_match2 thy out_ts
-         THEN print_tac "after prove_match2:"
-         THEN rest_tac
-      end;
-    val prems_tac = prove_prems2 in_ts ps 
-  in
-    print_tac "starting prove_clause2"
-    THEN etac @{thm bindE} 1
-    THEN (etac @{thm singleE'} 1)
-    THEN (TRY (etac @{thm Pair_inject} 1))
-    THEN print_tac "after singleE':"
-    THEN prems_tac
-  end;
- 
-fun prove_other_direction thy modes pred mode moded_clauses =
-  let
-    fun prove_clause clause i =
-      (if i < length moded_clauses then etac @{thm supE} 1 else all_tac)
-      THEN (prove_clause2 thy modes pred mode clause i)
-  in
-    (DETERM (TRY (rtac @{thm unit.induct} 1)))
-     THEN (REPEAT_DETERM (CHANGED (rewtac @{thm split_paired_all})))
-     THEN (rtac (predfun_intro_of thy pred mode) 1)
-     THEN (REPEAT_DETERM (rtac @{thm refl} 2))
-     THEN (EVERY (map2 prove_clause moded_clauses (1 upto (length moded_clauses))))
-  end;
-
-(** proof procedure **)
-
-fun prove_pred thy clauses preds modes pred mode (moded_clauses, compiled_term) =
-  let
-    val ctxt = ProofContext.init thy
-    val clauses = the (AList.lookup (op =) clauses pred)
-  in
-    Goal.prove ctxt (Term.add_free_names compiled_term []) [] compiled_term
-      (if !do_proofs then
-        (fn _ =>
-        rtac @{thm pred_iffI} 1
-        THEN prove_one_direction thy clauses preds modes pred mode moded_clauses
-        THEN print_tac "proved one direction"
-        THEN prove_other_direction thy modes pred mode moded_clauses
-        THEN print_tac "proved other direction")
-       else (fn _ => mycheat_tac thy 1))
-  end;
-
-(* composition of mode inference, definition, compilation and proof *)
-
-(** auxillary combinators for table of preds and modes **)
-
-fun map_preds_modes f preds_modes_table =
-  map (fn (pred, modes) =>
-    (pred, map (fn (mode, value) => (mode, f pred mode value)) modes)) preds_modes_table
-
-fun join_preds_modes table1 table2 =
-  map_preds_modes (fn pred => fn mode => fn value =>
-    (value, the (AList.lookup (op =) (the (AList.lookup (op =) table2 pred)) mode))) table1
-    
-fun maps_modes preds_modes_table =
-  map (fn (pred, modes) =>
-    (pred, map (fn (mode, value) => value) modes)) preds_modes_table  
-    
-fun compile_preds compfuns mk_fun_of use_size thy all_vs param_vs preds moded_clauses =
-  map_preds_modes (fn pred => compile_pred compfuns mk_fun_of use_size thy all_vs param_vs pred
-      (the (AList.lookup (op =) preds pred))) moded_clauses  
-  
-fun prove thy clauses preds modes moded_clauses compiled_terms =
-  map_preds_modes (prove_pred thy clauses preds modes)
-    (join_preds_modes moded_clauses compiled_terms)
-
-fun prove_by_skip thy _ _ _ _ compiled_terms =
-  map_preds_modes (fn pred => fn mode => fn t => Drule.standard (SkipProof.make_thm thy t))
-    compiled_terms
-    
-fun prepare_intrs thy prednames =
-  let
-    val intrs = maps (intros_of thy) prednames
-      |> map (Logic.unvarify o prop_of)
-    val nparams = nparams_of thy (hd prednames)
-    val extra_modes = all_modes_of thy |> filter_out (fn (name, _) => member (op =) prednames name)
-    val preds = distinct (op =) (map (dest_Const o fst o (strip_intro_concl nparams)) intrs)
-    val _ $ u = Logic.strip_imp_concl (hd intrs);
-    val params = List.take (snd (strip_comb u), nparams);
-    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 (Ts, Us) = chop nparams (binder_types T)
-    in
-      (AList.update op = (name, these (AList.lookup op = clauses name) @
-        [(ts2, prems)]) clauses,
-       AList.update op = (name, (map (fn U => (case strip_type U of
-                 (Rs as _ :: _, Type ("bool", [])) => SOME (length Rs)
-               | _ => NONE)) Ts,
-             length Us)) arities)
-    end;
-    val (clauses, arities) = fold add_clause intrs ([], []);
-  in (preds, nparams, all_vs, param_vs, extra_modes, clauses, arities) end;
-
-(** main function of predicate compiler **)
-
-fun add_equations_of steps prednames thy =
-  let
-    val _ = Output.tracing ("Starting predicate compiler for predicates " ^ commas prednames ^ "...")
-    val (preds, nparams, all_vs, param_vs, extra_modes, clauses, arities) =
-      prepare_intrs thy prednames
-    val _ = Output.tracing "Infering modes..."
-    val moded_clauses = #infer_modes steps thy extra_modes arities param_vs clauses 
-    val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses
-    val _ = print_modes modes
-    val _ = print_moded_clauses thy moded_clauses
-    val _ = Output.tracing "Defining executable functions..."
-    val thy' = fold (#create_definitions steps preds) modes thy
-      |> Theory.checkpoint
-    val _ = Output.tracing "Compiling equations..."
-    val compiled_terms =
-      (#compile_preds steps) thy' all_vs param_vs preds moded_clauses
-    val _ = print_compiled_terms thy' compiled_terms
-    val _ = Output.tracing "Proving equations..."
-    val result_thms = #prove steps thy' clauses preds (extra_modes @ modes)
-      moded_clauses compiled_terms
-    val qname = #qname steps
-    (* val attrib = gn thy => Attrib.attribute_i thy Code.add_eqn_attrib *)
-    val attrib = fn thy => Attrib.attribute_i thy (Attrib.internal (K (Thm.declaration_attribute
-      (fn thm => Context.mapping (Code.add_eqn thm) I))))
-    val thy'' = fold (fn (name, result_thms) => fn thy => snd (PureThy.add_thmss
-      [((Binding.qualify true (Long_Name.base_name name) (Binding.name qname), result_thms),
-        [attrib thy ])] thy))
-      (maps_modes result_thms) thy'
-      |> Theory.checkpoint
-  in
-    thy''
-  end
-
-fun extend' value_of edges_of key (G, visited) =
-  let
-    val (G', v) = case try (Graph.get_node G) key of
-        SOME v => (G, v)
-      | NONE => (Graph.new_node (key, value_of key) G, value_of key)
-    val (G'', visited') = fold (extend' value_of edges_of) (edges_of (key, v) \\ visited)
-      (G', key :: visited) 
-  in
-    (fold (Graph.add_edge o (pair key)) (edges_of (key, v)) G'', visited')
-  end;
-
-fun extend value_of edges_of key G = fst (extend' value_of edges_of key (G, [])) 
-  
-fun gen_add_equations steps names thy =
-  let
-    val thy' = PredData.map (fold (extend (fetch_pred_data thy) (depending_preds_of thy)) names) thy
-      |> Theory.checkpoint;
-    fun strong_conn_of gr keys =
-      Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr)
-    val scc = strong_conn_of (PredData.get thy') names
-    val thy'' = fold_rev
-      (fn preds => fn thy =>
-        if #are_not_defined steps thy preds then add_equations_of steps preds thy else thy)
-      scc thy' |> Theory.checkpoint
-  in thy'' end
-
-(* different instantiantions of the predicate compiler *)
-
-val add_equations = gen_add_equations
-  {infer_modes = infer_modes false,
-  create_definitions = create_definitions,
-  compile_preds = compile_preds PredicateCompFuns.compfuns mk_fun_of false,
-  prove = prove,
-  are_not_defined = (fn thy => forall (null o modes_of thy)),
-  qname = "equation"}
-
-val add_sizelim_equations = gen_add_equations
-  {infer_modes = infer_modes false,
-  create_definitions = sizelim_create_definitions,
-  compile_preds = compile_preds PredicateCompFuns.compfuns mk_sizelim_fun_of true,
-  prove = prove_by_skip,
-  are_not_defined = (fn thy => fn preds => true), (* TODO *)
-  qname = "sizelim_equation"
-  }
-  
-val add_quickcheck_equations = gen_add_equations
-  {infer_modes = infer_modes_with_generator,
-  create_definitions = rpred_create_definitions,
-  compile_preds = compile_preds RPredCompFuns.compfuns mk_generator_of true,
-  prove = prove_by_skip,
-  are_not_defined = (fn thy => fn preds => true), (* TODO *)
-  qname = "rpred_equation"}
-
-(** user interface **)
-
-(* generation of case rules from user-given introduction rules *)
-
-fun mk_casesrule ctxt nparams introrules =
-  let
-    val intros = map (Logic.unvarify o prop_of) introrules
-    val (pred, (params, args)) = strip_intro_concl nparams (hd intros)
-    val ([propname], ctxt1) = Variable.variant_fixes ["thesis"] ctxt
-    val prop = HOLogic.mk_Trueprop (Free (propname, HOLogic.boolT))
-    val (argnames, ctxt2) = Variable.variant_fixes
-      (map (fn i => "a" ^ string_of_int i) (1 upto (length args))) ctxt1
-    val argvs = map2 (curry Free) argnames (map fastype_of args)
-    fun mk_case intro =
-      let
-        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;
-
-(* code_pred_intro attribute *)
-
-fun attrib f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I);
-
-val code_pred_intros_attrib = attrib add_intro;
-
-local
-
-(* TODO: make TheoryDataFun to GenericDataFun & remove duplication of local theory and theory *)
-(* TODO: must create state to prove multiple cases *)
-fun generic_code_pred prep_const raw_const lthy =
-  let
-    val thy = ProofContext.theory_of lthy
-    val const = prep_const thy raw_const
-    val lthy' = LocalTheory.theory (PredData.map
-        (extend (fetch_pred_data thy) (depending_preds_of thy) const)) lthy
-      |> LocalTheory.checkpoint
-    val thy' = ProofContext.theory_of lthy'
-    val preds = Graph.all_preds (PredData.get thy') [const] |> filter_out (has_elim thy')
-    fun mk_cases const =
-      let
-        val nparams = nparams_of thy' const
-        val intros = intros_of thy' const
-      in mk_casesrule lthy' nparams intros end  
-    val cases_rules = map mk_cases preds
-    val cases =
-      map (fn case_rule => RuleCases.Case {fixes = [],
-        assumes = [("", Logic.strip_imp_prems case_rule)],
-        binds = [], cases = []}) cases_rules
-    val case_env = map2 (fn p => fn c => (Long_Name.base_name p, SOME c)) preds cases
-    val lthy'' = lthy'
-      |> fold Variable.auto_fixes cases_rules 
-      |> ProofContext.add_cases true case_env
-    fun after_qed thms goal_ctxt =
-      let
-        val global_thms = ProofContext.export goal_ctxt
-          (ProofContext.init (ProofContext.theory_of goal_ctxt)) (map the_single thms)
-      in
-        goal_ctxt |> LocalTheory.theory (fold set_elim global_thms #> add_equations [const])
-      end  
-  in
-    Proof.theorem_i NONE after_qed (map (single o (rpair [])) cases_rules) lthy''
-  end;
-
-structure P = OuterParse
-
-in
-
-val code_pred = generic_code_pred (K I);
-val code_pred_cmd = generic_code_pred Code.read_const
-
-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?*)
-
-val _ = OuterSyntax.local_theory_to_proof "code_pred"
-  "prove equations for predicate specified by intro/elim rules"
-  OuterKeyword.thy_goal (P.term_group >> code_pred_cmd)
-
-end
-
-(*FIXME
-- Naming of auxiliary rules necessary?
-- add default code equations P x y z = P_i_i_i x y z
-*)
-
-(* transformation for code generation *)
-
-val eval_ref = ref (NONE : (unit -> term Predicate.pred) option);
-
-(*FIXME turn this into an LCF-guarded preprocessor for comprehensions*)
-fun analyze_compr thy 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);
-    val (body, Ts, fp) = HOLogic.strip_psplits split;
-    val (pred as Const (name, T), all_args) = strip_comb body;
-    val (params, args) = chop (nparams_of thy name) all_args;
-    val user_mode = map_filter I (map_index
-      (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 modes = filter (fn Mode (_, is, _) => is = user_mode)
-      (modes_of_term (all_modes_of thy) (list_comb (pred, params)));
-    val m = case modes
-     of [] => error ("No mode possible for comprehension "
-                ^ Syntax.string_of_term_global thy t_compr)
-      | [m] => m
-      | 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 thy (m, list_comb (pred, params)), inargs);
-    val t_eval = if null outargs then t_pred else let
-        val outargs_bounds = map (fn Bound i => i) outargs;
-        val outargsTs = map (nth Ts) outargs_bounds;
-        val T_pred = HOLogic.mk_tupleT outargsTs;
-        val T_compr = HOLogic.mk_ptupleT fp Ts;
-        val arrange_bounds = map_index I outargs_bounds
-          |> sort (prod_ord (K EQUAL) int_ord)
-          |> map fst;
-        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 t_eval end;
-
-fun eval thy 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.eval_ref", eval_ref) Predicate.map thy t' []) end;
-
-fun values ctxt k t_compr =
-  let
-    val thy = ProofContext.theory_of ctxt;
-    val (T, t) = eval thy t_compr;
-    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
-  end;
-
-fun values_cmd modes 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 ty' = Term.type_of t';
-    val ctxt' = Variable.auto_fixes t' ctxt;
-    val p = PrintMode.with_modes modes (fn () =>
-      Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk,
-        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 _ = 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)));
-
-end;
-
-end;
-
--- a/src/Pure/Concurrent/future.ML	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/Pure/Concurrent/future.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -257,7 +257,7 @@
               "SCHEDULE: disposed " ^ string_of_int (length dead) ^ " dead worker threads")));
 
     val m = if ! do_shutdown then 0 else Multithreading.max_threads_value ();
-    val mm = if m = 9999 then 1 else (m * 3) div 2;
+    val mm = if m = 9999 then 1 else m * 2;
     val l = length (! workers);
     val _ = excessive := l - mm;
     val _ =
--- a/src/Pure/Isar/code.ML	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/Pure/Isar/code.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -505,9 +505,10 @@
 
 (*those following are permissive wrt. to overloaded constants!*)
 
+val head_eqn = dest_Const o fst o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of;
 fun const_typ_eqn thy thm =
   let
-    val (c, ty) = (dest_Const o fst o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm;
+    val (c, ty) = head_eqn thm;
     val c' = AxClass.unoverload_const thy (c, ty);
   in (c', ty) end;
 
@@ -523,8 +524,8 @@
 
 fun same_typscheme thy thms =
   let
-    fun tvars_of t = rev (Term.add_tvars t []);
-    val vss = map (tvars_of o Thm.prop_of) thms;
+    fun tvars_of T = rev (Term.add_tvarsT T []);
+    val vss = map (tvars_of o snd o head_eqn) thms;
     fun inter_sorts vs =
       fold (curry (Sorts.inter_sort (Sign.classes_of thy)) o snd) vs [];
     val sorts = map_transpose inter_sorts vss;
@@ -547,7 +548,7 @@
 fun case_certificate thm =
   let
     val ((head, raw_case_expr), cases) = (apfst Logic.dest_equals
-      o apsnd Logic.dest_conjunctions o Logic.dest_implies o Thm.prop_of) thm;
+      o apsnd Logic.dest_conjunctions o Logic.dest_implies o Thm.plain_prop_of) thm;
     val _ = case head of Free _ => true
       | Var _ => true
       | _ => raise TERM ("case_cert", []);
@@ -759,7 +760,7 @@
 end; (*struct*)
 
 
-(** type-safe interfaces for data depedent on executable code **)
+(** type-safe interfaces for data dependent on executable code **)
 
 functor Code_Data_Fun(Data: CODE_DATA_ARGS): CODE_DATA =
 struct
@@ -779,4 +780,49 @@
 
 end;
 
+(** datastructure to log definitions for preprocessing of the predicate compiler **)
+(* obviously a clone of Named_Thms *)
+
+signature PREDICATE_COMPILE_PREPROC_CONST_DEFS =
+sig
+  val get: Proof.context -> thm list
+  val add_thm: thm -> Context.generic -> Context.generic
+  val del_thm: thm -> Context.generic -> Context.generic
+  
+  val add_attribute : attribute
+  val del_attribute : attribute
+  
+  val add_attrib : Attrib.src
+  
+  val setup: theory -> theory
+end;
+
+structure Predicate_Compile_Preproc_Const_Defs : PREDICATE_COMPILE_PREPROC_CONST_DEFS =
+struct
+
+structure Data = GenericDataFun
+(
+  type T = thm list;
+  val empty = [];
+  val extend = I;
+  fun merge _ = Thm.merge_thms;
+);
+
+val get = Data.get o Context.Proof;
+
+val add_thm = Data.map o Thm.add_thm;
+val del_thm = Data.map o Thm.del_thm;
+
+val add_attribute = Thm.declaration_attribute add_thm;
+val del_attribute = Thm.declaration_attribute del_thm;
+
+val add_attrib = Attrib.internal (K add_attribute)
+
+val setup =
+  Attrib.setup (Binding.name "pred_compile_preproc") (Attrib.add_del add_attribute del_attribute)
+    ("declaration of definition for preprocessing of the predicate compiler") #>
+  PureThy.add_thms_dynamic (Binding.name "pred_compile_preproc", Data.get);
+
+end;
+
 structure Code : CODE = struct open Code; end;
--- a/src/Pure/Isar/constdefs.ML	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/Pure/Isar/constdefs.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -52,7 +52,7 @@
       thy
       |> Sign.add_consts_i [(b, T, mx)]
       |> PureThy.add_defs false [((name, def), atts)]
-      |-> (fn [thm] => Code.add_default_eqn thm);
+      |-> (fn [thm] => Code.add_default_eqn thm #> Context.theory_map (Predicate_Compile_Preproc_Const_Defs.add_thm thm));
   in ((c, T), thy') end;
 
 fun gen_constdefs prep_vars prep_prop prep_att (raw_structs, specs) thy =
--- a/src/Pure/Isar/specification.ML	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/Pure/Isar/specification.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -209,7 +209,8 @@
         (var, ((Binding.suffix_name "_raw" name, []), rhs));
     val ((def_name, [th']), lthy3) = lthy2
       |> LocalTheory.note Thm.definitionK
-        ((name, Code.add_default_eqn_attrib :: atts), [prove lthy2 th]);
+        ((name, Predicate_Compile_Preproc_Const_Defs.add_attrib :: Code.add_default_eqn_attrib :: atts),
+          [prove lthy2 th]);
 
     val lhs' = Morphism.term (LocalTheory.target_morphism lthy3) lhs;
     val _ =
--- a/src/Pure/envir.ML	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/Pure/envir.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -282,12 +282,9 @@
   let
     val funerr = "fastype: expected function type";
     fun fast Ts (f $ u) =
-          (case fast Ts f of
+          (case Type.devar tyenv (fast Ts f) of
             Type ("fun", [_, T]) => T
-          | TVar v =>
-              (case Type.lookup tyenv v of
-                SOME (Type ("fun", [_, T])) => T
-              | _ => raise TERM (funerr, [f $ u]))
+          | TVar v => raise TERM (funerr, [f $ u])
           | _ => raise TERM (funerr, [f $ u]))
       | fast Ts (Const (_, T)) = T
       | fast Ts (Free (_, T)) = T
--- a/src/Pure/type.ML	Thu Sep 24 17:25:42 2009 +0200
+++ b/src/Pure/type.ML	Thu Sep 24 17:26:05 2009 +0200
@@ -55,6 +55,7 @@
   exception TYPE_MATCH
   type tyenv = (sort * typ) Vartab.table
   val lookup: tyenv -> indexname * sort -> typ option
+  val devar: tyenv -> typ -> typ
   val typ_match: tsig -> typ * typ -> tyenv -> tyenv
   val typ_instance: tsig -> typ * typ -> bool
   val raw_match: typ * typ -> tyenv -> tyenv