merged
authorchaieb
Tue, 19 May 2009 14:13:37 +0100
changeset 32411 63975b7f79b1
parent 31201 3dde56615750 (diff)
parent 32410 624bd2ea7c1e (current diff)
child 32413 b3241e8e9716
merged
src/HOL/Library/Formal_Power_Series.thy
--- a/doc-src/Codegen/Thy/ML.thy	Tue May 19 14:13:23 2009 +0100
+++ b/doc-src/Codegen/Thy/ML.thy	Tue May 19 14:13:37 2009 +0100
@@ -25,11 +25,11 @@
   @{index_ML Code.add_eqn: "thm -> theory -> theory"} \\
   @{index_ML Code.del_eqn: "thm -> theory -> theory"} \\
   @{index_ML Code.add_eqnl: "string * (thm * bool) list lazy -> theory -> theory"} \\
-  @{index_ML Code.map_pre: "(simpset -> simpset) -> theory -> theory"} \\
-  @{index_ML Code.map_post: "(simpset -> simpset) -> theory -> theory"} \\
-  @{index_ML Code.add_functrans: "string * (theory -> (thm * bool) list -> (thm * bool) list option)
+  @{index_ML Code_Preproc.map_pre: "(simpset -> simpset) -> theory -> theory"} \\
+  @{index_ML Code_Preproc.map_post: "(simpset -> simpset) -> theory -> theory"} \\
+  @{index_ML Code_Preproc.add_functrans: "string * (theory -> (thm * bool) list -> (thm * bool) list option)
     -> theory -> theory"} \\
-  @{index_ML Code.del_functrans: "string -> theory -> theory"} \\
+  @{index_ML Code_Preproc.del_functrans: "string -> theory -> theory"} \\
   @{index_ML Code.add_datatype: "(string * typ) list -> theory -> theory"} \\
   @{index_ML Code.get_datatype: "theory -> string
     -> (string * sort) list * (string * typ list) list"} \\
@@ -48,10 +48,10 @@
      suspended code equations @{text lthms} for constant
      @{text const} to executable content.
 
-  \item @{ML Code.map_pre}~@{text "f"}~@{text "thy"} changes
+  \item @{ML Code_Preproc.map_pre}~@{text "f"}~@{text "thy"} changes
      the preprocessor simpset.
 
-  \item @{ML Code.add_functrans}~@{text "(name, f)"}~@{text "thy"} adds
+  \item @{ML Code_Preproc.add_functrans}~@{text "(name, f)"}~@{text "thy"} adds
      function transformer @{text f} (named @{text name}) to executable content;
      @{text f} is a transformer of the code equations belonging
      to a certain function definition, depending on the
@@ -59,7 +59,7 @@
      transformation took place;  otherwise, the whole process will be iterated
      with the new code equations.
 
-  \item @{ML Code.del_functrans}~@{text "name"}~@{text "thy"} removes
+  \item @{ML Code_Preproc.del_functrans}~@{text "name"}~@{text "thy"} removes
      function transformer named @{text name} from executable content.
 
   \item @{ML Code.add_datatype}~@{text cs}~@{text thy} adds
@@ -78,20 +78,16 @@
 
 text %mlref {*
   \begin{mldecls}
-  @{index_ML Code_Unit.read_const: "theory -> string -> string"} \\
-  @{index_ML Code_Unit.head_eqn: "theory -> thm -> string * ((string * sort) list * typ)"} \\
-  @{index_ML Code_Unit.rewrite_eqn: "simpset -> thm -> thm"} \\
+  @{index_ML Code.read_const: "theory -> string -> string"} \\
+  @{index_ML Code.rewrite_eqn: "simpset -> thm -> thm"} \\
   \end{mldecls}
 
   \begin{description}
 
-  \item @{ML Code_Unit.read_const}~@{text thy}~@{text s}
+  \item @{ML Code.read_const}~@{text thy}~@{text s}
      reads a constant as a concrete term expression @{text s}.
 
-  \item @{ML Code_Unit.head_eqn}~@{text thy}~@{text thm}
-     extracts the constant and its type from a code equation @{text thm}.
-
-  \item @{ML Code_Unit.rewrite_eqn}~@{text ss}~@{text thm}
+  \item @{ML Code.rewrite_eqn}~@{text ss}~@{text thm}
      rewrites a code equation @{text thm} with a simpset @{text ss};
      only arguments and right hand side are rewritten,
      not the head of the code equation.
--- a/doc-src/Codegen/Thy/document/Adaptation.tex	Tue May 19 14:13:23 2009 +0100
+++ b/doc-src/Codegen/Thy/document/Adaptation.tex	Tue May 19 14:13:37 2009 +0100
@@ -166,7 +166,7 @@
        \isa{index} which is mapped to target-language built-in integers.
        Useful for code setups which involve e.g. indexing of
        target-language arrays.
-    \item[\hyperlink{theory.Code-Message}{\mbox{\isa{Code{\isacharunderscore}Message}}}] provides an additional datatype
+    \item[\hyperlink{theory.String}{\mbox{\isa{String}}}] provides an additional datatype
        \isa{message{\isacharunderscore}string} which is isomorphic to strings;
        \isa{message{\isacharunderscore}string}s are mapped to target-language strings.
        Useful for code setups which involve e.g. printing (error) messages.
--- a/doc-src/Codegen/Thy/document/Codegen.tex	Tue May 19 14:13:23 2009 +0100
+++ b/doc-src/Codegen/Thy/document/Codegen.tex	Tue May 19 14:13:37 2009 +0100
@@ -1550,20 +1550,20 @@
 %
 \begin{isamarkuptext}%
 \begin{mldecls}
-  \indexml{Code\_Unit.read\_const}\verb|Code_Unit.read_const: theory -> string -> string| \\
-  \indexml{Code\_Unit.head\_func}\verb|Code_Unit.head_func: thm -> string * ((string * sort) list * typ)| \\
-  \indexml{Code\_Unit.rewrite\_func}\verb|Code_Unit.rewrite_func: MetaSimplifier.simpset -> thm -> thm| \\
+  \indexml{Code\_Unit.read\_const}\verb|Code.read_const: theory -> string -> string| \\
+  \indexml{Code\_Unit.head\_func}\verb|Code.head_func: thm -> string * ((string * sort) list * typ)| \\
+  \indexml{Code\_Unit.rewrite\_func}\verb|Code.rewrite_func: MetaSimplifier.simpset -> thm -> thm| \\
   \end{mldecls}
 
   \begin{description}
 
-  \item \verb|Code_Unit.read_const|~\isa{thy}~\isa{s}
+  \item \verb|Code.read_const|~\isa{thy}~\isa{s}
      reads a constant as a concrete term expression \isa{s}.
 
-  \item \verb|Code_Unit.head_func|~\isa{thm}
+  \item \verb|Code.head_func|~\isa{thm}
      extracts the constant and its type from a defining equation \isa{thm}.
 
-  \item \verb|Code_Unit.rewrite_func|~\isa{ss}~\isa{thm}
+  \item \verb|Code.rewrite_func|~\isa{ss}~\isa{thm}
      rewrites a defining equation \isa{thm} with a simpset \isa{ss};
      only arguments and right hand side are rewritten,
      not the head of the defining equation.
--- a/doc-src/Codegen/Thy/document/ML.tex	Tue May 19 14:13:23 2009 +0100
+++ b/doc-src/Codegen/Thy/document/ML.tex	Tue May 19 14:13:37 2009 +0100
@@ -55,11 +55,11 @@
   \indexdef{}{ML}{Code.add\_eqn}\verb|Code.add_eqn: thm -> theory -> theory| \\
   \indexdef{}{ML}{Code.del\_eqn}\verb|Code.del_eqn: thm -> theory -> theory| \\
   \indexdef{}{ML}{Code.add\_eqnl}\verb|Code.add_eqnl: string * (thm * bool) list lazy -> theory -> theory| \\
-  \indexdef{}{ML}{Code.map\_pre}\verb|Code.map_pre: (simpset -> simpset) -> theory -> theory| \\
-  \indexdef{}{ML}{Code.map\_post}\verb|Code.map_post: (simpset -> simpset) -> theory -> theory| \\
-  \indexdef{}{ML}{Code.add\_functrans}\verb|Code.add_functrans: string * (theory -> (thm * bool) list -> (thm * bool) list option)|\isasep\isanewline%
+  \indexdef{}{ML}{Code\_Preproc.map\_pre}\verb|Code_Preproc.map_pre: (simpset -> simpset) -> theory -> theory| \\
+  \indexdef{}{ML}{Code\_Preproc.map\_post}\verb|Code_Preproc.map_post: (simpset -> simpset) -> theory -> theory| \\
+  \indexdef{}{ML}{Code\_Preproc.add\_functrans}\verb|Code_Preproc.add_functrans: string * (theory -> (thm * bool) list -> (thm * bool) list option)|\isasep\isanewline%
 \verb|    -> theory -> theory| \\
-  \indexdef{}{ML}{Code.del\_functrans}\verb|Code.del_functrans: string -> theory -> theory| \\
+  \indexdef{}{ML}{Code\_Preproc.del\_functrans}\verb|Code_Preproc.del_functrans: string -> theory -> theory| \\
   \indexdef{}{ML}{Code.add\_datatype}\verb|Code.add_datatype: (string * typ) list -> theory -> theory| \\
   \indexdef{}{ML}{Code.get\_datatype}\verb|Code.get_datatype: theory -> string|\isasep\isanewline%
 \verb|    -> (string * sort) list * (string * typ list) list| \\
@@ -78,10 +78,10 @@
      suspended code equations \isa{lthms} for constant
      \isa{const} to executable content.
 
-  \item \verb|Code.map_pre|~\isa{f}~\isa{thy} changes
+  \item \verb|Code_Preproc.map_pre|~\isa{f}~\isa{thy} changes
      the preprocessor simpset.
 
-  \item \verb|Code.add_functrans|~\isa{{\isacharparenleft}name{\isacharcomma}\ f{\isacharparenright}}~\isa{thy} adds
+  \item \verb|Code_Preproc.add_functrans|~\isa{{\isacharparenleft}name{\isacharcomma}\ f{\isacharparenright}}~\isa{thy} adds
      function transformer \isa{f} (named \isa{name}) to executable content;
      \isa{f} is a transformer of the code equations belonging
      to a certain function definition, depending on the
@@ -89,7 +89,7 @@
      transformation took place;  otherwise, the whole process will be iterated
      with the new code equations.
 
-  \item \verb|Code.del_functrans|~\isa{name}~\isa{thy} removes
+  \item \verb|Code_Preproc.del_functrans|~\isa{name}~\isa{thy} removes
      function transformer named \isa{name} from executable content.
 
   \item \verb|Code.add_datatype|~\isa{cs}~\isa{thy} adds
@@ -124,20 +124,16 @@
 %
 \begin{isamarkuptext}%
 \begin{mldecls}
-  \indexdef{}{ML}{Code\_Unit.read\_const}\verb|Code_Unit.read_const: theory -> string -> string| \\
-  \indexdef{}{ML}{Code\_Unit.head\_eqn}\verb|Code_Unit.head_eqn: theory -> thm -> string * ((string * sort) list * typ)| \\
-  \indexdef{}{ML}{Code\_Unit.rewrite\_eqn}\verb|Code_Unit.rewrite_eqn: simpset -> thm -> thm| \\
+  \indexdef{}{ML}{Code\_Unit.read\_const}\verb|Code.read_const: theory -> string -> string| \\
+  \indexdef{}{ML}{Code\_Unit.rewrite\_eqn}\verb|Code.rewrite_eqn: simpset -> thm -> thm| \\
   \end{mldecls}
 
   \begin{description}
 
-  \item \verb|Code_Unit.read_const|~\isa{thy}~\isa{s}
+  \item \verb|Code.read_const|~\isa{thy}~\isa{s}
      reads a constant as a concrete term expression \isa{s}.
 
-  \item \verb|Code_Unit.head_eqn|~\isa{thy}~\isa{thm}
-     extracts the constant and its type from a code equation \isa{thm}.
-
-  \item \verb|Code_Unit.rewrite_eqn|~\isa{ss}~\isa{thm}
+  \item \verb|Code.rewrite_eqn|~\isa{ss}~\isa{thm}
      rewrites a code equation \isa{thm} with a simpset \isa{ss};
      only arguments and right hand side are rewritten,
      not the head of the code equation.
--- a/doc-src/Codegen/Thy/document/Program.tex	Tue May 19 14:13:23 2009 +0100
+++ b/doc-src/Codegen/Thy/document/Program.tex	Tue May 19 14:13:37 2009 +0100
@@ -891,8 +891,8 @@
 \hspace*{0pt}fun null [] = true\\
 \hspace*{0pt} ~| null (x ::~xs) = false;\\
 \hspace*{0pt}\\
-\hspace*{0pt}fun eq{\char95}nat (Suc a) Zero{\char95}nat = false\\
-\hspace*{0pt} ~| eq{\char95}nat Zero{\char95}nat (Suc a) = false\\
+\hspace*{0pt}fun eq{\char95}nat (Suc nat') Zero{\char95}nat = false\\
+\hspace*{0pt} ~| eq{\char95}nat Zero{\char95}nat (Suc nat') = false\\
 \hspace*{0pt} ~| eq{\char95}nat (Suc nat) (Suc nat') = eq{\char95}nat nat nat'\\
 \hspace*{0pt} ~| eq{\char95}nat Zero{\char95}nat Zero{\char95}nat = true;\\
 \hspace*{0pt}\\
--- a/doc-src/more_antiquote.ML	Tue May 19 14:13:23 2009 +0100
+++ b/doc-src/more_antiquote.ML	Tue May 19 14:13:37 2009 +0100
@@ -87,10 +87,10 @@
 fun pretty_code_thm src ctxt raw_const =
   let
     val thy = ProofContext.theory_of ctxt;
-    val const = Code_Unit.check_const thy raw_const;
-    val (_, funcgr) = Code_Wellsorted.obtain thy [const] [];
+    val const = Code.check_const thy raw_const;
+    val (_, funcgr) = Code_Preproc.obtain thy [const] [];
     fun holize thm = @{thm meta_eq_to_obj_eq} OF [thm];
-    val thms = Code_Wellsorted.eqns funcgr const
+    val thms = Code_Preproc.eqns funcgr const
       |> map_filter (fn (thm, linear) => if linear then SOME thm else NONE)
       |> map (holize o no_vars ctxt o AxClass.overload thy);
   in ThyOutput.output (ThyOutput.maybe_pretty_source (pretty_thm ctxt) src thms) end;
@@ -108,7 +108,7 @@
 local
 
   val parse_const_terms = Scan.repeat1 Args.term
-    >> (fn ts => fn thy => map (Code_Unit.check_const thy) ts);
+    >> (fn ts => fn thy => map (Code.check_const thy) ts);
   val parse_consts = Scan.lift (Args.parens (Args.$$$ "consts")) |-- parse_const_terms
     >> (fn mk_cs => fn thy => fn naming => map_filter (Code_Thingol.lookup_const naming) (mk_cs thy));
   val parse_types = Scan.lift (Args.parens (Args.$$$ "types") |-- Scan.repeat1 Args.name)
--- a/src/HOL/Bali/Example.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Bali/Example.thy	Tue May 19 14:13:37 2009 +0100
@@ -1075,10 +1075,7 @@
 lemma max_spec_Base_foo: "max_spec tprg S (ClassT Base) \<lparr>name=foo,parTs=[NT]\<rparr> = 
   {((ClassT Base, \<lparr>access=Public,static=False,pars=[z],resT=Class Base\<rparr>)
    , [Class Base])}"
-apply (unfold max_spec_def)
-apply (simp (no_asm) add: appl_methds_Base_foo)
-apply auto
-done
+by (simp add: max_spec_def appl_methds_Base_foo Collect_conv_if)
 
 section "well-typedness"
 
--- a/src/HOL/Code_Eval.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Code_Eval.thy	Tue May 19 14:13:37 2009 +0100
@@ -28,30 +28,23 @@
 lemma term_of_anything: "term_of x \<equiv> t"
   by (rule eq_reflection) (cases "term_of x", cases t, simp)
 
-ML {*
-structure Eval =
-struct
+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 ()))"
 
-fun mk_term f g (Const (c, ty)) =
-      @{term Const} $ HOLogic.mk_message_string c $ g ty
-  | mk_term f g (t1 $ t2) =
-      @{term App} $ mk_term f g t1 $ mk_term f g t2
-  | mk_term f g (Free v) = f v
-  | mk_term f g (Bound i) = Bound i
-  | mk_term f g (Abs (v, _, t)) = Abs (v, @{typ term}, mk_term f g t);
-
-fun mk_term_of ty t = Const (@{const_name term_of}, ty --> @{typ term}) $ t;
-
-end;
-*}
+lemma valapp_code [code, code inline]:
+  "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 *}
 
 setup {*
 let
-  fun add_term_of_def ty vs tyco thy =
+  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"};
@@ -64,66 +57,57 @@
       |> `(fn lthy => Syntax.check_term lthy eq)
       |-> (fn eq => Specification.definition (NONE, ((Binding.name (triv_name_of eq), []), eq)))
       |> snd
-      |> Class.prove_instantiation_instance (K (Class.intro_classes_tac []))
-      |> LocalTheory.exit_global
+      |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
     end;
-  fun interpretator ("prop", (raw_vs, _)) thy = thy
-    | interpretator (tyco, (raw_vs, _)) thy =
-        let
-          val has_inst = can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort term_of};
-          val constrain_sort =
-            curry (Sorts.inter_sort (Sign.classes_of thy)) @{sort term_of};
-          val vs = (map o apsnd) constrain_sort raw_vs;
-          val ty = Type (tyco, map TFree vs);
-        in
-          thy
-          |> Typerep.perhaps_add_def tyco
-          |> not has_inst ? add_term_of_def ty vs tyco
-        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
+      thy
+      |> need_inst ? add_term_of tyco raw_vs
+    end;
 in
-  Code.type_interpretation interpretator
+  Code.type_interpretation ensure_term_of
 end
 *}
 
 setup {*
 let
-  fun mk_term_of_eq ty vs tyco (c, tys) =
+  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));
-    in (map_aterms (fn Free (v, ty) => Var ((v, 0), ty) | t => t) t, Eval.mk_term
-      (fn (v, ty) => Eval.mk_term_of ty (Var ((v, 0), ty)))
-      (Typerep.mk (fn (v, sort) => Typerep.typerep (TFree (v, sort)))) t)
+      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 prove_term_of_eq ty eq thy =
+  fun add_term_of_code tyco raw_vs raw_cs thy =
     let
-      val cty = Thm.ctyp_of thy ty;
-      val (arg, rhs) = pairself (Thm.cterm_of thy) eq;
-      val thm = @{thm term_of_anything}
-        |> Drule.instantiate' [SOME cty] [SOME arg, SOME rhs]
-        |> Thm.varifyT;
+      val vs = map (fn (v, _) => (v, @{sort typerep})) 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
       thy
-      |> Code.add_eqn thm
+      |> has_inst ? add_term_of_code tyco raw_vs cs
     end;
-  fun interpretator ("prop", (raw_vs, _)) thy = thy
-    | interpretator (tyco, (raw_vs, raw_cs)) thy =
-        let
-          val constrain_sort =
-            curry (Sorts.inter_sort (Sign.classes_of thy)) @{sort term_of};
-          val vs = (map o apsnd) constrain_sort raw_vs;
-          val cs = (map o apsnd o map o map_atyps)
-            (fn TFree (v, sort) => TFree (v, constrain_sort sort)) raw_cs;
-          val ty = Type (tyco, map TFree vs);
-          val eqs = map (mk_term_of_eq ty vs tyco) cs;
-          val const = AxClass.param_of_inst thy (@{const_name term_of}, tyco);
-        in
-          thy
-          |> Code.del_eqns const
-          |> fold (prove_term_of_eq ty) eqs
-        end;
 in
-  Code.type_interpretation interpretator
+  Code.type_interpretation ensure_term_of_code
 end
 *}
 
@@ -159,38 +143,80 @@
 code_reserved Eval HOLogic
 
 
-subsection {* Evaluation setup *}
+subsubsection {* Syntax *}
 
-ML {*
-signature EVAL =
-sig
-  val mk_term: ((string * typ) -> term) -> (typ -> term) -> term -> term
-  val mk_term_of: typ -> term -> term
-  val eval_ref: (unit -> term) option ref
-  val eval_term: theory -> term -> term
-end;
+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)"
 
-structure Eval : EVAL =
-struct
-
-open Eval;
-
-val eval_ref = ref (NONE : (unit -> term) option);
-
-fun eval_term thy t =
-  t 
-  |> Eval.mk_term_of (fastype_of t)
-  |> (fn t => Code_ML.eval NONE ("Eval.eval_ref", eval_ref) I thy t []);
-
+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;
 *}
 
-setup {*
-  Value.add_evaluator ("code", Eval.eval_term o ProofContext.theory_of)
-*}
+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)
 
 
-subsubsection {* Syntax *}
+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)
+
+
+subsection {* Obfuscate *}
 
 print_translation {*
 let
@@ -204,8 +230,32 @@
 end
 *}
 
-hide const dummy_term
-hide (open) const Const App
-hide (open) const term_of
+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_Index.thy	Tue May 19 14:13:37 2009 +0100
@@ -0,0 +1,330 @@
+(* Author: Florian Haftmann, TU Muenchen *)
+
+header {* Type of indices *}
+
+theory Code_Index
+imports Main
+begin
+
+text {*
+  Indices are isomorphic to HOL @{typ nat} but
+  mapped to target-language builtin integers.
+*}
+
+subsection {* Datatype of indices *}
+
+typedef (open) index = "UNIV \<Colon> nat set"
+  morphisms nat_of of_nat by rule
+
+lemma of_nat_nat_of [simp]:
+  "of_nat (nat_of k) = k"
+  by (rule nat_of_inverse)
+
+lemma nat_of_of_nat [simp]:
+  "nat_of (of_nat n) = n"
+  by (rule of_nat_inverse) (rule UNIV_I)
+
+lemma [measure_function]:
+  "is_measure nat_of" by (rule is_measure_trivial)
+
+lemma index:
+  "(\<And>n\<Colon>index. PROP P n) \<equiv> (\<And>n\<Colon>nat. PROP P (of_nat n))"
+proof
+  fix n :: nat
+  assume "\<And>n\<Colon>index. PROP P n"
+  then show "PROP P (of_nat n)" .
+next
+  fix n :: index
+  assume "\<And>n\<Colon>nat. PROP P (of_nat n)"
+  then have "PROP P (of_nat (nat_of n))" .
+  then show "PROP P n" by simp
+qed
+
+lemma index_case:
+  assumes "\<And>n. k = of_nat n \<Longrightarrow> P"
+  shows P
+  by (rule assms [of "nat_of k"]) simp
+
+lemma index_induct_raw:
+  assumes "\<And>n. P (of_nat n)"
+  shows "P k"
+proof -
+  from assms have "P (of_nat (nat_of k))" .
+  then show ?thesis by simp
+qed
+
+lemma nat_of_inject [simp]:
+  "nat_of k = nat_of l \<longleftrightarrow> k = l"
+  by (rule nat_of_inject)
+
+lemma of_nat_inject [simp]:
+  "of_nat n = of_nat m \<longleftrightarrow> n = m"
+  by (rule of_nat_inject) (rule UNIV_I)+
+
+instantiation index :: zero
+begin
+
+definition [simp, code del]:
+  "0 = of_nat 0"
+
+instance ..
+
+end
+
+definition [simp]:
+  "Suc_index k = of_nat (Suc (nat_of k))"
+
+rep_datatype "0 \<Colon> index" Suc_index
+proof -
+  fix P :: "index \<Rightarrow> bool"
+  fix k :: index
+  assume "P 0" then have init: "P (of_nat 0)" by simp
+  assume "\<And>k. P k \<Longrightarrow> P (Suc_index k)"
+    then have "\<And>n. P (of_nat n) \<Longrightarrow> P (Suc_index (of_nat n))" .
+    then have step: "\<And>n. P (of_nat n) \<Longrightarrow> P (of_nat (Suc n))" by simp
+  from init step have "P (of_nat (nat_of k))"
+    by (induct "nat_of k") simp_all
+  then show "P k" by simp
+qed simp_all
+
+declare index_case [case_names nat, cases type: index]
+declare index.induct [case_names nat, induct type: index]
+
+lemma index_decr [termination_simp]:
+  "k \<noteq> Code_Index.of_nat 0 \<Longrightarrow> Code_Index.nat_of k - Suc 0 < Code_Index.nat_of k"
+  by (cases k) simp
+
+lemma [simp, code]:
+  "index_size = nat_of"
+proof (rule ext)
+  fix k
+  have "index_size k = nat_size (nat_of k)"
+    by (induct k rule: index.induct) (simp_all del: zero_index_def Suc_index_def, simp_all)
+  also have "nat_size (nat_of k) = nat_of k" by (induct "nat_of k") simp_all
+  finally show "index_size k = nat_of k" .
+qed
+
+lemma [simp, code]:
+  "size = nat_of"
+proof (rule ext)
+  fix k
+  show "size k = nat_of k"
+  by (induct k) (simp_all del: zero_index_def Suc_index_def, simp_all)
+qed
+
+lemmas [code del] = index.recs index.cases
+
+lemma [code]:
+  "eq_class.eq k l \<longleftrightarrow> eq_class.eq (nat_of k) (nat_of l)"
+  by (cases k, cases l) (simp add: eq)
+
+lemma [code nbe]:
+  "eq_class.eq (k::index) k \<longleftrightarrow> True"
+  by (rule HOL.eq_refl)
+
+
+subsection {* Indices as datatype of ints *}
+
+instantiation index :: number
+begin
+
+definition
+  "number_of = of_nat o nat"
+
+instance ..
+
+end
+
+lemma nat_of_number [simp]:
+  "nat_of (number_of k) = number_of k"
+  by (simp add: number_of_index_def nat_number_of_def number_of_is_id)
+
+code_datatype "number_of \<Colon> int \<Rightarrow> index"
+
+
+subsection {* Basic arithmetic *}
+
+instantiation index :: "{minus, ordered_semidom, semiring_div, linorder}"
+begin
+
+definition [simp, code del]:
+  "(1\<Colon>index) = of_nat 1"
+
+definition [simp, code del]:
+  "n + m = of_nat (nat_of n + nat_of m)"
+
+definition [simp, code del]:
+  "n - m = of_nat (nat_of n - nat_of m)"
+
+definition [simp, code del]:
+  "n * m = of_nat (nat_of n * nat_of m)"
+
+definition [simp, code del]:
+  "n div m = of_nat (nat_of n div nat_of m)"
+
+definition [simp, code del]:
+  "n mod m = of_nat (nat_of n mod nat_of m)"
+
+definition [simp, code del]:
+  "n \<le> m \<longleftrightarrow> nat_of n \<le> nat_of m"
+
+definition [simp, code del]:
+  "n < m \<longleftrightarrow> nat_of n < nat_of m"
+
+instance proof
+qed (auto simp add: index left_distrib div_mult_self1)
+
+end
+
+lemma zero_index_code [code inline, code]:
+  "(0\<Colon>index) = Numeral0"
+  by (simp add: number_of_index_def Pls_def)
+lemma [code post]: "Numeral0 = (0\<Colon>index)"
+  using zero_index_code ..
+
+lemma one_index_code [code inline, code]:
+  "(1\<Colon>index) = Numeral1"
+  by (simp add: number_of_index_def Pls_def Bit1_def)
+lemma [code post]: "Numeral1 = (1\<Colon>index)"
+  using one_index_code ..
+
+lemma plus_index_code [code nbe]:
+  "of_nat n + of_nat m = of_nat (n + m)"
+  by simp
+
+definition subtract_index :: "index \<Rightarrow> index \<Rightarrow> index" where
+  [simp, code del]: "subtract_index = op -"
+
+lemma subtract_index_code [code nbe]:
+  "subtract_index (of_nat n) (of_nat m) = of_nat (n - m)"
+  by simp
+
+lemma minus_index_code [code]:
+  "n - m = subtract_index n m"
+  by simp
+
+lemma times_index_code [code nbe]:
+  "of_nat n * of_nat m = of_nat (n * m)"
+  by simp
+
+lemma less_eq_index_code [code nbe]:
+  "of_nat n \<le> of_nat m \<longleftrightarrow> n \<le> m"
+  by simp
+
+lemma less_index_code [code nbe]:
+  "of_nat n < of_nat m \<longleftrightarrow> n < m"
+  by simp
+
+lemma Suc_index_minus_one: "Suc_index n - 1 = n" by simp
+
+lemma of_nat_code [code]:
+  "of_nat = Nat.of_nat"
+proof
+  fix n :: nat
+  have "Nat.of_nat n = of_nat n"
+    by (induct n) simp_all
+  then show "of_nat n = Nat.of_nat n"
+    by (rule sym)
+qed
+
+lemma index_not_eq_zero: "i \<noteq> of_nat 0 \<longleftrightarrow> i \<ge> 1"
+  by (cases i) auto
+
+definition nat_of_aux :: "index \<Rightarrow> nat \<Rightarrow> nat" where
+  "nat_of_aux i n = nat_of i + n"
+
+lemma nat_of_aux_code [code]:
+  "nat_of_aux i n = (if i = 0 then n else nat_of_aux (i - 1) (Suc n))"
+  by (auto simp add: nat_of_aux_def index_not_eq_zero)
+
+lemma nat_of_code [code]:
+  "nat_of i = nat_of_aux i 0"
+  by (simp add: nat_of_aux_def)
+
+definition div_mod_index ::  "index \<Rightarrow> index \<Rightarrow> index \<times> index" where
+  [code del]: "div_mod_index n m = (n div m, n mod m)"
+
+lemma [code]:
+  "div_mod_index n m = (if m = 0 then (0, n) else (n div m, n mod m))"
+  unfolding div_mod_index_def by auto
+
+lemma [code]:
+  "n div m = fst (div_mod_index n m)"
+  unfolding div_mod_index_def by simp
+
+lemma [code]:
+  "n mod m = snd (div_mod_index n m)"
+  unfolding div_mod_index_def by simp
+
+definition int_of :: "index \<Rightarrow> int" where
+  "int_of = Nat.of_nat o nat_of"
+
+lemma int_of_code [code]:
+  "int_of k = (if k = 0 then 0
+    else (if k mod 2 = 0 then 2 * int_of (k div 2) else 2 * int_of (k div 2) + 1))"
+  by (auto simp add: int_of_def mod_div_equality')
+
+lemma (in term_syntax) term_of_index_code [code]:
+  "Code_Eval.term_of k =
+    Code_Eval.termify (number_of :: int \<Rightarrow> int) <\<cdot>> Code_Eval.term_of_num (2::index) k"
+  by (simp only: term_of_anything)
+
+hide (open) const of_nat nat_of int_of
+
+
+subsection {* Code generator setup *}
+
+text {* Implementation of indices by bounded integers *}
+
+code_type index
+  (SML "int")
+  (OCaml "int")
+  (Haskell "Int")
+
+code_instance index :: eq
+  (Haskell -)
+
+setup {*
+  fold (Numeral.add_code @{const_name number_index_inst.number_of_index}
+    false false) ["SML", "OCaml", "Haskell"]
+*}
+
+code_reserved SML Int int
+code_reserved OCaml Pervasives int
+
+code_const "op + \<Colon> index \<Rightarrow> index \<Rightarrow> index"
+  (SML "Int.+/ ((_),/ (_))")
+  (OCaml "Pervasives.( + )")
+  (Haskell infixl 6 "+")
+
+code_const "subtract_index \<Colon> index \<Rightarrow> index \<Rightarrow> index"
+  (SML "Int.max/ (_/ -/ _,/ 0 : int)")
+  (OCaml "Pervasives.max/ (_/ -/ _)/ (0 : int) ")
+  (Haskell "max/ (_/ -/ _)/ (0 :: Int)")
+
+code_const "op * \<Colon> index \<Rightarrow> index \<Rightarrow> index"
+  (SML "Int.*/ ((_),/ (_))")
+  (OCaml "Pervasives.( * )")
+  (Haskell infixl 7 "*")
+
+code_const div_mod_index
+  (SML "(fn n => fn m =>/ if m = 0/ then (0, n) else/ (n div m, n mod m))")
+  (OCaml "(fun n -> fun m ->/ if m = 0/ then (0, n) else/ (n '/ m, n mod m))")
+  (Haskell "divMod")
+
+code_const "eq_class.eq \<Colon> index \<Rightarrow> index \<Rightarrow> bool"
+  (SML "!((_ : Int.int) = _)")
+  (OCaml "!((_ : int) = _)")
+  (Haskell infixl 4 "==")
+
+code_const "op \<le> \<Colon> index \<Rightarrow> index \<Rightarrow> bool"
+  (SML "Int.<=/ ((_),/ (_))")
+  (OCaml "!((_ : int) <= _)")
+  (Haskell infix 4 "<=")
+
+code_const "op < \<Colon> index \<Rightarrow> index \<Rightarrow> bool"
+  (SML "Int.</ ((_),/ (_))")
+  (OCaml "!((_ : int) < _)")
+  (Haskell infix 4 "<")
+
+end
--- a/src/HOL/Complex_Main.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Complex_Main.thy	Tue May 19 14:13:37 2009 +0100
@@ -9,6 +9,7 @@
   Ln
   Taylor
   Integration
+  Quickcheck
 begin
 
 end
--- a/src/HOL/Decision_Procs/Approximation.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Decision_Procs/Approximation.thy	Tue May 19 14:13:37 2009 +0100
@@ -460,7 +460,7 @@
 proof (cases "even n")
   case True
   obtain n' where "Suc n' = get_odd n" and "odd (Suc n')" using get_odd_ex by auto
-  hence "even n'" unfolding even_nat_Suc by auto
+  hence "even n'" unfolding even_Suc by auto
   have "arctan (real x) \<le> real (x * ub_arctan_horner prec (get_odd n) 1 (x * x))"
     unfolding `Suc n' = get_odd n`[symmetric] using arctan_0_1_bounds'[OF `0 \<le> real x` `real x \<le> 1` `even n'`] by auto
   moreover
@@ -470,7 +470,7 @@
 next
   case False hence "0 < n" by (rule odd_pos)
   from gr0_implies_Suc[OF this] obtain n' where "n = Suc n'" ..
-  from False[unfolded this even_nat_Suc]
+  from False[unfolded this even_Suc]
   have "even n'" and "even (Suc (Suc n'))" by auto
   have "get_odd n = Suc n'" unfolding get_odd_def if_P[OF False] using `n = Suc n'` .
 
--- a/src/HOL/Extraction/Higman.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Extraction/Higman.thy	Tue May 19 14:13:37 2009 +0100
@@ -349,9 +349,9 @@
 
 end
 
-function mk_word_aux :: "nat \<Rightarrow> seed \<Rightarrow> letter list \<times> seed" where
+function mk_word_aux :: "nat \<Rightarrow> Random.seed \<Rightarrow> letter list \<times> Random.seed" where
   "mk_word_aux k = (do
-     i \<leftarrow> range 10;
+     i \<leftarrow> Random.range 10;
      (if i > 7 \<and> k > 2 \<or> k > 1000 then return []
      else do
        let l = (if i mod 2 = 0 then A else B);
@@ -362,10 +362,10 @@
 by pat_completeness auto
 termination by (relation "measure ((op -) 1001)") auto
 
-definition mk_word :: "seed \<Rightarrow> letter list \<times> seed" where
+definition mk_word :: "Random.seed \<Rightarrow> letter list \<times> Random.seed" where
   "mk_word = mk_word_aux 0"
 
-primrec mk_word_s :: "nat \<Rightarrow> seed \<Rightarrow> letter list \<times> seed" where
+primrec mk_word_s :: "nat \<Rightarrow> Random.seed \<Rightarrow> letter list \<times> Random.seed" where
   "mk_word_s 0 = mk_word"
   | "mk_word_s (Suc n) = (do
        _ \<leftarrow> mk_word;
--- a/src/HOL/HOL.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/HOL.thy	Tue May 19 14:13:37 2009 +0100
@@ -1050,8 +1050,7 @@
     "(P | False) = P"  "(False | P) = P"
     "(P | P) = P"  "(P | (P | Q)) = (P | Q)" and
     "(ALL x. P) = P"  "(EX x. P) = P"  "EX x. x=t"  "EX x. t=x"
-    -- {* needed for the one-point-rule quantifier simplification procs *}
-    -- {* essential for termination!! *} and
+  and
     "!!P. (EX x. x=t & P(x)) = P(t)"
     "!!P. (EX x. t=x & P(x)) = P(t)"
     "!!P. (ALL x. x=t --> P(x)) = P(t)"
@@ -1331,7 +1330,7 @@
     of Abs (_, _, t') => count_loose t' 0 <= 1
      | _ => true;
 in fn _ => fn ss => fn ct => if is_trivial_let (Thm.term_of ct)
-  then SOME @{thm Let_def} (*no or one ocurrenc of bound variable*)
+  then SOME @{thm Let_def} (*no or one ocurrence of bound variable*)
   else let (*Norbert Schirmer's case*)
     val ctxt = Simplifier.the_context ss;
     val thy = ProofContext.theory_of ctxt;
@@ -1775,6 +1774,13 @@
 end
 *}
 
+subsubsection {* Generic code generator preprocessor setup *}
+
+setup {*
+  Code_Preproc.map_pre (K HOL_basic_ss)
+  #> Code_Preproc.map_post (K HOL_basic_ss)
+*}
+
 subsubsection {* Equality *}
 
 class eq =
@@ -1788,7 +1794,7 @@
 lemma eq_refl: "eq x x \<longleftrightarrow> True"
   unfolding eq by rule+
 
-lemma equals_eq [code inline]: "(op =) \<equiv> eq"
+lemma equals_eq: "(op =) \<equiv> eq"
   by (rule eq_reflection) (rule ext, rule ext, rule sym, rule eq_equals)
 
 declare equals_eq [symmetric, code post]
@@ -1797,6 +1803,14 @@
 
 declare equals_eq [code]
 
+setup {*
+  Code_Preproc.map_pre (fn simpset =>
+    simpset addsimprocs [Simplifier.simproc_i @{theory} "eq" [@{term "op ="}]
+      (fn thy => fn _ => fn t as Const (_, T) => case strip_type T
+        of ((T as Type _) :: _, _) => SOME @{thm equals_eq}
+         | _ => NONE)])
+*}
+
 
 subsubsection {* Generic code generator foundation *}
 
@@ -1836,17 +1850,32 @@
 
 lemmas [code, code unfold, symmetric, code post] = imp_conv_disj
 
+instantiation itself :: (type) eq
+begin
+
+definition eq_itself :: "'a itself \<Rightarrow> 'a itself \<Rightarrow> bool" where
+  "eq_itself x y \<longleftrightarrow> x = y"
+
+instance proof
+qed (fact eq_itself_def)
+
+end
+
+lemma eq_itself_code [code]:
+  "eq_class.eq TYPE('a) TYPE('a) \<longleftrightarrow> True"
+  by (simp add: eq)
+
 text {* Equality *}
 
 declare simp_thms(6) [code nbe]
 
+setup {*
+  Code.add_const_alias @{thm equals_eq}
+*}
+
 hide (open) const eq
 hide const eq
 
-setup {*
-  Code_Unit.add_const_alias @{thm equals_eq}
-*}
-
 text {* Cases *}
 
 lemma Let_case_cert:
@@ -1867,13 +1896,6 @@
 
 code_abort undefined
 
-subsubsection {* Generic code generator preprocessor *}
-
-setup {*
-  Code_Preproc.map_pre (K HOL_basic_ss)
-  #> Code_Preproc.map_post (K HOL_basic_ss)
-*}
-
 subsubsection {* Generic code generator target languages *}
 
 text {* type bool *}
@@ -1965,6 +1987,18 @@
 
 subsubsection {* Quickcheck *}
 
+ML {*
+structure Quickcheck_RecFun_Simp_Thms = NamedThmsFun
+(
+  val name = "quickcheck_recfun_simp"
+  val description = "simplification rules of recursive functions as needed by Quickcheck"
+)
+*}
+
+setup {*
+  Quickcheck_RecFun_Simp_Thms.setup
+*}
+
 setup {*
   Quickcheck.add_generator ("SML", Codegen.test_term)
 *}
--- a/src/HOL/IsaMakefile	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/IsaMakefile	Tue May 19 14:13:37 2009 +0100
@@ -206,6 +206,7 @@
 MAIN_DEPENDENCIES = $(PLAIN_DEPENDENCIES) \
   ATP_Linkup.thy \
   Code_Eval.thy \
+  Code_Index.thy \
   Equiv_Relations.thy \
   Groebner_Basis.thy \
   Hilbert_Choice.thy \
@@ -286,8 +287,10 @@
   Transcendental.thy \
   GCD.thy \
   Parity.thy \
+  Quickcheck.thy \
   Lubs.thy \
   PReal.thy \
+  Random.thy \
   Rational.thy \
   RComplete.thy \
   RealDef.thy \
@@ -333,12 +336,11 @@
   Library/comm_ring.ML Library/Coinductive_List.thy			\
   Library/AssocList.thy	Library/Formal_Power_Series.thy	\
   Library/Binomial.thy Library/Eval_Witness.thy				\
-  Library/Code_Index.thy Library/Code_Char.thy				\
+  Library/Code_Char.thy				\
   Library/Code_Char_chr.thy Library/Code_Integer.thy			\
   Library/Mapping.thy Library/Numeral_Type.thy Library/Reflection.thy	\
   Library/Boolean_Algebra.thy Library/Countable.thy			\
   Library/Diagonalize.thy Library/RBT.thy Library/Univ_Poly.thy		\
-  Library/Random.thy Library/Quickcheck.thy	\
   Library/Poly_Deriv.thy \
   Library/Polynomial.thy \
   Library/Preorder.thy \
@@ -851,7 +853,7 @@
   ex/Recdefs.thy ex/Records.thy ex/ReflectionEx.thy			\
   ex/Refute_Examples.thy ex/SAT_Examples.thy ex/SVC_Oracle.thy		\
   ex/Serbian.thy ex/Sqrt.thy ex/Sqrt_Script.thy \
-  ex/Sudoku.thy ex/Tarski.thy ex/Term_Of_Syntax.thy	\
+  ex/Sudoku.thy ex/Tarski.thy \
   ex/Termination.thy ex/Unification.thy ex/document/root.bib		\
   ex/document/root.tex ex/set.thy ex/svc_funcs.ML ex/svc_test.thy \
   ex/Predicate_Compile.thy ex/predicate_compile.ML ex/Predicate_Compile_ex.thy
--- a/src/HOL/Library/Binomial.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Library/Binomial.thy	Tue May 19 14:13:37 2009 +0100
@@ -88,9 +88,7 @@
 
 lemma card_s_0_eq_empty:
     "finite A ==> card {B. B \<subseteq> A & card B = 0} = 1"
-  apply (simp cong add: conj_cong add: finite_subset [THEN card_0_eq])
-  apply (simp cong add: rev_conj_cong)
-  done
+by (simp cong add: conj_cong add: finite_subset [THEN card_0_eq])
 
 lemma choose_deconstruct: "finite M ==> x \<notin> M
   ==> {s. s <= insert x M & card(s) = Suc k}
--- a/src/HOL/Library/Code_Index.thy	Tue May 19 14:13:23 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,336 +0,0 @@
-(* Author: Florian Haftmann, TU Muenchen *)
-
-header {* Type of indices *}
-
-theory Code_Index
-imports Main
-begin
-
-text {*
-  Indices are isomorphic to HOL @{typ nat} but
-  mapped to target-language builtin integers.
-*}
-
-subsection {* Datatype of indices *}
-
-typedef (open) index = "UNIV \<Colon> nat set"
-  morphisms nat_of of_nat by rule
-
-lemma of_nat_nat_of [simp]:
-  "of_nat (nat_of k) = k"
-  by (rule nat_of_inverse)
-
-lemma nat_of_of_nat [simp]:
-  "nat_of (of_nat n) = n"
-  by (rule of_nat_inverse) (rule UNIV_I)
-
-lemma [measure_function]:
-  "is_measure nat_of" by (rule is_measure_trivial)
-
-lemma index:
-  "(\<And>n\<Colon>index. PROP P n) \<equiv> (\<And>n\<Colon>nat. PROP P (of_nat n))"
-proof
-  fix n :: nat
-  assume "\<And>n\<Colon>index. PROP P n"
-  then show "PROP P (of_nat n)" .
-next
-  fix n :: index
-  assume "\<And>n\<Colon>nat. PROP P (of_nat n)"
-  then have "PROP P (of_nat (nat_of n))" .
-  then show "PROP P n" by simp
-qed
-
-lemma index_case:
-  assumes "\<And>n. k = of_nat n \<Longrightarrow> P"
-  shows P
-  by (rule assms [of "nat_of k"]) simp
-
-lemma index_induct_raw:
-  assumes "\<And>n. P (of_nat n)"
-  shows "P k"
-proof -
-  from assms have "P (of_nat (nat_of k))" .
-  then show ?thesis by simp
-qed
-
-lemma nat_of_inject [simp]:
-  "nat_of k = nat_of l \<longleftrightarrow> k = l"
-  by (rule nat_of_inject)
-
-lemma of_nat_inject [simp]:
-  "of_nat n = of_nat m \<longleftrightarrow> n = m"
-  by (rule of_nat_inject) (rule UNIV_I)+
-
-instantiation index :: zero
-begin
-
-definition [simp, code del]:
-  "0 = of_nat 0"
-
-instance ..
-
-end
-
-definition [simp]:
-  "Suc_index k = of_nat (Suc (nat_of k))"
-
-rep_datatype "0 \<Colon> index" Suc_index
-proof -
-  fix P :: "index \<Rightarrow> bool"
-  fix k :: index
-  assume "P 0" then have init: "P (of_nat 0)" by simp
-  assume "\<And>k. P k \<Longrightarrow> P (Suc_index k)"
-    then have "\<And>n. P (of_nat n) \<Longrightarrow> P (Suc_index (of_nat n))" .
-    then have step: "\<And>n. P (of_nat n) \<Longrightarrow> P (of_nat (Suc n))" by simp
-  from init step have "P (of_nat (nat_of k))"
-    by (induct "nat_of k") simp_all
-  then show "P k" by simp
-qed simp_all
-
-declare index_case [case_names nat, cases type: index]
-declare index.induct [case_names nat, induct type: index]
-
-lemma index_decr [termination_simp]:
-  "k \<noteq> Code_Index.of_nat 0 \<Longrightarrow> Code_Index.nat_of k - Suc 0 < Code_Index.nat_of k"
-  by (cases k) simp
-
-lemma [simp, code]:
-  "index_size = nat_of"
-proof (rule ext)
-  fix k
-  have "index_size k = nat_size (nat_of k)"
-    by (induct k rule: index.induct) (simp_all del: zero_index_def Suc_index_def, simp_all)
-  also have "nat_size (nat_of k) = nat_of k" by (induct "nat_of k") simp_all
-  finally show "index_size k = nat_of k" .
-qed
-
-lemma [simp, code]:
-  "size = nat_of"
-proof (rule ext)
-  fix k
-  show "size k = nat_of k"
-  by (induct k) (simp_all del: zero_index_def Suc_index_def, simp_all)
-qed
-
-lemmas [code del] = index.recs index.cases
-
-lemma [code]:
-  "eq_class.eq k l \<longleftrightarrow> eq_class.eq (nat_of k) (nat_of l)"
-  by (cases k, cases l) (simp add: eq)
-
-lemma [code nbe]:
-  "eq_class.eq (k::index) k \<longleftrightarrow> True"
-  by (rule HOL.eq_refl)
-
-
-subsection {* Indices as datatype of ints *}
-
-instantiation index :: number
-begin
-
-definition
-  "number_of = of_nat o nat"
-
-instance ..
-
-end
-
-lemma nat_of_number [simp]:
-  "nat_of (number_of k) = number_of k"
-  by (simp add: number_of_index_def nat_number_of_def number_of_is_id)
-
-code_datatype "number_of \<Colon> int \<Rightarrow> index"
-
-
-subsection {* Basic arithmetic *}
-
-instantiation index :: "{minus, ordered_semidom, semiring_div, linorder}"
-begin
-
-definition [simp, code del]:
-  "(1\<Colon>index) = of_nat 1"
-
-definition [simp, code del]:
-  "n + m = of_nat (nat_of n + nat_of m)"
-
-definition [simp, code del]:
-  "n - m = of_nat (nat_of n - nat_of m)"
-
-definition [simp, code del]:
-  "n * m = of_nat (nat_of n * nat_of m)"
-
-definition [simp, code del]:
-  "n div m = of_nat (nat_of n div nat_of m)"
-
-definition [simp, code del]:
-  "n mod m = of_nat (nat_of n mod nat_of m)"
-
-definition [simp, code del]:
-  "n \<le> m \<longleftrightarrow> nat_of n \<le> nat_of m"
-
-definition [simp, code del]:
-  "n < m \<longleftrightarrow> nat_of n < nat_of m"
-
-instance proof
-qed (auto simp add: index left_distrib div_mult_self1)
-
-end
-
-lemma zero_index_code [code inline, code]:
-  "(0\<Colon>index) = Numeral0"
-  by (simp add: number_of_index_def Pls_def)
-lemma [code post]: "Numeral0 = (0\<Colon>index)"
-  using zero_index_code ..
-
-lemma one_index_code [code inline, code]:
-  "(1\<Colon>index) = Numeral1"
-  by (simp add: number_of_index_def Pls_def Bit1_def)
-lemma [code post]: "Numeral1 = (1\<Colon>index)"
-  using one_index_code ..
-
-lemma plus_index_code [code nbe]:
-  "of_nat n + of_nat m = of_nat (n + m)"
-  by simp
-
-definition subtract_index :: "index \<Rightarrow> index \<Rightarrow> index" where
-  [simp, code del]: "subtract_index = op -"
-
-lemma subtract_index_code [code nbe]:
-  "subtract_index (of_nat n) (of_nat m) = of_nat (n - m)"
-  by simp
-
-lemma minus_index_code [code]:
-  "n - m = subtract_index n m"
-  by simp
-
-lemma times_index_code [code nbe]:
-  "of_nat n * of_nat m = of_nat (n * m)"
-  by simp
-
-lemma less_eq_index_code [code nbe]:
-  "of_nat n \<le> of_nat m \<longleftrightarrow> n \<le> m"
-  by simp
-
-lemma less_index_code [code nbe]:
-  "of_nat n < of_nat m \<longleftrightarrow> n < m"
-  by simp
-
-lemma Suc_index_minus_one: "Suc_index n - 1 = n" by simp
-
-lemma of_nat_code [code]:
-  "of_nat = Nat.of_nat"
-proof
-  fix n :: nat
-  have "Nat.of_nat n = of_nat n"
-    by (induct n) simp_all
-  then show "of_nat n = Nat.of_nat n"
-    by (rule sym)
-qed
-
-lemma index_not_eq_zero: "i \<noteq> of_nat 0 \<longleftrightarrow> i \<ge> 1"
-  by (cases i) auto
-
-definition nat_of_aux :: "index \<Rightarrow> nat \<Rightarrow> nat" where
-  "nat_of_aux i n = nat_of i + n"
-
-lemma nat_of_aux_code [code]:
-  "nat_of_aux i n = (if i = 0 then n else nat_of_aux (i - 1) (Suc n))"
-  by (auto simp add: nat_of_aux_def index_not_eq_zero)
-
-lemma nat_of_code [code]:
-  "nat_of i = nat_of_aux i 0"
-  by (simp add: nat_of_aux_def)
-
-definition div_mod_index ::  "index \<Rightarrow> index \<Rightarrow> index \<times> index" where
-  [code del]: "div_mod_index n m = (n div m, n mod m)"
-
-lemma [code]:
-  "div_mod_index n m = (if m = 0 then (0, n) else (n div m, n mod m))"
-  unfolding div_mod_index_def by auto
-
-lemma [code]:
-  "n div m = fst (div_mod_index n m)"
-  unfolding div_mod_index_def by simp
-
-lemma [code]:
-  "n mod m = snd (div_mod_index n m)"
-  unfolding div_mod_index_def by simp
-
-hide (open) const of_nat nat_of
-
-subsection {* ML interface *}
-
-ML {*
-structure Index =
-struct
-
-fun mk k = HOLogic.mk_number @{typ index} k;
-
-end;
-*}
-
-
-subsection {* Code generator setup *}
-
-text {* Implementation of indices by bounded integers *}
-
-code_type index
-  (SML "int")
-  (OCaml "int")
-  (Haskell "Int")
-
-code_instance index :: eq
-  (Haskell -)
-
-setup {*
-  fold (Numeral.add_code @{const_name number_index_inst.number_of_index}
-    false false) ["SML", "OCaml", "Haskell"]
-*}
-
-code_reserved SML Int int
-code_reserved OCaml Pervasives int
-
-code_const "op + \<Colon> index \<Rightarrow> index \<Rightarrow> index"
-  (SML "Int.+/ ((_),/ (_))")
-  (OCaml "Pervasives.( + )")
-  (Haskell infixl 6 "+")
-
-code_const "subtract_index \<Colon> index \<Rightarrow> index \<Rightarrow> index"
-  (SML "Int.max/ (_/ -/ _,/ 0 : int)")
-  (OCaml "Pervasives.max/ (_/ -/ _)/ (0 : int) ")
-  (Haskell "max/ (_/ -/ _)/ (0 :: Int)")
-
-code_const "op * \<Colon> index \<Rightarrow> index \<Rightarrow> index"
-  (SML "Int.*/ ((_),/ (_))")
-  (OCaml "Pervasives.( * )")
-  (Haskell infixl 7 "*")
-
-code_const div_mod_index
-  (SML "(fn n => fn m =>/ if m = 0/ then (0, n) else/ (n div m, n mod m))")
-  (OCaml "(fun n -> fun m ->/ if m = 0/ then (0, n) else/ (n '/ m, n mod m))")
-  (Haskell "divMod")
-
-code_const "eq_class.eq \<Colon> index \<Rightarrow> index \<Rightarrow> bool"
-  (SML "!((_ : Int.int) = _)")
-  (OCaml "!((_ : int) = _)")
-  (Haskell infixl 4 "==")
-
-code_const "op \<le> \<Colon> index \<Rightarrow> index \<Rightarrow> bool"
-  (SML "Int.<=/ ((_),/ (_))")
-  (OCaml "!((_ : int) <= _)")
-  (Haskell infix 4 "<=")
-
-code_const "op < \<Colon> index \<Rightarrow> index \<Rightarrow> bool"
-  (SML "Int.</ ((_),/ (_))")
-  (OCaml "!((_ : int) < _)")
-  (Haskell infix 4 "<")
-
-text {* Evaluation *}
-
-lemma [code, code del]:
-  "(Code_Eval.term_of \<Colon> index \<Rightarrow> term) = Code_Eval.term_of" ..
-
-code_const "Code_Eval.term_of \<Colon> index \<Rightarrow> term"
-  (SML "HOLogic.mk'_number/ HOLogic.indexT/ (IntInf.fromInt/ _)")
-
-end
--- a/src/HOL/Library/Code_Integer.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Library/Code_Integer.thy	Tue May 19 14:13:37 2009 +0100
@@ -5,7 +5,7 @@
 header {* Pretty integer literals for code generation *}
 
 theory Code_Integer
-imports Main
+imports Main Code_Index
 begin
 
 text {*
@@ -91,15 +91,17 @@
   (OCaml "Big'_int.lt'_big'_int")
   (Haskell infix 4 "<")
 
+code_const Code_Index.int_of
+  (SML "IntInf.fromInt")
+  (OCaml "Big'_int.big'_int'_of'_int")
+  (Haskell "toEnum")
+
 code_reserved SML IntInf
 code_reserved OCaml Big_int
 
 text {* Evaluation *}
 
-lemma [code, code del]:
-  "(Code_Eval.term_of \<Colon> int \<Rightarrow> term) = Code_Eval.term_of" ..
-
 code_const "Code_Eval.term_of \<Colon> int \<Rightarrow> term"
-  (SML "HOLogic.mk'_number/ HOLogic.intT")
+  (Eval "HOLogic.mk'_number/ HOLogic.intT")
 
 end
\ No newline at end of file
--- a/src/HOL/Library/Enum.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Library/Enum.thy	Tue May 19 14:13:37 2009 +0100
@@ -116,9 +116,8 @@
     by (simp add: length_n_lists)
 qed
 
-lemma map_of_zip_map:
-  fixes f :: "'a\<Colon>enum \<Rightarrow> 'b\<Colon>enum"
-  shows "map_of (zip xs (map f xs)) = (\<lambda>x. if x \<in> set xs then Some (f x) else None)"
+lemma map_of_zip_map: (*FIXME move to Map.thy*)
+  "map_of (zip xs (map f xs)) = (\<lambda>x. if x \<in> set xs then Some (f x) else None)"
   by (induct xs) (simp_all add: expand_fun_eq)
 
 lemma map_of_zip_enum_is_Some:
--- a/src/HOL/Library/Formal_Power_Series.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Library/Formal_Power_Series.thy	Tue May 19 14:13:37 2009 +0100
@@ -917,8 +917,7 @@
 proof-
   have eq: "(1 + X) * ?r = 1"
     unfolding minus_one_power_iff
-    apply (auto simp add: ring_simps fps_eq_iff)
-    by presburger+
+    by (auto simp add: ring_simps fps_eq_iff)
   show ?thesis by (auto simp add: eq intro: fps_inverse_unique)
 qed
 
@@ -2103,6 +2102,80 @@
   ultimately show ?thesis by (cases n, auto)
 qed
 
+lemma fps_compose_uminus: "- (a::'a::ring_1 fps) oo c = - (a oo c)"
+  by (simp add: fps_eq_iff fps_compose_nth ring_simps setsum_negf[symmetric])
+
+lemma fps_compose_sub_distrib:
+  shows "(a - b) oo (c::'a::ring_1 fps) = (a oo c) - (b oo c)"
+  unfolding diff_minus fps_compose_uminus fps_compose_add_distrib ..
+
+lemma X_fps_compose:"X oo a = Abs_fps (\<lambda>n. if n = 0 then (0::'a::comm_ring_1) else a$n)"
+  by (simp add: fps_eq_iff fps_compose_nth mult_delta_left setsum_delta power_Suc)
+
+lemma fps_inverse_compose:
+  assumes b0: "(b$0 :: 'a::field) = 0" and a0: "a$0 \<noteq> 0"
+  shows "inverse a oo b = inverse (a oo b)"
+proof-
+  let ?ia = "inverse a"
+  let ?ab = "a oo b"
+  let ?iab = "inverse ?ab"
+
+from a0 have ia0: "?ia $ 0 \<noteq> 0" by (simp )
+from a0 have ab0: "?ab $ 0 \<noteq> 0" by (simp add: fps_compose_def)
+thm inverse_mult_eq_1[OF ab0]
+have "(?ia oo b) *  (a oo b) = 1"
+unfolding fps_compose_mult_distrib[OF b0, symmetric]
+unfolding inverse_mult_eq_1[OF a0]
+fps_compose_1 ..
+
+then have "(?ia oo b) *  (a oo b) * ?iab  = 1 * ?iab" by simp
+then have "(?ia oo b) *  (?iab * (a oo b))  = ?iab" by simp
+then show ?thesis 
+  unfolding inverse_mult_eq_1[OF ab0] by simp
+qed
+
+lemma fps_divide_compose:
+  assumes c0: "(c$0 :: 'a::field) = 0" and b0: "b$0 \<noteq> 0"
+  shows "(a/b) oo c = (a oo c) / (b oo c)"
+    unfolding fps_divide_def fps_compose_mult_distrib[OF c0]
+    fps_inverse_compose[OF c0 b0] ..
+
+lemma gp: assumes a0: "a$0 = (0::'a::field)"
+  shows "(Abs_fps (\<lambda>n. 1)) oo a = 1/(1 - a)" (is "?one oo a = _")
+proof-
+  have o0: "?one $ 0 \<noteq> 0" by simp
+  have th0: "(1 - X) $ 0 \<noteq> (0::'a)" by simp  
+  from fps_inverse_gp[where ?'a = 'a]
+  have "inverse ?one = 1 - X" by (simp add: fps_eq_iff)
+  hence "inverse (inverse ?one) = inverse (1 - X)" by simp
+  hence th: "?one = 1/(1 - X)" unfolding fps_inverse_idempotent[OF o0] 
+    by (simp add: fps_divide_def)
+  show ?thesis unfolding th
+    unfolding fps_divide_compose[OF a0 th0]
+    fps_compose_1 fps_compose_sub_distrib X_fps_compose_startby0[OF a0] ..
+qed
+
+lemma fps_const_power[simp]: "fps_const (c::'a::ring_1) ^ n = fps_const (c^n)"
+by (induct n, auto)
+
+lemma fps_compose_radical:
+  assumes b0: "b$0 = (0::'a::{field, ring_char_0})"
+  and ra0: "r (Suc k) (a$0) ^ Suc k = a$0"
+  and a0: "a$0 \<noteq> 0"
+  shows "fps_radical r (Suc k)  a oo b = fps_radical r (Suc k) (a oo b)"
+proof-
+  let ?r = "fps_radical r (Suc k)"
+  let ?ab = "a oo b"
+  have ab0: "?ab $ 0 = a$0" by (simp add: fps_compose_def)
+  from ab0 a0 ra0 have rab0: "?ab $ 0 \<noteq> 0" "r (Suc k) (?ab $ 0) ^ Suc k = ?ab $ 0" by simp_all
+  have th00: "r (Suc k) ((a oo b) $ 0) = (fps_radical r (Suc k) a oo b) $ 0"
+    by (simp add: ab0 fps_compose_def)
+  have th0: "(?r a oo b) ^ (Suc k) = a  oo b"
+    unfolding fps_compose_power[OF b0]
+    unfolding iffD1[OF power_radical[of a r k], OF a0 ra0]  .. 
+  from iffD1[OF radical_unique[where r=r and k=k and b= ?ab and a = "?r a oo b", OF rab0(2) th00 rab0(1)], OF th0] show ?thesis  . 
+qed
+
 lemma fps_const_mult_apply_left:
   "fps_const c * (a oo b) = (fps_const c * a) oo b"
   by (simp add: fps_eq_iff fps_compose_nth setsum_right_distrib mult_assoc)
@@ -2273,15 +2346,6 @@
 lemma E_nth_deriv[simp]: "fps_nth_deriv n (E (a::'a::{field, ring_char_0})) = (fps_const a)^n * (E a)"
   by (induct n, auto simp add: power_Suc)
 
-lemma fps_compose_uminus: "- (a::'a::ring_1 fps) oo c = - (a oo c)"
-  by (simp add: fps_eq_iff fps_compose_nth ring_simps setsum_negf[symmetric])
-
-lemma fps_compose_sub_distrib:
-  shows "(a - b) oo (c::'a::ring_1 fps) = (a oo c) - (b oo c)"
-  unfolding diff_minus fps_compose_uminus fps_compose_add_distrib ..
-
-lemma X_fps_compose:"X oo a = Abs_fps (\<lambda>n. if n = 0 then (0::'a::comm_ring_1) else a$n)"
-  by (simp add: fps_eq_iff fps_compose_nth mult_delta_left setsum_delta power_Suc)
 
 lemma X_compose_E[simp]: "X oo E (a::'a::{field}) = E a - 1"
   by (simp add: fps_eq_iff X_fps_compose)
@@ -2309,9 +2373,7 @@
   (is "inverse ?l = ?r")
 proof-
   have th: "?l * ?r = 1"
-    apply (auto simp add: ring_simps fps_eq_iff X_mult_nth  minus_one_power_iff)
-    apply presburger+
-    done
+    by (auto simp add: ring_simps fps_eq_iff minus_one_power_iff)
   have th': "?l $ 0 \<noteq> 0" by (simp add: )
   from fps_inverse_unique[OF th' th] show ?thesis .
 qed
@@ -2327,6 +2389,7 @@
   unfolding inverse_one_plus_X
   by (simp add: L_def fps_eq_iff power_Suc del: of_nat_Suc)
 
+
 lemma L_nth: "L $ n = (- 1) ^ Suc n / of_nat n"
   by (simp add: L_def)
 
--- a/src/HOL/Library/Library.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Library/Library.thy	Tue May 19 14:13:37 2009 +0100
@@ -9,7 +9,6 @@
   Boolean_Algebra
   Char_ord
   Code_Char_chr
-  Code_Index
   Code_Integer
   Coinductive_List
   Commutative_Ring
@@ -45,11 +44,9 @@
   Preorder
   Primes
   Product_Vector
-  Quickcheck
   Quicksort
   Quotient
   Ramsey
-  Random
   Reflection
   RBT
   State_Monad
--- a/src/HOL/Library/Pocklington.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Library/Pocklington.thy	Tue May 19 14:13:37 2009 +0100
@@ -382,8 +382,9 @@
 (* Euler totient function.                                                   *)
 
 definition phi_def: "\<phi> n = card { m. 0 < m \<and> m <= n \<and> coprime m n }"
+
 lemma phi_0[simp]: "\<phi> 0 = 0"
-  unfolding phi_def by (auto simp add: card_eq_0_iff)
+  unfolding phi_def by auto
 
 lemma phi_finite[simp]: "finite ({ m. 0 < m \<and> m <= n \<and> coprime m n })"
 proof-
--- a/src/HOL/Library/Quickcheck.thy	Tue May 19 14:13:23 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,148 +0,0 @@
-(* Author: Florian Haftmann, TU Muenchen *)
-
-header {* A simple counterexample generator *}
-
-theory Quickcheck
-imports Random Code_Eval Map
-begin
-
-subsection {* The @{text random} class *}
-
-class random = typerep +
-  fixes random :: "index \<Rightarrow> seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> seed"
-
-text {* Type @{typ "'a itself"} *}
-
-instantiation itself :: ("{type, typerep}") random
-begin
-
-definition
-  "random _ = Pair (TYPE('a), \<lambda>u. Code_Eval.Const (STR ''TYPE'') TYPEREP('a))"
-
-instance ..
-
-end
-
-
-subsection {* Quickcheck generator *}
-
-ML {*
-structure StateMonad =
-struct
-
-fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT);
-fun liftT' sT = sT --> sT;
-
-fun return T sT x = Const (@{const_name Pair}, T --> liftT T sT) $ x;
-
-fun scomp T1 T2 sT f g = Const (@{const_name scomp},
-  liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g;
-
-end;
-
-structure Quickcheck =
-struct
-
-open Quickcheck;
-
-val eval_ref : (unit -> int -> int * int -> term list option * (int * int)) option ref = ref NONE;
-
-val target = "Quickcheck";
-
-fun mk_generator_expr thy prop tys =
-  let
-    val bound_max = length tys - 1;
-    val bounds = map_index (fn (i, ty) =>
-      (2 * (bound_max - i) + 1, 2 * (bound_max - i), 2 * i, ty)) tys;
-    val result = list_comb (prop, map (fn (i, _, _, _) => Bound i) bounds);
-    val terms = HOLogic.mk_list @{typ term} (map (fn (_, i, _, _) => Bound i $ @{term "()"}) bounds);
-    val check = @{term "If \<Colon> bool \<Rightarrow> term list option \<Rightarrow> term list option \<Rightarrow> term list option"}
-      $ result $ @{term "None \<Colon> term list option"} $ (@{term "Some \<Colon> term list \<Rightarrow> term list option "} $ terms);
-    val return = @{term "Pair \<Colon> term list option \<Rightarrow> seed \<Rightarrow> term list option \<times> seed"};
-    fun mk_termtyp ty = HOLogic.mk_prodT (ty, @{typ "unit \<Rightarrow> term"});
-    fun mk_split ty = Sign.mk_const thy
-      (@{const_name split}, [ty, @{typ "unit \<Rightarrow> term"}, StateMonad.liftT @{typ "term list option"} @{typ seed}]);
-    fun mk_scomp_split ty t t' =
-      StateMonad.scomp (mk_termtyp ty) @{typ "term list option"} @{typ seed} t (*FIXME*)
-        (mk_split ty $ Abs ("", ty, Abs ("", @{typ "unit \<Rightarrow> term"}, t')));
-    fun mk_bindclause (_, _, i, ty) = mk_scomp_split ty
-      (Sign.mk_const thy (@{const_name random}, [ty]) $ Bound i)
-    val t = fold_rev mk_bindclause bounds (return $ check);
-  in Abs ("n", @{typ index}, t) end;
-
-fun compile_generator_expr thy t =
-  let
-    val tys = (map snd o fst o strip_abs) t;
-    val t' = mk_generator_expr thy t tys;
-    val f = Code_ML.eval (SOME target) ("Quickcheck.eval_ref", eval_ref)
-      (fn proc => fn g => fn s => g s #>> (Option.map o map) proc) thy t' [];
-  in f #> Random_Engine.run end;
-
-end
-*}
-
-setup {*
-  Code_Target.extend_target (Quickcheck.target, (Code_ML.target_Eval, K I))
-  #> Quickcheck.add_generator ("code", Quickcheck.compile_generator_expr o ProofContext.theory_of)
-*}
-
-
-subsection {* Type @{typ "'a \<Rightarrow> 'b"} *}
-
-ML {*
-structure Random_Engine =
-struct
-
-open Random_Engine;
-
-fun random_fun (T1 : typ) (T2 : typ) (eq : 'a -> 'a -> bool) (term_of : 'a -> term)
-    (random : Random_Engine.seed -> ('b * (unit -> term)) * Random_Engine.seed)
-    (random_split : Random_Engine.seed -> Random_Engine.seed * Random_Engine.seed)
-    (seed : Random_Engine.seed) =
-  let
-    val (seed', seed'') = random_split seed;
-    val state = ref (seed', [], Const (@{const_name undefined}, T1 --> T2));
-    val fun_upd = Const (@{const_name fun_upd},
-      (T1 --> T2) --> T1 --> T2 --> T1 --> T2);
-    fun random_fun' x =
-      let
-        val (seed, fun_map, f_t) = ! state;
-      in case AList.lookup (uncurry eq) fun_map x
-       of SOME y => y
-        | NONE => let
-              val t1 = term_of x;
-              val ((y, t2), seed') = random seed;
-              val fun_map' = (x, y) :: fun_map;
-              val f_t' = fun_upd $ f_t $ t1 $ t2 ();
-              val _ = state := (seed', fun_map', f_t');
-            in y end
-      end;
-    fun term_fun' () = #3 (! state);
-  in ((random_fun', term_fun'), seed'') end;
-
-end
-*}
-
-axiomatization
-  random_fun_aux :: "typerep \<Rightarrow> typerep \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> term)
-    \<Rightarrow> (seed \<Rightarrow> ('b \<times> (unit \<Rightarrow> term)) \<times> seed) \<Rightarrow> (seed \<Rightarrow> seed \<times> seed)
-    \<Rightarrow> seed \<Rightarrow> (('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)) \<times> seed"
-
-code_const random_fun_aux (Quickcheck "Random'_Engine.random'_fun")
-  -- {* With enough criminal energy this can be abused to derive @{prop False};
-  for this reason we use a distinguished target @{text Quickcheck}
-  not spoiling the regular trusted code generation *}
-
-instantiation "fun" :: ("{eq, term_of}", "{type, random}") random
-begin
-
-definition random_fun :: "index \<Rightarrow> seed \<Rightarrow> (('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)) \<times> seed" where
-  "random n = random_fun_aux TYPEREP('a) TYPEREP('b) (op =) Code_Eval.term_of (random n) split_seed"
-
-instance ..
-
-end
-
-code_reserved Quickcheck Random_Engine
-
-end
--- a/src/HOL/Library/Random.thy	Tue May 19 14:13:23 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,147 +0,0 @@
-(* Author: Florian Haftmann, TU Muenchen *)
-
-header {* A HOL random engine *}
-
-theory Random
-imports Code_Index
-begin
-
-notation fcomp (infixl "o>" 60)
-notation scomp (infixl "o\<rightarrow>" 60)
-
-
-subsection {* Auxiliary functions *}
-
-definition inc_shift :: "index \<Rightarrow> index \<Rightarrow> index" where
-  "inc_shift v k = (if v = k then 1 else k + 1)"
-
-definition minus_shift :: "index \<Rightarrow> index \<Rightarrow> index \<Rightarrow> index" where
-  "minus_shift r k l = (if k < l then r + k - l else k - l)"
-
-fun log :: "index \<Rightarrow> index \<Rightarrow> index" where
-  "log b i = (if b \<le> 1 \<or> i < b then 1 else 1 + log b (i div b))"
-
-
-subsection {* Random seeds *}
-
-types seed = "index \<times> index"
-
-primrec "next" :: "seed \<Rightarrow> index \<times> seed" where
-  "next (v, w) = (let
-     k =  v div 53668;
-     v' = minus_shift 2147483563 (40014 * (v mod 53668)) (k * 12211);
-     l =  w div 52774;
-     w' = minus_shift 2147483399 (40692 * (w mod 52774)) (l * 3791);
-     z =  minus_shift 2147483562 v' (w' + 1) + 1
-   in (z, (v', w')))"
-
-lemma next_not_0:
-  "fst (next s) \<noteq> 0"
-  by (cases s) (auto simp add: minus_shift_def Let_def)
-
-primrec seed_invariant :: "seed \<Rightarrow> bool" where
-  "seed_invariant (v, w) \<longleftrightarrow> 0 < v \<and> v < 9438322952 \<and> 0 < w \<and> True"
-
-lemma if_same: "(if b then f x else f y) = f (if b then x else y)"
-  by (cases b) simp_all
-
-definition split_seed :: "seed \<Rightarrow> seed \<times> seed" where
-  "split_seed s = (let
-     (v, w) = s;
-     (v', w') = snd (next s);
-     v'' = inc_shift 2147483562 v;
-     s'' = (v'', w');
-     w'' = inc_shift 2147483398 w;
-     s''' = (v', w'')
-   in (s'', s'''))"
-
-
-subsection {* Base selectors *}
-
-fun iterate :: "index \<Rightarrow> ('b \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a" where
-  "iterate k f x = (if k = 0 then Pair x else f x o\<rightarrow> iterate (k - 1) f)"
-
-definition range :: "index \<Rightarrow> seed \<Rightarrow> index \<times> seed" where
-  "range k = iterate (log 2147483561 k)
-      (\<lambda>l. next o\<rightarrow> (\<lambda>v. Pair (v + l * 2147483561))) 1
-    o\<rightarrow> (\<lambda>v. Pair (v mod k))"
-
-lemma range:
-  "k > 0 \<Longrightarrow> fst (range k s) < k"
-  by (simp add: range_def scomp_apply split_def del: log.simps iterate.simps)
-
-definition select :: "'a list \<Rightarrow> seed \<Rightarrow> 'a \<times> seed" where
-  "select xs = range (Code_Index.of_nat (length xs))
-    o\<rightarrow> (\<lambda>k. Pair (nth xs (Code_Index.nat_of k)))"
-     
-lemma select:
-  assumes "xs \<noteq> []"
-  shows "fst (select xs s) \<in> set xs"
-proof -
-  from assms have "Code_Index.of_nat (length xs) > 0" by simp
-  with range have
-    "fst (range (Code_Index.of_nat (length xs)) s) < Code_Index.of_nat (length xs)" by best
-  then have
-    "Code_Index.nat_of (fst (range (Code_Index.of_nat (length xs)) s)) < length xs" by simp
-  then show ?thesis
-    by (simp add: scomp_apply split_beta select_def)
-qed
-
-definition select_default :: "index \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> seed \<Rightarrow> 'a \<times> seed" where
-  [code del]: "select_default k x y = range k
-     o\<rightarrow> (\<lambda>l. Pair (if l + 1 < k then x else y))"
-
-lemma select_default_zero:
-  "fst (select_default 0 x y s) = y"
-  by (simp add: scomp_apply split_beta select_default_def)
-
-lemma select_default_code [code]:
-  "select_default k x y = (if k = 0
-    then range 1 o\<rightarrow> (\<lambda>_. Pair y)
-    else range k o\<rightarrow> (\<lambda>l. Pair (if l + 1 < k then x else y)))"
-proof
-  fix s
-  have "snd (range (Code_Index.of_nat 0) s) = snd (range (Code_Index.of_nat 1) s)"
-    by (simp add: range_def scomp_Pair scomp_apply split_beta)
-  then show "select_default k x y s = (if k = 0
-    then range 1 o\<rightarrow> (\<lambda>_. Pair y)
-    else range k o\<rightarrow> (\<lambda>l. Pair (if l + 1 < k then x else y))) s"
-    by (cases "k = 0") (simp_all add: select_default_def scomp_apply split_beta)
-qed
-
-
-subsection {* @{text ML} interface *}
-
-ML {*
-structure Random_Engine =
-struct
-
-type seed = int * int;
-
-local
-
-val seed = ref 
-  (let
-    val now = Time.toMilliseconds (Time.now ());
-    val (q, s1) = IntInf.divMod (now, 2147483562);
-    val s2 = q mod 2147483398;
-  in (s1 + 1, s2 + 1) end);
-
-in
-
-fun run f =
-  let
-    val (x, seed') = f (! seed);
-    val _ = seed := seed'
-  in x end;
-
-end;
-
-end;
-*}
-
-no_notation fcomp (infixl "o>" 60)
-no_notation scomp (infixl "o\<rightarrow>" 60)
-
-end
-
--- a/src/HOL/List.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/List.thy	Tue May 19 14:13:37 2009 +0100
@@ -1299,6 +1299,25 @@
   show ?case by (clarsimp simp add: Cons nth_append)
 qed
 
+lemma Skolem_list_nth:
+  "(ALL i<k. EX x. P i x) = (EX xs. size xs = k & (ALL i<k. P i (xs!i)))"
+  (is "_ = (EX xs. ?P k xs)")
+proof(induct k)
+  case 0 show ?case by simp
+next
+  case (Suc k)
+  show ?case (is "?L = ?R" is "_ = (EX xs. ?P' xs)")
+  proof
+    assume "?R" thus "?L" using Suc by auto
+  next
+    assume "?L"
+    with Suc obtain x xs where "?P k xs & P k x" by (metis less_Suc_eq)
+    hence "?P'(xs@[x])" by(simp add:nth_append less_Suc_eq)
+    thus "?R" ..
+  qed
+qed
+
+
 subsubsection {* @{text list_update} *}
 
 lemma length_list_update [simp]: "length(xs[i:=x]) = length xs"
@@ -2894,6 +2913,10 @@
   "sorted (xs@ys) = (sorted xs & sorted ys & (\<forall>x \<in> set xs. \<forall>y \<in> set ys. x\<le>y))"
 by (induct xs) (auto simp add:sorted_Cons)
 
+lemma sorted_nth_mono:
+  "sorted xs \<Longrightarrow> i <= j \<Longrightarrow> j < length xs \<Longrightarrow> xs!i <= xs!j"
+by (induct xs arbitrary: i j) (auto simp:nth_Cons' sorted_Cons)
+
 lemma set_insort: "set(insort x xs) = insert x (set xs)"
 by (induct xs) auto
 
@@ -3646,10 +3669,14 @@
 
 lemmas in_set_code [code unfold] = mem_iff [symmetric]
 
-lemma empty_null [code inline]:
+lemma empty_null:
   "xs = [] \<longleftrightarrow> null xs"
 by (cases xs) simp_all
 
+lemma [code inline]:
+  "eq_class.eq xs [] \<longleftrightarrow> null xs"
+by (simp add: eq empty_null)
+
 lemmas null_empty [code post] =
   empty_null [symmetric]
 
--- a/src/HOL/MacLaurin.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/MacLaurin.thy	Tue May 19 14:13:37 2009 +0100
@@ -552,10 +552,6 @@
     "[|x = y; abs u \<le> (v::real) |] ==> \<bar>(x + u) - y\<bar> \<le> v"
 by auto
 
-text {* TODO: move to Parity.thy *}
-lemma nat_odd_1 [simp]: "odd (1::nat)"
-  unfolding even_nat_def by simp
-
 lemma Maclaurin_sin_bound:
   "abs(sin x - (\<Sum>m=0..<n. (if even m then 0 else (-1 ^ ((m - Suc 0) div 2)) / real (fact m)) *
   x ^ m))  \<le> inverse(real (fact n)) * \<bar>x\<bar> ^ n"
--- a/src/HOL/MicroJava/BV/BVExample.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/MicroJava/BV/BVExample.thy	Tue May 19 14:13:37 2009 +0100
@@ -69,10 +69,10 @@
 lemma subcls1:
   "subcls1 E = (\<lambda>C D. (C, D) \<in> {(list_name,Object), (test_name,Object), (Xcpt NullPointer, Object),
                 (Xcpt ClassCast, Object), (Xcpt OutOfMemory, Object)})"
-  apply (simp add: subcls1_def2)
-  apply (simp add: name_defs class_defs system_defs E_def class_def)
-  apply (auto simp: expand_fun_eq split: split_if_asm)
-  done
+apply (simp add: subcls1_def2)
+apply (simp add: name_defs class_defs system_defs E_def class_def)
+apply (auto simp: expand_fun_eq)
+done
 
 text {* The subclass relation is acyclic; hence its converse is well founded: *}
 lemma notin_rtrancl:
--- a/src/HOL/NanoJava/TypeRel.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/NanoJava/TypeRel.thy	Tue May 19 14:13:37 2009 +0100
@@ -56,7 +56,7 @@
   "subcls1 = 
     (SIGMA C: {C. is_class C} . {D. C\<noteq>Object \<and> super (the (class C)) = D})"
 apply (unfold subcls1_def is_class_def)
-apply auto
+apply (auto split:split_if_asm)
 done
 
 lemma finite_subcls1: "finite subcls1"
--- a/src/HOL/Nat.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Nat.thy	Tue May 19 14:13:37 2009 +0100
@@ -1199,7 +1199,7 @@
 definition funpow :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" where
   funpow_code_def [code post]: "funpow = compow"
 
-lemmas [code inline] = funpow_code_def [symmetric]
+lemmas [code unfold] = funpow_code_def [symmetric]
 
 lemma [code]:
   "funpow 0 f = id"
--- a/src/HOL/Nat_Numeral.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Nat_Numeral.thy	Tue May 19 14:13:37 2009 +0100
@@ -316,13 +316,13 @@
 lemma nat_number_of [simp]: "nat (number_of w) = number_of w"
 by (simp add: nat_number_of_def)
 
-lemma nat_numeral_0_eq_0 [simp]: "Numeral0 = (0::nat)"
+lemma nat_numeral_0_eq_0 [simp, code post]: "Numeral0 = (0::nat)"
 by (simp add: nat_number_of_def)
 
 lemma nat_numeral_1_eq_1 [simp]: "Numeral1 = (1::nat)"
 by (simp add: nat_1 nat_number_of_def)
 
-lemma numeral_1_eq_Suc_0: "Numeral1 = Suc 0"
+lemma numeral_1_eq_Suc_0 [code post]: "Numeral1 = Suc 0"
 by (simp add: nat_numeral_1_eq_1)
 
 
--- a/src/HOL/Nominal/nominal_inductive.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Nominal/nominal_inductive.ML	Tue May 19 14:13:37 2009 +0100
@@ -561,7 +561,7 @@
             (strong_raw_induct, [ind_case_names, RuleCases.consumes 0])
           else (strong_raw_induct RSN (2, rev_mp),
             [ind_case_names, RuleCases.consumes 1]);
-        val ((_, [strong_induct']), ctxt') = LocalTheory.note Thm.theoremK
+        val ((_, [strong_induct']), ctxt') = LocalTheory.note Thm.generatedK
           ((rec_qualified (Binding.name "strong_induct"),
             map (Attrib.internal o K) (#2 strong_induct)), [#1 strong_induct])
           ctxt;
@@ -569,12 +569,12 @@
           ProjectRule.projects ctxt (1 upto length names) strong_induct'
       in
         ctxt' |>
-        LocalTheory.note Thm.theoremK
+        LocalTheory.note Thm.generatedK
           ((rec_qualified (Binding.name "strong_inducts"),
             [Attrib.internal (K ind_case_names),
              Attrib.internal (K (RuleCases.consumes 1))]),
            strong_inducts) |> snd |>
-        LocalTheory.notes Thm.theoremK (map (fn ((name, elim), (_, cases)) =>
+        LocalTheory.notes Thm.generatedK (map (fn ((name, elim), (_, cases)) =>
             ((Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "strong_cases"),
               [Attrib.internal (K (RuleCases.case_names (map snd cases))),
                Attrib.internal (K (RuleCases.consumes 1))]), [([elim], [])]))
@@ -664,7 +664,7 @@
       end) atoms
   in
     ctxt |>
-    LocalTheory.notes Thm.theoremK (map (fn (name, ths) =>
+    LocalTheory.notes Thm.generatedK (map (fn (name, ths) =>
         ((Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "eqvt"),
           [Attrib.internal (K NominalThmDecls.eqvt_add)]), [(ths, [])]))
       (names ~~ transp thss)) |> snd
--- a/src/HOL/Nominal/nominal_inductive2.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Nominal/nominal_inductive2.ML	Tue May 19 14:13:37 2009 +0100
@@ -461,7 +461,7 @@
             (strong_raw_induct, [ind_case_names, RuleCases.consumes 0])
           else (strong_raw_induct RSN (2, rev_mp),
             [ind_case_names, RuleCases.consumes 1]);
-        val ((_, [strong_induct']), ctxt') = LocalTheory.note Thm.theoremK
+        val ((_, [strong_induct']), ctxt') = LocalTheory.note Thm.generatedK
           ((rec_qualified (Binding.name "strong_induct"),
             map (Attrib.internal o K) (#2 strong_induct)), [#1 strong_induct])
           ctxt;
@@ -469,7 +469,7 @@
           ProjectRule.projects ctxt' (1 upto length names) strong_induct'
       in
         ctxt' |>
-        LocalTheory.note Thm.theoremK
+        LocalTheory.note Thm.generatedK
           ((rec_qualified (Binding.name "strong_inducts"),
             [Attrib.internal (K ind_case_names),
              Attrib.internal (K (RuleCases.consumes 1))]),
--- a/src/HOL/Nominal/nominal_primrec.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Nominal/nominal_primrec.ML	Tue May 19 14:13:37 2009 +0100
@@ -367,11 +367,11 @@
       (fn thss => fn goal_ctxt =>
          let
            val simps = ProofContext.export goal_ctxt lthy' (flat thss);
-           val (simps', lthy'') = fold_map (LocalTheory.note Thm.theoremK)
+           val (simps', lthy'') = fold_map (LocalTheory.note Thm.generatedK)
              (names_atts' ~~ map single simps) lthy'
          in
            lthy''
-           |> LocalTheory.note Thm.theoremK ((qualify (Binding.name "simps"),
+           |> LocalTheory.note Thm.generatedK ((qualify (Binding.name "simps"),
                 map (Attrib.internal o K)
                     [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add]),
                 maps snd simps')
--- a/src/HOL/Option.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Option.thy	Tue May 19 14:13:37 2009 +0100
@@ -63,10 +63,8 @@
 lemma set_empty_eq [simp]: "(set xo = {}) = (xo = None)"
   by (cases xo) auto
 
-definition
-  map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a option \<Rightarrow> 'b option"
-where
-  [code del]: "map = (%f y. case y of None => None | Some x => Some (f x))"
+definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a option \<Rightarrow> 'b option" where
+  "map = (%f y. case y of None => None | Some x => Some (f x))"
 
 lemma option_map_None [simp, code]: "map f None = None"
   by (simp add: map_def)
@@ -95,14 +93,21 @@
 
 subsubsection {* Code generator setup *}
 
-definition
-  is_none :: "'a option \<Rightarrow> bool" where
-  is_none_none [code post, symmetric, code inline]: "is_none x \<longleftrightarrow> x = None"
+definition is_none :: "'a option \<Rightarrow> bool" where
+  [code post]: "is_none x \<longleftrightarrow> x = None"
 
 lemma is_none_code [code]:
   shows "is_none None \<longleftrightarrow> True"
     and "is_none (Some x) \<longleftrightarrow> False"
-  unfolding is_none_none [symmetric] by simp_all
+  unfolding is_none_def by simp_all
+
+lemma is_none_none:
+  "is_none x \<longleftrightarrow> x = None"
+  by (simp add: is_none_def)
+
+lemma [code inline]:
+  "eq_class.eq x None \<longleftrightarrow> is_none x"
+  by (simp add: eq is_none_none)
 
 hide (open) const is_none
 
--- a/src/HOL/Parity.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Parity.thy	Tue May 19 14:13:37 2009 +0100
@@ -29,6 +29,18 @@
 end
 
 
+lemma even_zero_int[simp]: "even (0::int)" by presburger
+
+lemma odd_one_int[simp]: "odd (1::int)" by presburger
+
+lemma even_zero_nat[simp]: "even (0::nat)" by presburger
+
+lemma odd_zero_nat [simp]: "odd (1::nat)" by presburger
+
+declare even_def[of "number_of v", standard, simp]
+
+declare even_nat_def[of "number_of v", standard, simp]
+
 subsection {* Even and odd are mutually exclusive *}
 
 lemma int_pos_lt_two_imp_zero_or_one:
@@ -54,66 +66,47 @@
 lemma odd_times_odd: "odd (x::int) ==> odd y ==> odd (x * y)" 
   by (simp add: even_def zmod_zmult1_eq)
 
-lemma even_product[presburger]: "even((x::int) * y) = (even x | even y)"
+lemma even_product[simp,presburger]: "even((x::int) * y) = (even x | even y)"
   apply (auto simp add: even_times_anything anything_times_even)
   apply (rule ccontr)
   apply (auto simp add: odd_times_odd)
   done
 
 lemma even_plus_even: "even (x::int) ==> even y ==> even (x + y)"
-  by presburger
+by presburger
 
 lemma even_plus_odd: "even (x::int) ==> odd y ==> odd (x + y)"
-  by presburger
+by presburger
 
 lemma odd_plus_even: "odd (x::int) ==> even y ==> odd (x + y)"
-  by presburger
+by presburger
 
 lemma odd_plus_odd: "odd (x::int) ==> odd y ==> even (x + y)" by presburger
 
-lemma even_sum[presburger]: "even ((x::int) + y) = ((even x & even y) | (odd x & odd y))"
-  by presburger
+lemma even_sum[simp,presburger]:
+  "even ((x::int) + y) = ((even x & even y) | (odd x & odd y))"
+by presburger
 
-lemma even_neg[presburger, algebra]: "even (-(x::int)) = even x" by presburger
+lemma even_neg[simp,presburger,algebra]: "even (-(x::int)) = even x"
+by presburger
 
-lemma even_difference:
+lemma even_difference[simp]:
     "even ((x::int) - y) = ((even x & even y) | (odd x & odd y))" by presburger
 
-lemma even_pow_gt_zero:
-    "even (x::int) ==> 0 < n ==> even (x^n)"
-  by (induct n) (auto simp add: even_product)
-
-lemma odd_pow_iff[presburger, algebra]: 
-  "odd ((x::int) ^ n) \<longleftrightarrow> (n = 0 \<or> odd x)"
-  apply (induct n, simp_all)
-  apply presburger
-  apply (case_tac n, auto)
-  apply (simp_all add: even_product)
-  done
+lemma even_power[simp,presburger]: "even ((x::int)^n) = (even x & n \<noteq> 0)"
+by (induct n) auto
 
-lemma odd_pow: "odd x ==> odd((x::int)^n)" by (simp add: odd_pow_iff)
-
-lemma even_power[presburger]: "even ((x::int)^n) = (even x & 0 < n)"
-  apply (auto simp add: even_pow_gt_zero)
-  apply (erule contrapos_pp, erule odd_pow)
-  apply (erule contrapos_pp, simp add: even_def)
-  done
-
-lemma even_zero[presburger]: "even (0::int)" by presburger
-
-lemma odd_one[presburger]: "odd (1::int)" by presburger
-
-lemmas even_odd_simps [simp] = even_def[of "number_of v",standard] even_zero
-  odd_one even_product even_sum even_neg even_difference even_power
+lemma odd_pow: "odd x ==> odd((x::int)^n)" by simp
 
 
 subsection {* Equivalent definitions *}
 
 lemma two_times_even_div_two: "even (x::int) ==> 2 * (x div 2) = x" 
-  by presburger
+by presburger
 
-lemma two_times_odd_div_two_plus_one: "odd (x::int) ==>
-    2 * (x div 2) + 1 = x" by presburger
+lemma two_times_odd_div_two_plus_one:
+  "odd (x::int) ==> 2 * (x div 2) + 1 = x"
+by presburger
 
 lemma even_equiv_def: "even (x::int) = (EX y. x = 2 * y)" by presburger
 
@@ -122,45 +115,45 @@
 subsection {* even and odd for nats *}
 
 lemma pos_int_even_equiv_nat_even: "0 \<le> x ==> even x = even (nat x)"
-  by (simp add: even_nat_def)
-
-lemma even_nat_product[presburger, algebra]: "even((x::nat) * y) = (even x | even y)"
-  by (simp add: even_nat_def int_mult)
+by (simp add: even_nat_def)
 
-lemma even_nat_sum[presburger, algebra]: "even ((x::nat) + y) =
-    ((even x & even y) | (odd x & odd y))" by presburger
+lemma even_product_nat[simp,presburger,algebra]:
+  "even((x::nat) * y) = (even x | even y)"
+by (simp add: even_nat_def int_mult)
 
-lemma even_nat_difference[presburger, algebra]:
-    "even ((x::nat) - y) = (x < y | (even x & even y) | (odd x & odd y))"
+lemma even_sum_nat[simp,presburger,algebra]:
+  "even ((x::nat) + y) = ((even x & even y) | (odd x & odd y))"
 by presburger
 
-lemma even_nat_Suc[presburger, algebra]: "even (Suc x) = odd x" by presburger
-
-lemma even_nat_power[presburger, algebra]: "even ((x::nat)^y) = (even x & 0 < y)"
-  by (simp add: even_nat_def int_power)
+lemma even_difference_nat[simp,presburger,algebra]:
+  "even ((x::nat) - y) = (x < y | (even x & even y) | (odd x & odd y))"
+by presburger
 
-lemma even_nat_zero[presburger]: "even (0::nat)" by presburger
+lemma even_Suc[simp,presburger,algebra]: "even (Suc x) = odd x"
+by presburger
 
-lemmas even_odd_nat_simps [simp] = even_nat_def[of "number_of v",standard]
-  even_nat_zero even_nat_Suc even_nat_product even_nat_sum even_nat_power
+lemma even_power_nat[simp,presburger,algebra]:
+  "even ((x::nat)^y) = (even x & 0 < y)"
+by (simp add: even_nat_def int_power)
 
 
 subsection {* Equivalent definitions *}
 
-lemma nat_lt_two_imp_zero_or_one: "(x::nat) < Suc (Suc 0) ==>
-    x = 0 | x = Suc 0" by presburger
+lemma nat_lt_two_imp_zero_or_one:
+  "(x::nat) < Suc (Suc 0) ==> x = 0 | x = Suc 0"
+by presburger
 
 lemma even_nat_mod_two_eq_zero: "even (x::nat) ==> x mod (Suc (Suc 0)) = 0"
-  by presburger
+by presburger
 
 lemma odd_nat_mod_two_eq_one: "odd (x::nat) ==> x mod (Suc (Suc 0)) = Suc 0"
 by presburger
 
 lemma even_nat_equiv_def: "even (x::nat) = (x mod Suc (Suc 0) = 0)"
-  by presburger
+by presburger
 
 lemma odd_nat_equiv_def: "odd (x::nat) = (x mod Suc (Suc 0) = Suc 0)"
-  by presburger
+by presburger
 
 lemma even_nat_div_two_times_two: "even (x::nat) ==>
     Suc (Suc 0) * (x div Suc (Suc 0)) = x" by presburger
@@ -169,10 +162,10 @@
     Suc( Suc (Suc 0) * (x div Suc (Suc 0))) = x" by presburger
 
 lemma even_nat_equiv_def2: "even (x::nat) = (EX y. x = Suc (Suc 0) * y)"
-  by presburger
+by presburger
 
 lemma odd_nat_equiv_def2: "odd (x::nat) = (EX y. x = Suc(Suc (Suc 0) * y))"
-  by presburger
+by presburger
 
 
 subsection {* Parity and powers *}
@@ -183,7 +176,7 @@
   apply (induct x)
   apply (rule conjI)
   apply simp
-  apply (insert even_nat_zero, blast)
+  apply (insert even_zero_nat, blast)
   apply (simp add: power_Suc)
   done
 
--- a/src/HOL/Power.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Power.thy	Tue May 19 14:13:37 2009 +0100
@@ -452,4 +452,13 @@
   from power_strict_increasing_iff [OF this] less show ?thesis ..
 qed
 
+
+subsection {* Code generator tweak *}
+
+lemma power_power_power [code, code unfold, code inline del]:
+  "power = power.power (1::'a::{power}) (op *)"
+  unfolding power_def power.power_def ..
+
+declare power.power.simps [code]
+
 end
--- a/src/HOL/Predicate.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Predicate.thy	Tue May 19 14:13:37 2009 +0100
@@ -610,7 +610,7 @@
     (simp_all add: Seq_def single_less_eq_eval contained_less_eq)
 
 lemma eq_pred_code [code]:
-  fixes P Q :: "'a::eq pred"
+  fixes P Q :: "'a pred"
   shows "eq_class.eq P Q \<longleftrightarrow> P \<le> Q \<and> Q \<le> P"
   unfolding eq by auto
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Quickcheck.thy	Tue May 19 14:13:37 2009 +0100
@@ -0,0 +1,212 @@
+(* Author: Florian Haftmann, TU Muenchen *)
+
+header {* A simple counterexample generator *}
+
+theory Quickcheck
+imports Main Real Random
+begin
+
+notation fcomp (infixl "o>" 60)
+notation scomp (infixl "o\<rightarrow>" 60)
+
+
+subsection {* The @{text random} class *}
+
+class random = typerep +
+  fixes random :: "index \<Rightarrow> Random.seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> Random.seed"
+
+
+subsection {* Quickcheck generator *}
+
+ML {*
+structure Quickcheck =
+struct
+
+open Quickcheck;
+
+val eval_ref : (unit -> int -> int * int -> term list option * (int * int)) option ref = ref NONE;
+
+val target = "Quickcheck";
+
+fun mk_generator_expr thy prop tys =
+  let
+    val bound_max = length tys - 1;
+    val bounds = map_index (fn (i, ty) =>
+      (2 * (bound_max - i) + 1, 2 * (bound_max - i), 2 * i, ty)) tys;
+    val result = list_comb (prop, map (fn (i, _, _, _) => Bound i) bounds);
+    val terms = HOLogic.mk_list @{typ term} (map (fn (_, i, _, _) => Bound i $ @{term "()"}) bounds);
+    val check = @{term "If \<Colon> bool \<Rightarrow> term list option \<Rightarrow> term list option \<Rightarrow> term list option"}
+      $ result $ @{term "None \<Colon> term list option"} $ (@{term "Some \<Colon> term list \<Rightarrow> term list option "} $ terms);
+    val return = @{term "Pair \<Colon> term list option \<Rightarrow> Random.seed \<Rightarrow> term list option \<times> Random.seed"};
+    fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT);
+    fun mk_termtyp ty = HOLogic.mk_prodT (ty, @{typ "unit \<Rightarrow> term"});
+    fun mk_scomp T1 T2 sT f g = Const (@{const_name scomp},
+      liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g;
+    fun mk_split ty = Sign.mk_const thy
+      (@{const_name split}, [ty, @{typ "unit \<Rightarrow> term"}, liftT @{typ "term list option"} @{typ Random.seed}]);
+    fun mk_scomp_split ty t t' =
+      mk_scomp (mk_termtyp ty) @{typ "term list option"} @{typ Random.seed} t
+        (mk_split ty $ Abs ("", ty, Abs ("", @{typ "unit \<Rightarrow> term"}, t')));
+    fun mk_bindclause (_, _, i, ty) = mk_scomp_split ty
+      (Sign.mk_const thy (@{const_name random}, [ty]) $ Bound i);
+  in Abs ("n", @{typ index}, fold_rev mk_bindclause bounds (return $ check)) end;
+
+fun compile_generator_expr thy t =
+  let
+    val tys = (map snd o fst o strip_abs) t;
+    val t' = mk_generator_expr thy t tys;
+    val f = Code_ML.eval (SOME target) ("Quickcheck.eval_ref", eval_ref)
+      (fn proc => fn g => fn s => g s #>> (Option.map o map) proc) thy t' [];
+  in f #> Random_Engine.run end;
+
+end
+*}
+
+setup {*
+  Code_Target.extend_target (Quickcheck.target, (Code_ML.target_Eval, K I))
+  #> Quickcheck.add_generator ("code", Quickcheck.compile_generator_expr o ProofContext.theory_of)
+*}
+
+
+subsection {* Fundamental types*}
+
+instantiation bool :: random
+begin
+
+definition
+  "random i = Random.range i o\<rightarrow>
+    (\<lambda>k. Pair (if (k div 2 = 0) then Code_Eval.valtermify True else Code_Eval.valtermify False))"
+
+instance ..
+
+end
+
+instantiation itself :: (typerep) random
+begin
+
+definition random_itself :: "index \<Rightarrow> Random.seed \<Rightarrow> ('a itself \<times> (unit \<Rightarrow> term)) \<times> Random.seed" where
+  "random_itself _ = Pair (Code_Eval.valtermify TYPE('a))"
+
+instance ..
+
+end
+
+text {* Type @{typ "'a \<Rightarrow> 'b"} *}
+
+ML {*
+structure Random_Engine =
+struct
+
+open Random_Engine;
+
+fun random_fun (T1 : typ) (T2 : typ) (eq : 'a -> 'a -> bool) (term_of : 'a -> term)
+    (random : Random_Engine.seed -> ('b * (unit -> term)) * Random_Engine.seed)
+    (random_split : Random_Engine.seed -> Random_Engine.seed * Random_Engine.seed)
+    (seed : Random_Engine.seed) =
+  let
+    val (seed', seed'') = random_split seed;
+    val state = ref (seed', [], Const (@{const_name undefined}, T1 --> T2));
+    val fun_upd = Const (@{const_name fun_upd},
+      (T1 --> T2) --> T1 --> T2 --> T1 --> T2);
+    fun random_fun' x =
+      let
+        val (seed, fun_map, f_t) = ! state;
+      in case AList.lookup (uncurry eq) fun_map x
+       of SOME y => y
+        | NONE => let
+              val t1 = term_of x;
+              val ((y, t2), seed') = random seed;
+              val fun_map' = (x, y) :: fun_map;
+              val f_t' = fun_upd $ f_t $ t1 $ t2 ();
+              val _ = state := (seed', fun_map', f_t');
+            in y end
+      end;
+    fun term_fun' () = #3 (! state);
+  in ((random_fun', term_fun'), seed'') end;
+
+end
+*}
+
+axiomatization random_fun_aux :: "typerep \<Rightarrow> typerep \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> term)
+  \<Rightarrow> (Random.seed \<Rightarrow> ('b \<times> (unit \<Rightarrow> term)) \<times> Random.seed) \<Rightarrow> (Random.seed \<Rightarrow> Random.seed \<times> Random.seed)
+  \<Rightarrow> Random.seed \<Rightarrow> (('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)) \<times> Random.seed"
+
+code_const random_fun_aux (Quickcheck "Random'_Engine.random'_fun")
+  -- {* With enough criminal energy this can be abused to derive @{prop False};
+  for this reason we use a distinguished target @{text Quickcheck}
+  not spoiling the regular trusted code generation *}
+
+instantiation "fun" :: ("{eq, term_of}", "{type, random}") random
+begin
+
+definition random_fun :: "index \<Rightarrow> Random.seed \<Rightarrow> (('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)) \<times> Random.seed" where
+  "random n = random_fun_aux TYPEREP('a) TYPEREP('b) (op =) Code_Eval.term_of (random n) Random.split_seed"
+
+instance ..
+
+end
+
+code_reserved Quickcheck Random_Engine
+
+
+subsection {* Numeric types *}
+
+instantiation nat :: random
+begin
+
+definition random_nat :: "index \<Rightarrow> Random.seed \<Rightarrow> (nat \<times> (unit \<Rightarrow> Code_Eval.term)) \<times> Random.seed" where
+  "random_nat i = Random.range (i + 1) o\<rightarrow> (\<lambda>k. Pair (
+     let n = Code_Index.nat_of k
+     in (n, \<lambda>_. Code_Eval.term_of n)))"
+
+instance ..
+
+end
+
+instantiation int :: random
+begin
+
+definition
+  "random i = Random.range (2 * i + 1) o\<rightarrow> (\<lambda>k. Pair (
+     let j = (if k \<ge> i then Code_Index.int_of (k - i) else - Code_Index.int_of (i - k))
+     in (j, \<lambda>_. Code_Eval.term_of j)))"
+
+instance ..
+
+end
+
+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 inline]: "valterm_fract k l = Code_Eval.valtermify Fract {\<cdot>} k {\<cdot>} l"
+
+instantiation rat :: random
+begin
+
+definition
+  "random i = random i o\<rightarrow> (\<lambda>num. Random.range i o\<rightarrow> (\<lambda>denom. Pair (
+     let j = Code_Index.int_of (denom + 1)
+     in valterm_fract num (j, \<lambda>u. Code_Eval.term_of j))))"
+
+instance ..
+
+end
+
+definition (in term_syntax)
+  valterm_ratreal :: "rat \<times> (unit \<Rightarrow> Code_Eval.term) \<Rightarrow> real \<times> (unit \<Rightarrow> Code_Eval.term)" where
+  [code inline]: "valterm_ratreal k = Code_Eval.valtermify Ratreal {\<cdot>} k"
+
+instantiation real :: random
+begin
+
+definition
+  "random i = random i o\<rightarrow> (\<lambda>r. Pair (valterm_ratreal r))"
+
+instance ..
+
+end
+
+
+no_notation fcomp (infixl "o>" 60)
+no_notation scomp (infixl "o\<rightarrow>" 60)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Random.thy	Tue May 19 14:13:37 2009 +0100
@@ -0,0 +1,178 @@
+(* Author: Florian Haftmann, TU Muenchen *)
+
+header {* A HOL random engine *}
+
+theory Random
+imports Code_Index
+begin
+
+notation fcomp (infixl "o>" 60)
+notation scomp (infixl "o\<rightarrow>" 60)
+
+
+subsection {* Auxiliary functions *}
+
+definition inc_shift :: "index \<Rightarrow> index \<Rightarrow> index" where
+  "inc_shift v k = (if v = k then 1 else k + 1)"
+
+definition minus_shift :: "index \<Rightarrow> index \<Rightarrow> index \<Rightarrow> index" where
+  "minus_shift r k l = (if k < l then r + k - l else k - l)"
+
+fun log :: "index \<Rightarrow> index \<Rightarrow> index" where
+  "log b i = (if b \<le> 1 \<or> i < b then 1 else 1 + log b (i div b))"
+
+
+subsection {* Random seeds *}
+
+types seed = "index \<times> index"
+
+primrec "next" :: "seed \<Rightarrow> index \<times> seed" where
+  "next (v, w) = (let
+     k =  v div 53668;
+     v' = minus_shift 2147483563 (40014 * (v mod 53668)) (k * 12211);
+     l =  w div 52774;
+     w' = minus_shift 2147483399 (40692 * (w mod 52774)) (l * 3791);
+     z =  minus_shift 2147483562 v' (w' + 1) + 1
+   in (z, (v', w')))"
+
+lemma next_not_0:
+  "fst (next s) \<noteq> 0"
+  by (cases s) (auto simp add: minus_shift_def Let_def)
+
+primrec seed_invariant :: "seed \<Rightarrow> bool" where
+  "seed_invariant (v, w) \<longleftrightarrow> 0 < v \<and> v < 9438322952 \<and> 0 < w \<and> True"
+
+lemma if_same: "(if b then f x else f y) = f (if b then x else y)"
+  by (cases b) simp_all
+
+definition split_seed :: "seed \<Rightarrow> seed \<times> seed" where
+  "split_seed s = (let
+     (v, w) = s;
+     (v', w') = snd (next s);
+     v'' = inc_shift 2147483562 v;
+     s'' = (v'', w');
+     w'' = inc_shift 2147483398 w;
+     s''' = (v', w'')
+   in (s'', s'''))"
+
+
+subsection {* Base selectors *}
+
+fun iterate :: "index \<Rightarrow> ('b \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a" where
+  "iterate k f x = (if k = 0 then Pair x else f x o\<rightarrow> iterate (k - 1) f)"
+
+definition range :: "index \<Rightarrow> seed \<Rightarrow> index \<times> seed" where
+  "range k = iterate (log 2147483561 k)
+      (\<lambda>l. next o\<rightarrow> (\<lambda>v. Pair (v + l * 2147483561))) 1
+    o\<rightarrow> (\<lambda>v. Pair (v mod k))"
+
+lemma range:
+  "k > 0 \<Longrightarrow> fst (range k s) < k"
+  by (simp add: range_def scomp_apply split_def del: log.simps iterate.simps)
+
+definition select :: "'a list \<Rightarrow> seed \<Rightarrow> 'a \<times> seed" where
+  "select xs = range (Code_Index.of_nat (length xs))
+    o\<rightarrow> (\<lambda>k. Pair (nth xs (Code_Index.nat_of k)))"
+     
+lemma select:
+  assumes "xs \<noteq> []"
+  shows "fst (select xs s) \<in> set xs"
+proof -
+  from assms have "Code_Index.of_nat (length xs) > 0" by simp
+  with range have
+    "fst (range (Code_Index.of_nat (length xs)) s) < Code_Index.of_nat (length xs)" by best
+  then have
+    "Code_Index.nat_of (fst (range (Code_Index.of_nat (length xs)) s)) < length xs" by simp
+  then show ?thesis
+    by (simp add: scomp_apply split_beta select_def)
+qed
+
+primrec pick :: "(index \<times> 'a) list \<Rightarrow> index \<Rightarrow> 'a" where
+  "pick (x # xs) i = (if i < fst x then snd x else pick xs (i - fst x))"
+
+lemma pick_member:
+  "i < listsum (map fst xs) \<Longrightarrow> pick xs i \<in> set (map snd xs)"
+  by (induct xs arbitrary: i) simp_all
+
+lemma pick_drop_zero:
+  "pick (filter (\<lambda>(k, _). k > 0) xs) = pick xs"
+  by (induct xs) (auto simp add: expand_fun_eq)
+
+definition select_weight :: "(index \<times> 'a) list \<Rightarrow> seed \<Rightarrow> 'a \<times> seed" where
+  "select_weight xs = range (listsum (map fst xs))
+   o\<rightarrow> (\<lambda>k. Pair (pick xs k))"
+
+lemma select_weight_member:
+  assumes "0 < listsum (map fst xs)"
+  shows "fst (select_weight xs s) \<in> set (map snd xs)"
+proof -
+  from range assms
+    have "fst (range (listsum (map fst xs)) s) < listsum (map fst xs)" .
+  with pick_member
+    have "pick xs (fst (range (listsum (map fst xs)) s)) \<in> set (map snd xs)" .
+  then show ?thesis by (simp add: select_weight_def scomp_def split_def) 
+qed
+
+definition select_default :: "index \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> seed \<Rightarrow> 'a \<times> seed" where
+  [code del]: "select_default k x y = range k
+     o\<rightarrow> (\<lambda>l. Pair (if l + 1 < k then x else y))"
+
+lemma select_default_zero:
+  "fst (select_default 0 x y s) = y"
+  by (simp add: scomp_apply split_beta select_default_def)
+
+lemma select_default_code [code]:
+  "select_default k x y = (if k = 0
+    then range 1 o\<rightarrow> (\<lambda>_. Pair y)
+    else range k o\<rightarrow> (\<lambda>l. Pair (if l + 1 < k then x else y)))"
+proof
+  fix s
+  have "snd (range (Code_Index.of_nat 0) s) = snd (range (Code_Index.of_nat 1) s)"
+    by (simp add: range_def scomp_Pair scomp_apply split_beta)
+  then show "select_default k x y s = (if k = 0
+    then range 1 o\<rightarrow> (\<lambda>_. Pair y)
+    else range k o\<rightarrow> (\<lambda>l. Pair (if l + 1 < k then x else y))) s"
+    by (cases "k = 0") (simp_all add: select_default_def scomp_apply split_beta)
+qed
+
+
+subsection {* @{text ML} interface *}
+
+ML {*
+structure Random_Engine =
+struct
+
+type seed = int * int;
+
+local
+
+val seed = ref 
+  (let
+    val now = Time.toMilliseconds (Time.now ());
+    val (q, s1) = IntInf.divMod (now, 2147483562);
+    val s2 = q mod 2147483398;
+  in (s1 + 1, s2 + 1) end);
+
+in
+
+fun run f =
+  let
+    val (x, seed') = f (! seed);
+    val _ = seed := seed'
+  in x end;
+
+end;
+
+end;
+*}
+
+hide (open) type seed
+hide (open) const inc_shift minus_shift log "next" seed_invariant split_seed
+  iterate range select pick select_weight select_default
+hide (open) fact log_def
+
+no_notation fcomp (infixl "o>" 60)
+no_notation scomp (infixl "o\<rightarrow>" 60)
+
+end
+
--- a/src/HOL/Set.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Set.thy	Tue May 19 14:13:37 2009 +0100
@@ -1034,6 +1034,33 @@
 
 subsubsection {* Set reasoning tools *}
 
+text{* Elimination of @{text"{x. \<dots> & x=t & \<dots>}"}. *}
+
+lemma Collect_conv_if: "{x. x=a & P x} = (if P a then {a} else {})"
+by auto
+
+lemma Collect_conv_if2: "{x. a=x & P x} = (if P a then {a} else {})"
+by auto
+
+text {*
+Simproc for pulling @{text "x=t"} in @{text "{x. \<dots> & x=t & \<dots>}"}
+to the front (and similarly for @{text "t=x"}):
+*}
+
+ML{*
+  local
+    val Coll_perm_tac = rtac @{thm Collect_cong} 1 THEN rtac @{thm iffI} 1 THEN
+    ALLGOALS(EVERY'[REPEAT_DETERM o (etac @{thm conjE}),
+                    DEPTH_SOLVE_1 o (ares_tac [@{thm conjI}])])
+  in
+    val defColl_regroup = Simplifier.simproc (the_context ())
+      "defined Collect" ["{x. P x & Q x}"]
+      (Quantifier1.rearrange_Coll Coll_perm_tac)
+  end;
+
+  Addsimprocs [defColl_regroup];
+*}
+
 text {*
   Rewrite rules for boolean case-splitting: faster than @{text
   "split_if [split]"}.
--- a/src/HOL/String.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/String.thy	Tue May 19 14:13:37 2009 +0100
@@ -55,7 +55,7 @@
    (fn n => fn m => Drule.instantiate' [] [SOME n, SOME m] @{thm nibble_pair_of_char.simps})
       nibbles nibbles;
 in
-  PureThy.note_thmss Thm.lemmaK [((Binding.name "nibble_pair_of_char_simps", []), [(thms, [])])]
+  PureThy.note_thmss Thm.definitionK [((Binding.name "nibble_pair_of_char_simps", []), [(thms, [])])]
   #-> (fn [(_, thms)] => fold_rev Code.add_eqn thms)
 end
 *}
--- a/src/HOL/Tools/datatype_codegen.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Tools/datatype_codegen.ML	Tue May 19 14:13:37 2009 +0100
@@ -6,7 +6,7 @@
 
 signature DATATYPE_CODEGEN =
 sig
-  val mk_eq: theory -> string -> thm list
+  val mk_eq_eqns: theory -> string -> (thm * bool) list
   val mk_case_cert: theory -> string -> thm
   val setup: theory -> theory
 end;
@@ -309,18 +309,6 @@
 
 (** generic code generator **)
 
-(* specification *)
-
-fun add_datatype_spec vs dtco cos thy =
-  let
-    val cs = map (fn (c, tys) => (c, tys ---> Type (dtco, map TFree vs))) cos;
-  in
-    thy
-    |> try (Code.add_datatype cs)
-    |> the_default thy
-  end;
-
-
 (* case certificates *)
 
 fun mk_case_cert thy tyco =
@@ -354,88 +342,40 @@
     |> Thm.varifyT
   end;
 
-fun add_datatype_cases dtco thy =
-  let
-    val {case_rewrites, ...} = DatatypePackage.the_datatype thy dtco;
-    val cert = mk_case_cert thy dtco;
-    fun add_case_liberal thy = thy
-      |> try (Code.add_case cert)
-      |> the_default thy;
-  in
-    thy
-    |> add_case_liberal
-    |> fold_rev Code.add_default_eqn case_rewrites
-  end;
-
 
 (* equality *)
 
-local
-
-val not_sym = @{thm HOL.not_sym};
-val not_false_true = iffD2 OF [nth @{thms HOL.simp_thms} 7, TrueI];
-val refl = @{thm refl};
-val eqTrueI = @{thm eqTrueI};
-
-fun mk_distinct cos =
-  let
-    fun sym_product [] = []
-      | sym_product (x::xs) = map (pair x) xs @ sym_product xs;
-    fun mk_co_args (co, tys) ctxt =
-      let
-        val names = Name.invents ctxt "a" (length tys);
-        val ctxt' = fold Name.declare names ctxt;
-        val vs = map2 (curry Free) names tys;
-      in (vs, ctxt') end;
-    fun mk_dist ((co1, tys1), (co2, tys2)) =
-      let
-        val ((xs1, xs2), _) = Name.context
-          |> mk_co_args (co1, tys1)
-          ||>> mk_co_args (co2, tys2);
-        val prem = HOLogic.mk_eq
-          (list_comb (co1, xs1), list_comb (co2, xs2));
-        val t = HOLogic.mk_not prem;
-      in HOLogic.mk_Trueprop t end;
-  in map mk_dist (sym_product cos) end;
-
-in
-
-fun mk_eq thy dtco =
+fun mk_eq_eqns thy dtco =
   let
-    val (vs, cs) = DatatypePackage.the_datatype_spec thy dtco;
-    fun mk_triv_inject co =
-      let
-        val ct' = Thm.cterm_of thy
-          (Const (co, Type (dtco, map (fn (v, sort) => TVar ((v, 0), sort)) vs)))
-        val cty' = Thm.ctyp_of_term ct';
-        val SOME (ct, cty) = fold_aterms (fn Var (v, ty) =>
-          (K o SOME) (Thm.cterm_of thy (Var (v, Thm.typ_of cty')), Thm.ctyp_of thy ty) | _ => I)
-          (Thm.prop_of refl) NONE;
-      in eqTrueI OF [Thm.instantiate ([(cty, cty')], [(ct, ct')]) refl] end;
-    val inject1 = map_filter (fn (co, []) => SOME (mk_triv_inject co) | _ => NONE) cs
-    val inject2 = (#inject o DatatypePackage.the_datatype thy) dtco;
-    val ctxt = ProofContext.init thy;
-    val simpset = Simplifier.context ctxt
-      (Simplifier.empty_ss addsimprocs [DatatypePackage.distinct_simproc]);
-    val cos = map (fn (co, tys) =>
-        (Const (co, tys ---> Type (dtco, map TFree vs)), tys)) cs;
-    val tac = ALLGOALS (simp_tac simpset)
-      THEN ALLGOALS (ProofContext.fact_tac [not_false_true, TrueI]);
-    val distinct =
-      mk_distinct cos
-      |> map (fn t => Goal.prove_global thy [] [] t (K tac))
-      |> (fn thms => thms @ map (fn thm => not_sym OF [thm]) thms)
-  in inject1 @ inject2 @ distinct end;
+    val (vs, cos) = DatatypePackage.the_datatype_spec thy dtco;
+    val { descr, index, inject = inject_thms, ... } = DatatypePackage.the_datatype thy dtco;
+    val ty = Type (dtco, map TFree vs);
+    fun mk_eq (t1, t2) = Const (@{const_name eq_class.eq}, ty --> ty --> HOLogic.boolT)
+      $ t1 $ t2;
+    fun true_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.true_const);
+    fun false_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.false_const);
+    val triv_injects = map_filter
+     (fn (c, []) => SOME (HOLogic.mk_Trueprop (true_eq (Const (c, ty), Const (c, ty))))
+       | _ => NONE) cos;
+    fun prep_inject (trueprop $ (equiv $ (_ $ t1 $ t2) $ rhs)) =
+      trueprop $ (equiv $ mk_eq (t1, t2) $ rhs);
+    val injects = map prep_inject (nth (DatatypeProp.make_injs [descr] vs) index);
+    fun prep_distinct (trueprop $ (not $ (_ $ t1 $ t2))) =
+      [trueprop $ false_eq (t1, t2), trueprop $ false_eq (t2, t1)];
+    val distincts = maps prep_distinct (snd (nth (DatatypeProp.make_distincts [descr] vs) index));
+    val refl = HOLogic.mk_Trueprop (true_eq (Free ("x", ty), Free ("x", ty)));
+    val simpset = Simplifier.context (ProofContext.init thy) (HOL_basic_ss
+      addsimps (map Simpdata.mk_eq (@{thm eq} :: @{thm eq_True} :: inject_thms))
+      addsimprocs [DatatypePackage.distinct_simproc]);
+    fun prove prop = Goal.prove_global thy [] [] prop (K (ALLGOALS (simp_tac simpset)))
+      |> Simpdata.mk_eq;
+  in map (rpair true o prove) (triv_injects @ injects @ distincts) @ [(prove refl, false)] end;
 
-end;
-
-fun add_datatypes_equality vs dtcos thy =
+fun add_equality vs dtcos thy =
   let
-    val vs' = (map o apsnd)
-      (curry (Sorts.inter_sort (Sign.classes_of thy)) [HOLogic.class_eq]) vs;
     fun add_def dtco lthy =
       let
-        val ty = Type (dtco, map TFree vs');
+        val ty = Type (dtco, map TFree vs);
         fun mk_side const_name = Const (const_name, ty --> ty --> HOLogic.boolT)
           $ Free ("x", ty) $ Free ("y", ty);
         val def = HOLogic.mk_Trueprop (HOLogic.mk_eq
@@ -448,52 +388,60 @@
       in (thm', lthy') end;
     fun tac thms = Class.intro_classes_tac []
       THEN ALLGOALS (ProofContext.fact_tac thms);
-    fun mk_eq' thy dtco = mk_eq thy dtco
-      |> map (Code_Unit.constrain_thm thy [HOLogic.class_eq])
-      |> map Simpdata.mk_eq
-      |> map (MetaSimplifier.rewrite_rule [Thm.transfer thy @{thm equals_eq}])
-      |> map (AxClass.unoverload thy);
     fun add_eq_thms dtco thy =
       let
-        val ty = Type (dtco, map TFree vs');
+        val const = AxClass.param_of_inst thy (@{const_name eq_class.eq}, dtco);
         val thy_ref = Theory.check_thy thy;
-        val const = AxClass.param_of_inst thy (@{const_name eq_class.eq}, dtco);
-        val eq_refl = @{thm HOL.eq_refl}
-          |> Thm.instantiate
-              ([pairself (Thm.ctyp_of thy) (TVar (("'a", 0), @{sort eq}), Logic.varifyT ty)], [])
-          |> Simpdata.mk_eq
-          |> AxClass.unoverload thy;
-        fun mk_thms () = (eq_refl, false)
-          :: rev (map (rpair true) (mk_eq' (Theory.deref thy_ref) dtco));
+        fun mk_thms () = rev ((mk_eq_eqns (Theory.deref thy_ref) dtco));
       in
         Code.add_eqnl (const, Lazy.lazy mk_thms) thy
       end;
   in
     thy
-    |> TheoryTarget.instantiation (dtcos, vs', [HOLogic.class_eq])
+    |> TheoryTarget.instantiation (dtcos, vs, [HOLogic.class_eq])
     |> fold_map add_def dtcos
-    |-> (fn thms => Class.prove_instantiation_instance (K (tac thms))
-    #> LocalTheory.exit_global
-    #> fold Code.del_eqn thms
-    #> fold add_eq_thms dtcos)
+    |-> (fn def_thms => Class.prove_instantiation_exit_result (map o Morphism.thm)
+         (fn _ => fn def_thms => tac def_thms) def_thms)
+    |-> (fn def_thms => fold Code.del_eqn def_thms)
+    |> fold add_eq_thms dtcos
+  end;
+
+
+(* liberal addition of code data for datatypes *)
+
+fun mk_constr_consts thy vs dtco cos =
+  let
+    val cs = map (fn (c, tys) => (c, tys ---> Type (dtco, map TFree vs))) cos;
+    val cs' = map (fn c_ty as (_, ty) => (AxClass.unoverload_const thy c_ty, ty)) cs;
+  in if is_some (try (Code.constrset_of_consts thy) cs')
+    then SOME cs
+    else NONE
   end;
 
+fun add_all_code dtcos thy =
+  let
+    val (vs :: _, coss) = (split_list o map (DatatypePackage.the_datatype_spec thy)) dtcos;
+    val any_css = map2 (mk_constr_consts thy vs) dtcos coss;
+    val css = if exists is_none any_css then []
+      else map_filter I any_css;
+    val case_rewrites = maps (#case_rewrites o DatatypePackage.the_datatype thy) dtcos;
+    val certs = map (mk_case_cert thy) dtcos;
+  in
+    if null css then thy
+    else thy
+      |> fold Code.add_datatype css
+      |> fold_rev Code.add_default_eqn case_rewrites
+      |> fold Code.add_case certs
+      |> add_equality vs dtcos
+   end;
+
+
 
 (** theory setup **)
 
-fun add_datatype_code dtcos thy =
-  let
-    val (vs :: _, coss) = (split_list o map (DatatypePackage.the_datatype_spec thy)) dtcos;
-  in
-    thy
-    |> fold2 (add_datatype_spec vs) dtcos coss
-    |> fold add_datatype_cases dtcos
-    |> add_datatypes_equality vs dtcos
-  end;
-
 val setup = 
   add_codegen "datatype" datatype_codegen
   #> add_tycodegen "datatype" datatype_tycodegen
-  #> DatatypePackage.interpretation add_datatype_code
+  #> DatatypePackage.interpretation add_all_code
 
 end;
--- a/src/HOL/Tools/function_package/fundef_package.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Tools/function_package/fundef_package.ML	Tue May 19 14:13:37 2009 +0100
@@ -37,14 +37,15 @@
 val simp_attribs = map (Attrib.internal o K)
     [Simplifier.simp_add,
      Code.add_default_eqn_attribute,
-     Nitpick_Const_Simp_Thms.add]
+     Nitpick_Const_Simp_Thms.add,
+     Quickcheck_RecFun_Simp_Thms.add]
 
 val psimp_attribs = map (Attrib.internal o K)
     [Simplifier.simp_add,
      Nitpick_Const_Psimp_Thms.add]
 
 fun note_theorem ((name, atts), ths) =
-  LocalTheory.note Thm.theoremK ((Binding.qualified_name name, atts), ths)
+  LocalTheory.note Thm.generatedK ((Binding.qualified_name name, atts), ths)
 
 fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_"
 
@@ -55,7 +56,7 @@
                    |> map (apfst (apfst extra_qualify))
 
       val (saved_spec_simps, lthy) =
-        fold_map (LocalTheory.note Thm.theoremK) spec lthy
+        fold_map (LocalTheory.note Thm.generatedK) spec lthy
 
       val saved_simps = flat (map snd saved_spec_simps)
       val simps_by_f = sort saved_simps
--- a/src/HOL/Tools/hologic.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Tools/hologic.ML	Tue May 19 14:13:37 2009 +0100
@@ -119,6 +119,11 @@
   val message_stringT: typ
   val mk_message_string: string -> term
   val dest_message_string: term -> string
+  val mk_typerep: typ -> term
+  val mk_term_of: typ -> term -> term
+  val reflect_term: term -> term
+  val mk_return: typ -> typ -> term -> term
+  val mk_ST: ((term * typ) * (string * typ) option)  list -> term -> typ -> typ option * typ -> term
 end;
 
 structure HOLogic: HOLOGIC =
@@ -591,4 +596,41 @@
       dest_string t
   | dest_message_string t = raise TERM ("dest_message_string", [t]);
 
+
+(* typerep and term *)
+
+val typerepT = Type ("Typerep.typerep", []);
+
+fun mk_typerep T = Const ("Typerep.typerep_class.typerep",
+  Term.itselfT T --> typerepT) $ Logic.mk_type T;
+
+val termT = Type ("Code_Eval.term", []);
+
+fun mk_term_of T t = Const ("Code_Eval.term_of_class.term_of", T --> termT) $ t;
+
+fun reflect_term (Const (c, T)) =
+      Const ("Code_Eval.Const", message_stringT --> typerepT --> termT)
+        $ mk_message_string c $ mk_typerep T
+  | reflect_term (t1 $ t2) =
+      Const ("Code_Eval.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;
+
+
+(* open state monads *)
+
+fun mk_return T U x = pair_const T U $ x;
+
+fun mk_ST clauses t U (someT, V) =
+  let
+    val R = case someT of SOME T => mk_prodT (T, V) | NONE => V
+    fun mk_clause ((t, U), SOME (v, T)) (t', U') =
+          (Const ("Product_Type.scomp", (U --> mk_prodT (T, U')) --> (T --> U' --> R) --> U --> R)
+            $ t $ lambda (Free (v, T)) t', U)
+      | mk_clause ((t, U), NONE) (t', U') =
+          (Const ("Product_Type.fcomp", (U --> U') --> (U' --> R) --> U --> R)
+            $ t $ t', U)
+  in fold_rev mk_clause clauses (t, U) |> fst end;
+
 end;
--- a/src/HOL/Tools/inductive_package.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Tools/inductive_package.ML	Tue May 19 14:13:37 2009 +0100
@@ -454,7 +454,7 @@
     val facts = args |> map (fn ((a, atts), props) =>
       ((a, map (prep_att thy) atts),
         map (Thm.no_attributes o single o mk_cases lthy o prep_prop lthy) props));
-  in lthy |> LocalTheory.notes Thm.theoremK facts |>> map snd end;
+  in lthy |> LocalTheory.notes Thm.generatedK facts |>> map snd end;
 
 val inductive_cases = gen_inductive_cases Attrib.intern_src Syntax.read_prop;
 val inductive_cases_i = gen_inductive_cases (K I) Syntax.check_prop;
@@ -849,7 +849,7 @@
       |> Specification.read_spec (cnames_syn @ pnames_syn) intro_srcs;
     val (cs, ps) = chop (length cnames_syn) vars;
     val monos = Attrib.eval_thms lthy raw_monos;
-    val flags = {quiet_mode = false, verbose = verbose, kind = Thm.theoremK,
+    val flags = {quiet_mode = false, verbose = verbose, kind = Thm.generatedK,
       alt_name = Binding.empty, coind = coind, no_elim = false, no_ind = false,
       skip_mono = false, fork_mono = not int};
   in
--- a/src/HOL/Tools/inductive_realizer.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Tools/inductive_realizer.ML	Tue May 19 14:13:37 2009 +0100
@@ -349,7 +349,7 @@
 
     val (ind_info, thy3') = thy2 |>
       InductivePackage.add_inductive_global (serial_string ())
-        {quiet_mode = false, verbose = false, kind = Thm.theoremK, alt_name = Binding.empty,
+        {quiet_mode = false, verbose = false, kind = Thm.generatedK, alt_name = Binding.empty,
           coind = false, no_elim = false, no_ind = false, skip_mono = false, fork_mono = false}
         rlzpreds rlzparams (map (fn (rintr, intr) =>
           ((Binding.name (Long_Name.base_name (name_of_thm intr)), []),
--- a/src/HOL/Tools/primrec_package.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Tools/primrec_package.ML	Tue May 19 14:13:37 2009 +0100
@@ -247,14 +247,14 @@
     val spec' = (map o apfst)
       (fn (b, attrs) => (qualify b, Code.add_default_eqn_attrib :: attrs)) spec;
     val simp_atts = map (Attrib.internal o K)
-      [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add];
+      [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add, Quickcheck_RecFun_Simp_Thms.add];
   in
     lthy
     |> set_group ? LocalTheory.set_group (serial_string ())
     |> fold_map (LocalTheory.define Thm.definitionK o make_def lthy fixes fs) defs
     |-> (fn defs => `(fn ctxt => prove_spec ctxt names1 rec_rewrites defs spec'))
-    |-> (fn simps => fold_map (LocalTheory.note Thm.theoremK) simps)
-    |-> (fn simps' => LocalTheory.note Thm.theoremK
+    |-> (fn simps => fold_map (LocalTheory.note Thm.generatedK) simps)
+    |-> (fn simps' => LocalTheory.note Thm.generatedK
           ((qualify (Binding.qualified_name "simps"), simp_atts), maps snd simps'))
     |>> snd
   end handle PrimrecError (msg, some_eqn) =>
--- a/src/HOL/Tools/recdef_package.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Tools/recdef_package.ML	Tue May 19 14:13:37 2009 +0100
@@ -208,7 +208,7 @@
                congs wfs name R eqs;
     val rules = (map o map) fst (partition_eq (eq_snd (op = : int * int -> bool)) rules_idx);
     val simp_att = if null tcs then [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add,
-      Code.add_default_eqn_attribute] else [];
+      Code.add_default_eqn_attribute, Quickcheck_RecFun_Simp_Thms.add] else [];
 
     val ((simps' :: rules', [induct']), thy) =
       thy
--- a/src/HOL/Tools/recfun_codegen.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Tools/recfun_codegen.ML	Tue May 19 14:13:37 2009 +0100
@@ -26,10 +26,10 @@
 fun add_thm NONE thm thy = Code.add_eqn thm thy
   | add_thm (SOME module_name) thm thy =
       let
-        val (thm', _) = Code_Unit.mk_eqn thy (K false) (thm, true)
+        val (thm', _) = Code.mk_eqn thy (K false) (thm, true)
       in
         thy
-        |> ModuleData.map (Symtab.update (Code_Unit.const_eqn thy thm', module_name))
+        |> ModuleData.map (Symtab.update (Code.const_eqn thy thm', module_name))
         |> Code.add_eqn thm'
       end;
 
@@ -44,10 +44,10 @@
 fun expand_eta thy [] = []
   | expand_eta thy (thms as thm :: _) =
       let
-        val (_, ty) = Code_Unit.const_typ_eqn thm;
+        val (_, ty) = Code.const_typ_eqn thm;
       in if null (Term.add_tvarsT ty []) orelse (null o fst o strip_type) ty
         then thms
-        else map (Code_Unit.expand_eta thy 1) thms
+        else map (Code.expand_eta thy 1) thms
       end;
 
 fun retrieve_equations thy (c, T) = if c = @{const_name "op ="} then NONE else
@@ -58,7 +58,7 @@
       |> map_filter (fn (thm, linear) => if linear then SOME thm else NONE)
       |> expand_eta thy
       |> map_filter (meta_eq_to_obj_eq thy)
-      |> Code_Unit.norm_varnames thy
+      |> Code.norm_varnames thy
       |> map (rpair opt_name)
   in if null thms then NONE else SOME thms end;
 
--- a/src/HOL/Tools/record_package.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Tools/record_package.ML	Tue May 19 14:13:37 2009 +0100
@@ -1464,7 +1464,7 @@
     val tname = Binding.name (Long_Name.base_name name);
   in
     thy
-    |> TypecopyPackage.add_typecopy (Binding.suffix_name ext_typeN tname, alphas) repT NONE
+    |> TypecopyPackage.typecopy (Binding.suffix_name ext_typeN tname, alphas) repT NONE
     |-> (fn (name, _) => `(fn thy => get_thms thy name))
   end;
 
--- a/src/HOL/Tools/typecopy_package.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Tools/typecopy_package.ML	Tue May 19 14:13:37 2009 +0100
@@ -14,12 +14,12 @@
     proj: string * typ,
     proj_def: thm
   }
-  val add_typecopy: binding * string list -> typ -> (binding * binding) option
+  val typecopy: binding * string list -> typ -> (binding * binding) option
     -> theory -> (string * info) * theory
   val get_typecopies: theory -> string list
   val get_info: theory -> string -> info option
   val interpretation: (string -> theory -> theory) -> theory -> theory
-  val add_typecopy_default_code: string -> theory -> theory
+  val add_default_code: string -> theory -> theory
   val print_typecopies: theory -> unit
   val setup: theory -> theory
 end;
@@ -71,7 +71,10 @@
 structure TypecopyInterpretation = InterpretationFun(type T = string val eq = op =);
 val interpretation = TypecopyInterpretation.interpretation;
 
-fun add_typecopy (raw_tyco, raw_vs) raw_ty constr_proj thy =
+
+(* declaring typecopies *)
+
+fun typecopy (raw_tyco, raw_vs) raw_ty constr_proj thy =
   let
     val ty = Sign.certify_typ thy raw_ty;
     val vs =
@@ -108,28 +111,26 @@
   end;
 
 
-(* code generator setup *)
+(* default code setup *)
 
-fun add_typecopy_default_code tyco thy =
+fun add_default_code tyco thy =
   let
     val SOME { constr = constr_name, proj = (proj, _), proj_def = proj_eq, vs = raw_vs,
-      typ = raw_ty_rep, ... } =  get_info thy tyco;
-    val inst_meet = Sorts.meet_sort_typ (Sign.classes_of thy)
-      (Logic.varifyT raw_ty_rep, [HOLogic.class_eq]) handle Sorts.CLASS_ERROR _ => I;
-    val ty_inst = Logic.unvarifyT o inst_meet o Logic.varifyT;
-    val vs = (map dest_TFree o snd o dest_Type o ty_inst)
-      (Type (tyco, map TFree raw_vs));
-    val ty_rep = ty_inst raw_ty_rep;
+      typ = ty_rep, ... } =  get_info thy tyco;
     val SOME { Rep_inject = proj_inject, ... } = TypedefPackage.get_info thy tyco;
-    val ty_constr = Logic.unvarifyT (Sign.the_const_type thy constr_name);
-    val constr = (constr_name, ty_constr)
+    val constr = (constr_name, Logic.unvarifyT (Sign.the_const_type thy constr_name));
+    val vs = (map dest_TFree o snd o dest_Type) (Type (tyco, map TFree raw_vs));
     val ty = Type (tyco, map TFree vs);
-    fun mk_eq ty t_x t_y = Const (@{const_name eq_class.eq}, ty --> ty --> HOLogic.boolT)
+    val proj = Const (proj, ty --> ty_rep);
+    val (t_x, t_y) = (Free ("x", ty), Free ("y", ty));
+    val eq_lhs = Const (@{const_name eq_class.eq}, ty --> ty --> HOLogic.boolT)
       $ t_x $ t_y;
-    fun mk_proj t = Const (proj, ty --> ty_rep) $ t;
-    val (t_x, t_y) = (Free ("x", ty), Free ("y", ty));
-    val def_eq = (HOLogic.mk_Trueprop o HOLogic.mk_eq)
-      (mk_eq ty t_x t_y, HOLogic.mk_eq (mk_proj t_x, mk_proj t_y));
+    val eq_rhs = HOLogic.mk_eq (proj $ t_x, proj $ t_y);
+    val eq = (HOLogic.mk_Trueprop o HOLogic.mk_eq) (eq_lhs, eq_rhs);
+    fun tac eq_thm = Class.intro_classes_tac []
+      THEN (Simplifier.rewrite_goals_tac
+        (map Simpdata.mk_eq [eq_thm, @{thm eq}, proj_inject]))
+          THEN ALLGOALS (rtac @{thm refl});
     fun mk_eq_refl thy = @{thm HOL.eq_refl}
       |> Thm.instantiate
          ([pairself (Thm.ctyp_of thy) (TVar (("'a", 0), @{sort eq}), Logic.varifyT ty)], [])
@@ -139,22 +140,18 @@
     |> Code.add_datatype [constr]
     |> Code.add_eqn proj_eq
     |> TheoryTarget.instantiation ([tyco], vs, [HOLogic.class_eq])
-    |> `(fn lthy => Syntax.check_term lthy def_eq)
-    |-> (fn def_eq => Specification.definition
-         (NONE, (Attrib.empty_binding, def_eq)))
-    |-> (fn (_, (_, def_thm)) =>
+    |> `(fn lthy => Syntax.check_term lthy eq)
+    |-> (fn eq => Specification.definition
+         (NONE, (Attrib.empty_binding, eq)))
+    |-> (fn (_, (_, eq_thm)) =>
        Class.prove_instantiation_exit_result Morphism.thm
-         (fn _ => fn def_thm => Class.intro_classes_tac []
-            THEN (Simplifier.rewrite_goals_tac
-              (map Simpdata.mk_eq [def_thm, @{thm eq}, proj_inject]))
-                THEN ALLGOALS (rtac @{thm refl})) def_thm)
-    |-> (fn def_thm => Code.add_eqn def_thm)
-    |> `(fn thy => mk_eq_refl thy)
-    |-> (fn refl_thm => Code.add_nbe_eqn refl_thm)
+         (fn _ => fn eq_thm => tac eq_thm) eq_thm)
+    |-> (fn eq_thm => Code.add_eqn eq_thm)
+    |> (fn thy => Code.add_nbe_eqn (mk_eq_refl thy) thy)
   end;
 
 val setup =
   TypecopyInterpretation.init
-  #> interpretation add_typecopy_default_code
+  #> interpretation add_default_code
 
 end;
--- a/src/HOL/Transcendental.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Transcendental.thy	Tue May 19 14:13:37 2009 +0100
@@ -173,7 +173,7 @@
     have "(\<lambda> n. if even n then f (n div 2) else 0) sums y"
       unfolding sums_def setsum_shift_lb_Suc0_0_upt[where f="?s", OF `?s 0 = 0`, symmetric]
                 image_Suc_atLeastLessThan[symmetric] setsum_reindex[OF inj_Suc, unfolded comp_def]
-                even_nat_Suc Suc_m1 if_eq .
+                even_Suc Suc_m1 if_eq .
   } from sums_add[OF g_sums this]
   show ?thesis unfolding if_sum .
 qed
--- a/src/HOL/Typerep.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/Typerep.thy	Tue May 19 14:13:37 2009 +0100
@@ -35,28 +35,18 @@
 end
 *}
 
-ML {*
-structure Typerep =
-struct
+setup {*
+let
 
-fun mk f (Type (tyco, tys)) =
-      @{term Typerep} $ HOLogic.mk_message_string tyco
-        $ HOLogic.mk_list @{typ typerep} (map (mk f) tys)
-  | mk f (TFree v) =
-      f v;
-
-fun typerep ty =
-  Const (@{const_name typerep}, Term.itselfT ty --> @{typ typerep})
-    $ Logic.mk_type ty;
-
-fun add_def tyco thy =
+fun add_typerep tyco thy =
   let
     val sorts = replicate (Sign.arity_number thy tyco) @{sort typerep};
     val vs = Name.names Name.context "'a" sorts;
     val ty = Type (tyco, map TFree vs);
     val lhs = Const (@{const_name typerep}, Term.itselfT ty --> @{typ typerep})
       $ Free ("T", Term.itselfT ty);
-    val rhs = mk (typerep o TFree) ty;
+    val rhs = @{term Typerep} $ HOLogic.mk_message_string tyco
+      $ HOLogic.mk_list @{typ typerep} (map (HOLogic.mk_typerep o TFree) vs);
     val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
   in
     thy
@@ -64,23 +54,20 @@
     |> `(fn lthy => Syntax.check_term lthy eq)
     |-> (fn eq => Specification.definition (NONE, (Attrib.empty_binding, eq)))
     |> snd
-    |> Class.prove_instantiation_instance (K (Class.intro_classes_tac []))
-    |> LocalTheory.exit_global
+    |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
   end;
 
-fun perhaps_add_def tyco thy =
-  let
-    val inst = can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort typerep}
-  in if inst then thy else add_def tyco thy end;
+fun ensure_typerep tyco thy = if not (can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort typerep})
+  andalso can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort type}
+  then add_typerep tyco thy else thy;
+
+in
 
-end;
-*}
+add_typerep @{type_name fun}
+#> TypedefPackage.interpretation ensure_typerep
+#> Code.type_interpretation (ensure_typerep o fst)
 
-setup {*
-  Typerep.add_def @{type_name fun}
-  #> Typerep.add_def @{type_name itself}
-  #> Typerep.add_def @{type_name bool}
-  #> TypedefPackage.interpretation Typerep.perhaps_add_def
+end
 *}
 
 lemma [code]:
--- a/src/HOL/ex/Predicate_Compile.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/ex/Predicate_Compile.thy	Tue May 19 14:13:37 2009 +0100
@@ -22,9 +22,7 @@
 val pred_ref = ref (NONE : (unit -> term Predicate.pred) option);
 
 fun eval_pred thy t =
-  t 
-  |> Eval.mk_term_of (fastype_of t)
-  |> (fn t => Code_ML.eval NONE ("Predicate.pred_ref", pred_ref) @{code pred_map} thy t []);
+  Code_ML.eval NONE ("Predicate.pred_ref", pred_ref) @{code pred_map} thy (HOLogic.mk_term_of (fastype_of t) t) [];
 
 fun eval_pred_elems thy t T length =
   t |> eval_pred thy |> yieldn length |> fst |> HOLogic.mk_list T;
--- a/src/HOL/ex/Predicate_Compile_ex.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/ex/Predicate_Compile_ex.thy	Tue May 19 14:13:37 2009 +0100
@@ -12,6 +12,8 @@
 
 thm even.equation
 
+ML_val {* Predicate.yieldn 10 @{code even_0} *}
+
 
 inductive append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
     append_Nil: "append [] xs xs"
@@ -22,6 +24,8 @@
 
 thm append.equation
 
+ML_val {* Predicate.yieldn 10 (@{code append_3} [1, 2, 3]) *}
+
 
 inductive partition :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
   for f where
--- a/src/HOL/ex/Quickcheck_Generators.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/ex/Quickcheck_Generators.thy	Tue May 19 14:13:37 2009 +0100
@@ -8,27 +8,11 @@
 
 subsection {* Datatypes *}
 
-definition
-  collapse :: "('a \<Rightarrow> ('a \<Rightarrow> 'b \<times> 'a) \<times> 'a) \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a" where
+definition collapse :: "('a \<Rightarrow> ('a \<Rightarrow> 'b \<times> 'a) \<times> 'a) \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a" where
   "collapse f = (do g \<leftarrow> f; g done)"
 
-ML {*
-structure StateMonad =
-struct
-
-fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT);
-fun liftT' sT = sT --> sT;
-
-fun return T sT x = Const (@{const_name return}, T --> liftT T sT) $ x;
-
-fun scomp T1 T2 sT f g = Const (@{const_name scomp},
-  liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g;
-
-end;
-*}
-
 lemma random'_if:
-  fixes random' :: "index \<Rightarrow> index \<Rightarrow> seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> seed"
+  fixes random' :: "index \<Rightarrow> index \<Rightarrow> Random.seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> Random.seed"
   assumes "random' 0 j = (\<lambda>s. undefined)"
     and "\<And>i. random' (Suc_index i) j = rhs2 i"
   shows "random' i j s = (if i = 0 then undefined else rhs2 (i - 1) s)"
@@ -36,15 +20,18 @@
 
 setup {*
 let
+  fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT);
+  fun scomp T1 T2 sT f g = Const (@{const_name scomp},
+    liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g;
   exception REC of string;
   exception TYP of string;
   fun mk_collapse thy ty = Sign.mk_const thy
-    (@{const_name collapse}, [@{typ seed}, ty]);
+    (@{const_name collapse}, [@{typ Random.seed}, ty]);
   fun term_ty ty = HOLogic.mk_prodT (ty, @{typ "unit \<Rightarrow> term"});
   fun mk_split thy ty ty' = Sign.mk_const thy
-    (@{const_name split}, [ty, @{typ "unit \<Rightarrow> term"}, StateMonad.liftT (term_ty ty') @{typ seed}]);
+    (@{const_name split}, [ty, @{typ "unit \<Rightarrow> term"}, liftT (term_ty ty') @{typ Random.seed}]);
   fun mk_scomp_split thy ty ty' t t' =
-    StateMonad.scomp (term_ty ty) (term_ty ty') @{typ seed} t
+    scomp (term_ty ty) (term_ty ty') @{typ Random.seed} t
       (mk_split thy ty ty' $ Abs ("", ty, Abs ("", @{typ "unit \<Rightarrow> term"}, t')))
   fun mk_cons thy this_ty (c, args) =
     let
@@ -54,10 +41,10 @@
       val t_indices = map (curry ( op * ) 2) (length tys - 1 downto 0);
       val c_indices = map (curry ( op + ) 1) t_indices;
       val c_t = list_comb (c, map Bound c_indices);
-      val t_t = Abs ("", @{typ unit}, Eval.mk_term Free Typerep.typerep
+      val t_t = Abs ("", @{typ unit}, HOLogic.reflect_term
         (list_comb (c, map (fn k => Bound (k + 1)) t_indices))
         |> map_aterms (fn t as Bound _ => t $ @{term "()"} | t => t));
-      val return = StateMonad.return (term_ty this_ty) @{typ seed}
+      val return = HOLogic.mk_return (term_ty this_ty) @{typ Random.seed}
         (HOLogic.mk_prod (c_t, t_t));
       val t = fold_rev (fn ((ty, _), random) =>
         mk_scomp_split thy ty this_ty random)
@@ -67,21 +54,21 @@
   fun mk_conss thy ty [] = NONE
     | mk_conss thy ty [(_, t)] = SOME t
     | mk_conss thy ty ts = SOME (mk_collapse thy (term_ty ty) $
-          (Sign.mk_const thy (@{const_name select}, [StateMonad.liftT (term_ty ty) @{typ seed}]) $
-            HOLogic.mk_list (StateMonad.liftT (term_ty ty) @{typ seed}) (map snd ts)));
+          (Sign.mk_const thy (@{const_name Random.select}, [liftT (term_ty ty) @{typ Random.seed}]) $
+            HOLogic.mk_list (liftT (term_ty ty) @{typ Random.seed}) (map snd ts)));
   fun mk_clauses thy ty (tyco, (ts_rec, ts_atom)) = 
     let
       val SOME t_atom = mk_conss thy ty ts_atom;
     in case mk_conss thy ty ts_rec
      of SOME t_rec => mk_collapse thy (term_ty ty) $
-          (Sign.mk_const thy (@{const_name select_default}, [StateMonad.liftT (term_ty ty) @{typ seed}]) $
+          (Sign.mk_const thy (@{const_name Random.select_default}, [liftT (term_ty ty) @{typ Random.seed}]) $
              @{term "i\<Colon>index"} $ t_rec $ t_atom)
       | NONE => t_atom
     end;
   fun mk_random_eqs thy vs tycos =
     let
       val this_ty = Type (hd tycos, map TFree vs);
-      val this_ty' = StateMonad.liftT (term_ty this_ty) @{typ seed};
+      val this_ty' = liftT (term_ty this_ty) @{typ Random.seed};
       val random_name = Long_Name.base_name @{const_name random};
       val random'_name = random_name ^ "_" ^ Class.type_name (hd tycos) ^ "'";
       fun random ty = Sign.mk_const thy (@{const_name random}, [ty]);
@@ -100,8 +87,8 @@
         |> (map o apsnd) (List.partition fst)
         |> map (mk_clauses thy this_ty)
       val eqss = map ((apsnd o map) (HOLogic.mk_Trueprop o HOLogic.mk_eq) o (fn rhs => ((this_ty, random'), [
-          (random' $ @{term "0\<Colon>index"} $ @{term "j\<Colon>index"}, Abs ("s", @{typ seed},
-            Const (@{const_name undefined}, HOLogic.mk_prodT (term_ty this_ty, @{typ seed})))),
+          (random' $ @{term "0\<Colon>index"} $ @{term "j\<Colon>index"}, Abs ("s", @{typ Random.seed},
+            Const (@{const_name undefined}, HOLogic.mk_prodT (term_ty this_ty, @{typ Random.seed})))),
           (random' $ @{term "Suc_index i"} $ @{term "j\<Colon>index"}, rhs)
         ]))) rhss;
     in eqss end;
@@ -113,7 +100,7 @@
           val ((this_ty, random'), eqs') = singleton (mk_random_eqs thy vs) tyco;
           val eq = (HOLogic.mk_Trueprop o HOLogic.mk_eq)
             (Sign.mk_const thy (@{const_name random}, [this_ty]) $ @{term "i\<Colon>index"},
-               random' $ @{term "i\<Colon>index"} $ @{term "i\<Colon>index"})
+               random' $ @{term "max (i\<Colon>index) 1"} $ @{term "i\<Colon>index"})
           val del_func = Attrib.internal (fn _ => Thm.declaration_attribute
             (fn thm => Context.mapping (Code.del_eqn thm) I));
           fun add_code simps lthy =
@@ -146,32 +133,15 @@
         end
     | random_inst tycos thy = raise REC
         ("Will not generate random elements for mutual recursive type(s) " ^ commas (map quote tycos));
-  fun add_random_inst tycos thy = random_inst tycos thy
-     handle REC msg => (warning msg; thy)
-          | TYP msg => (warning msg; thy)
+  fun add_random_inst [@{type_name bool}] thy = thy
+    | add_random_inst [@{type_name nat}] thy = thy
+    | add_random_inst tycos thy = random_inst tycos thy
+        handle REC msg => (warning msg; thy)
+             | TYP msg => (warning msg; thy)
 in DatatypePackage.interpretation add_random_inst end
 *}
 
 
-subsection {* Type @{typ int} *}
-
-instantiation int :: random
-begin
-
-definition
-  "random n = (do
-     (b, _) \<leftarrow> random n;
-     (m, t) \<leftarrow> random n;
-     return (if b then (int m, \<lambda>u. Code_Eval.App (Code_Eval.Const (STR ''Int.int'') TYPEREP(nat \<Rightarrow> int)) (t ()))
-       else (- int m, \<lambda>u. Code_Eval.App (Code_Eval.Const (STR ''HOL.uminus_class.uminus'') TYPEREP(int \<Rightarrow> int))
-         (Code_Eval.App (Code_Eval.Const (STR ''Int.int'') TYPEREP(nat \<Rightarrow> int)) (t ()))))
-   done)"
-
-instance ..
-
-end
-
-
 subsection {* Examples *}
 
 theorem "map g (map f xs) = map (g o f) xs"
@@ -294,4 +264,8 @@
   quickcheck [generator = code]
   oops
 
+lemma "int (nat k) = k"
+  quickcheck [generator = code]
+  oops
+
 end
--- a/src/HOL/ex/ROOT.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/ex/ROOT.ML	Tue May 19 14:13:37 2009 +0100
@@ -11,7 +11,6 @@
   "Word",
   "Eval_Examples",
   "Quickcheck_Generators",
-  "Term_Of_Syntax",
   "Codegenerator",
   "Codegenerator_Pretty",
   "NormalForm",
--- a/src/HOL/ex/Term_Of_Syntax.thy	Tue May 19 14:13:23 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,60 +0,0 @@
-(*  Title:      HOL/ex/Term_Of_Syntax.thy
-    ID:         $Id$
-    Author:     Florian Haftmann, TU Muenchen
-*)
-
-header {* Input syntax for term_of functions *}
-
-theory Term_Of_Syntax
-imports Code_Eval
-begin
-
-setup {*
-  Sign.declare_const [] ((Binding.name "rterm_of", @{typ "'a \<Rightarrow> 'b"}), NoSyn)
-  #> snd
-*}
-
-notation (output)
-  rterm_of ("\<guillemotleft>_\<guillemotright>")
-
-locale rterm_syntax =
-  fixes rterm_of_syntax :: "'a \<Rightarrow> 'b" ("\<guillemotleft>_\<guillemotright>")
-
-parse_translation {*
-let
-  fun rterm_of_tr [t] = Lexicon.const @{const_name rterm_of} $ t
-    | rterm_of_tr ts = raise TERM ("rterm_of_tr", ts);
-in
-  [(Syntax.fixedN ^ "rterm_of_syntax", rterm_of_tr)]
-end
-*}
-
-setup {*
-let
-  val subst_rterm_of = Eval.mk_term
-    (fn (v, _) => error ("illegal free variable in term quotation: " ^ quote v))
-    (Typerep.mk (fn (v, sort) => Typerep.typerep (TFree (v, sort))));
-  fun subst_rterm_of' (Const (@{const_name rterm_of}, _), [t]) = subst_rterm_of t
-    | subst_rterm_of' (Const (@{const_name rterm_of}, _), _) =
-        error ("illegal number of arguments for " ^ quote @{const_name rterm_of})
-    | subst_rterm_of' (t, ts) = list_comb (t, map (subst_rterm_of' o strip_comb) ts);
-  fun subst_rterm_of'' t = 
-    let
-      val t' = subst_rterm_of' (strip_comb t);
-    in if t aconv t'
-      then NONE
-      else SOME t'
-    end;
-  fun check_rterm_of ts ctxt =
-    let
-      val ts' = map subst_rterm_of'' ts;
-    in if exists is_some ts'
-      then SOME (map2 the_default ts ts', ctxt)
-      else NONE
-    end;
-in
-  Context.theory_map (Syntax.add_term_check 0 "rterm_of" check_rterm_of)
-end;
-*}
-
-end
--- a/src/HOL/ex/predicate_compile.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOL/ex/predicate_compile.ML	Tue May 19 14:13:37 2009 +0100
@@ -18,8 +18,9 @@
   val code_pred_cmd: string -> Proof.context -> Proof.state
   val print_alternative_rules: theory -> theory (*FIXME diagnostic command?*)
   val do_proofs: bool ref
-  val pred_intros: theory -> string -> thm list
-  val get_nparams: theory -> string -> int
+  val pred_intros : theory -> string -> thm list
+  val get_nparams : theory -> string -> int
+  val pred_term_of : theory -> term -> term option
 end;
 
 structure Predicate_Compile : PREDICATE_COMPILE =
@@ -270,7 +271,7 @@
 datatype hmode = Mode of mode * int list * hmode option list; (*FIXME don't understand
   why there is another mode type!?*)
 
-fun modes_of modes t =
+fun modes_of_term modes t =
   let
     val ks = 1 upto length (binder_types (fastype_of t));
     val default = [Mode (([], ks), ks, [])];
@@ -288,7 +289,7 @@
           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 modes arg)))
+                  (fn Mode (_, js', _) => js=js') (modes_of_term modes arg)))
                     (iss ~~ args1)))
           end
         end)) (AList.lookup op = modes name)
@@ -317,13 +318,13 @@
             term_vs t subset vs andalso
             forall is_eqT dupTs
           end)
-            (modes_of modes t handle Option =>
+            (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 modes t handle Option =>
+            (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
@@ -1392,7 +1393,7 @@
     val (predname, _) = dest_Const pred
     fun after_qed [[th]] lthy'' =
       lthy''
-      |> LocalTheory.note Thm.theoremK
+      |> LocalTheory.note Thm.generatedK
            ((Binding.empty, [Attrib.internal (K add_elim_attrib)]), [th])
       |> snd
       |> LocalTheory.theory (prove_equation predname NONE)
@@ -1405,7 +1406,7 @@
 in
 
 val code_pred = generic_code_pred (K I);
-val code_pred_cmd = generic_code_pred Code_Unit.read_const
+val code_pred_cmd = generic_code_pred Code.read_const
 
 val setup =
   Attrib.setup @{binding code_ind_intros} (Scan.succeed (attrib add_intro_thm))
@@ -1426,4 +1427,27 @@
 - Naming of auxiliary rules necessary?
 *)
 
+(* transformation for code generation *)
+
+fun pred_term_of thy t = let
+   val (vars, body) = strip_abs t
+   val (pred, all_args) = strip_comb body
+   val (name, T) = dest_Const pred 
+   val (params, args) = chop (get_nparams thy name) all_args
+   val user_mode = flat (map_index
+      (fn (i, t) => case t of Bound j => if j < length vars then [] else [i+1] | _ => [i+1])
+        args)
+  val (inargs, _) = get_args user_mode args
+  val all_modes = Symtab.dest (#modes (IndCodegenData.get thy))
+  val modes = filter (fn Mode (_, is, _) => is = user_mode) (modes_of_term all_modes (list_comb (pred, params)))
+  fun compile m = list_comb (compile_expr thy all_modes (SOME m, list_comb (pred, params)), inargs)
+  in
+    case modes of
+      []  => (let val _ = error "No mode possible for this term" in NONE end)
+    | [m] => SOME (compile m)
+    | ms  => (let val _ = warning "Multiple modes possible for this term"
+        in SOME (compile (hd ms)) end)
+  end;
+
 end;
+
--- a/src/HOLCF/Algebraic.thy	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOLCF/Algebraic.thy	Tue May 19 14:13:37 2009 +0100
@@ -29,6 +29,11 @@
   thus "P i j" using step trans by (rule less_Suc_induct)
 qed
 
+definition
+  eventual_iterate :: "('a \<rightarrow> 'a::cpo) \<Rightarrow> ('a \<rightarrow> 'a)"
+where
+  "eventual_iterate f = eventual (\<lambda>n. iterate n\<cdot>f)"
+
 text {* A pre-deflation is like a deflation, but not idempotent. *}
 
 locale pre_deflation =
@@ -103,9 +108,10 @@
 abbreviation
   d :: "'a \<rightarrow> 'a"
 where
-  "d \<equiv> eventual (\<lambda>n. iterate n\<cdot>f)"
+  "d \<equiv> eventual_iterate f"
 
 lemma MOST_d: "MOST n. P (iterate n\<cdot>f) \<Longrightarrow> P d"
+unfolding eventual_iterate_def
 using eventually_constant_iterate by (rule MOST_eventual)
 
 lemma f_d: "f\<cdot>(d\<cdot>x) = d\<cdot>x"
@@ -134,6 +140,7 @@
 proof
   fix x :: 'a
   have "d \<in> range (\<lambda>n. iterate n\<cdot>f)"
+    unfolding eventual_iterate_def
     using eventually_constant_iterate
     by (rule eventual_mem_range)
   then obtain n where n: "d = iterate n\<cdot>f" ..
@@ -153,9 +160,17 @@
     by (simp add: d_fixed_iff)
 qed
 
+lemma deflation_d: "deflation d"
+using finite_deflation_d
+by (rule finite_deflation_imp_deflation)
+
 end
 
-lemma pre_deflation_d_f:
+lemma finite_deflation_eventual_iterate:
+  "pre_deflation d \<Longrightarrow> finite_deflation (eventual_iterate d)"
+by (rule pre_deflation.finite_deflation_d)
+
+lemma pre_deflation_oo:
   assumes "finite_deflation d"
   assumes f: "\<And>x. f\<cdot>x \<sqsubseteq> x"
   shows "pre_deflation (d oo f)"
@@ -171,13 +186,13 @@
 lemma eventual_iterate_oo_fixed_iff:
   assumes "finite_deflation d"
   assumes f: "\<And>x. f\<cdot>x \<sqsubseteq> x"
-  shows "eventual (\<lambda>n. iterate n\<cdot>(d oo f))\<cdot>x = x \<longleftrightarrow> d\<cdot>x = x \<and> f\<cdot>x = x"
+  shows "eventual_iterate (d oo f)\<cdot>x = x \<longleftrightarrow> d\<cdot>x = x \<and> f\<cdot>x = x"
 proof -
   interpret d: finite_deflation d by fact
   let ?e = "d oo f"
   interpret e: pre_deflation "d oo f"
     using `finite_deflation d` f
-    by (rule pre_deflation_d_f)
+    by (rule pre_deflation_oo)
   let ?g = "eventual (\<lambda>n. iterate n\<cdot>?e)"
   show ?thesis
     apply (subst e.d_fixed_iff)
@@ -192,6 +207,94 @@
     done
 qed
 
+lemma eventual_mono:
+  assumes A: "eventually_constant A"
+  assumes B: "eventually_constant B"
+  assumes below: "\<And>n. A n \<sqsubseteq> B n"
+  shows "eventual A \<sqsubseteq> eventual B"
+proof -
+  from A have "MOST n. A n = eventual A"
+    by (rule MOST_eq_eventual)
+  then have "MOST n. eventual A \<sqsubseteq> B n"
+    by (rule MOST_mono) (erule subst, rule below)
+  with B show "eventual A \<sqsubseteq> eventual B"
+    by (rule MOST_eventual)
+qed
+
+lemma eventual_iterate_mono:
+  assumes f: "pre_deflation f" and g: "pre_deflation g" and "f \<sqsubseteq> g"
+  shows "eventual_iterate f \<sqsubseteq> eventual_iterate g"
+unfolding eventual_iterate_def
+apply (rule eventual_mono)
+apply (rule pre_deflation.eventually_constant_iterate [OF f])
+apply (rule pre_deflation.eventually_constant_iterate [OF g])
+apply (rule monofun_cfun_arg [OF `f \<sqsubseteq> g`])
+done
+
+lemma cont2cont_eventual_iterate_oo:
+  assumes d: "finite_deflation d"
+  assumes cont: "cont f" and below: "\<And>x y. f x\<cdot>y \<sqsubseteq> y"
+  shows "cont (\<lambda>x. eventual_iterate (d oo f x))"
+    (is "cont ?e")
+proof (rule contI2)
+  show "monofun ?e"
+    apply (rule monofunI)
+    apply (rule eventual_iterate_mono)
+    apply (rule pre_deflation_oo [OF d below])
+    apply (rule pre_deflation_oo [OF d below])
+    apply (rule monofun_cfun_arg)
+    apply (erule cont2monofunE [OF cont])
+    done
+next
+  fix Y :: "nat \<Rightarrow> 'b"
+  assume Y: "chain Y"
+  with cont have fY: "chain (\<lambda>i. f (Y i))"
+    by (rule ch2ch_cont)
+  assume eY: "chain (\<lambda>i. ?e (Y i))"
+  have lub_below: "\<And>x. f (\<Squnion>i. Y i)\<cdot>x \<sqsubseteq> x"
+    by (rule admD [OF _ Y], simp add: cont, rule below)
+  have "deflation (?e (\<Squnion>i. Y i))"
+    apply (rule pre_deflation.deflation_d)
+    apply (rule pre_deflation_oo [OF d lub_below])
+    done
+  then show "?e (\<Squnion>i. Y i) \<sqsubseteq> (\<Squnion>i. ?e (Y i))"
+  proof (rule deflation.belowI)
+    fix x :: 'a
+    assume "?e (\<Squnion>i. Y i)\<cdot>x = x"
+    hence "d\<cdot>x = x" and "f (\<Squnion>i. Y i)\<cdot>x = x"
+      by (simp_all add: eventual_iterate_oo_fixed_iff [OF d lub_below])
+    hence "(\<Squnion>i. f (Y i)\<cdot>x) = x"
+      apply (simp only: cont2contlubE [OF cont Y])
+      apply (simp only: contlub_cfun_fun [OF fY])
+      done
+    have "compact (d\<cdot>x)"
+      using d by (rule finite_deflation.compact)
+    then have "compact x"
+      using `d\<cdot>x = x` by simp
+    then have "compact (\<Squnion>i. f (Y i)\<cdot>x)"
+      using `(\<Squnion>i. f (Y i)\<cdot>x) = x` by simp
+    then have "\<exists>n. max_in_chain n (\<lambda>i. f (Y i)\<cdot>x)"
+      by - (rule compact_imp_max_in_chain, simp add: fY, assumption)
+    then obtain n where n: "max_in_chain n (\<lambda>i. f (Y i)\<cdot>x)" ..
+    then have "f (Y n)\<cdot>x = x"
+      using `(\<Squnion>i. f (Y i)\<cdot>x) = x` fY by (simp add: maxinch_is_thelub)
+    with `d\<cdot>x = x` have "?e (Y n)\<cdot>x = x"
+      by (simp add: eventual_iterate_oo_fixed_iff [OF d below])
+    moreover have "?e (Y n)\<cdot>x \<sqsubseteq> (\<Squnion>i. ?e (Y i)\<cdot>x)"
+      by (rule is_ub_thelub, simp add: eY)
+    ultimately have "x \<sqsubseteq> (\<Squnion>i. ?e (Y i))\<cdot>x"
+      by (simp add: contlub_cfun_fun eY)
+    also have "(\<Squnion>i. ?e (Y i))\<cdot>x \<sqsubseteq> x"
+      apply (rule deflation.below)
+      apply (rule admD [OF adm_deflation eY])
+      apply (rule pre_deflation.deflation_d)
+      apply (rule pre_deflation_oo [OF d below])
+      done
+    finally show "(\<Squnion>i. ?e (Y i))\<cdot>x = x" ..
+  qed
+qed
+
+
 subsection {* Type constructor for finite deflations *}
 
 defaultsort profinite
@@ -214,6 +317,10 @@
 lemma finite_deflation_Rep_fin_defl: "finite_deflation (Rep_fin_defl d)"
 using Rep_fin_defl by simp
 
+lemma deflation_Rep_fin_defl: "deflation (Rep_fin_defl d)"
+using finite_deflation_Rep_fin_defl
+by (rule finite_deflation_imp_deflation)
+
 interpretation Rep_fin_defl: finite_deflation "Rep_fin_defl d"
 by (rule finite_deflation_Rep_fin_defl)
 
@@ -244,27 +351,69 @@
 subsection {* Take function for finite deflations *}
 
 definition
+  defl_approx :: "nat \<Rightarrow> ('a \<rightarrow> 'a) \<Rightarrow> ('a \<rightarrow> 'a)"
+where
+  "defl_approx i d = eventual_iterate (approx i oo d)"
+
+lemma finite_deflation_defl_approx:
+  "deflation d \<Longrightarrow> finite_deflation (defl_approx i d)"
+unfolding defl_approx_def
+apply (rule pre_deflation.finite_deflation_d)
+apply (rule pre_deflation_oo)
+apply (rule finite_deflation_approx)
+apply (erule deflation.below)
+done
+
+lemma deflation_defl_approx:
+  "deflation d \<Longrightarrow> deflation (defl_approx i d)"
+apply (rule finite_deflation_imp_deflation)
+apply (erule finite_deflation_defl_approx)
+done
+
+lemma defl_approx_fixed_iff:
+  "deflation d \<Longrightarrow> defl_approx i d\<cdot>x = x \<longleftrightarrow> approx i\<cdot>x = x \<and> d\<cdot>x = x"
+unfolding defl_approx_def
+apply (rule eventual_iterate_oo_fixed_iff)
+apply (rule finite_deflation_approx)
+apply (erule deflation.below)
+done
+
+lemma defl_approx_below:
+  "\<lbrakk>a \<sqsubseteq> b; deflation a; deflation b\<rbrakk> \<Longrightarrow> defl_approx i a \<sqsubseteq> defl_approx i b"
+apply (rule deflation.belowI)
+apply (erule deflation_defl_approx)
+apply (simp add: defl_approx_fixed_iff)
+apply (erule (1) deflation.belowD)
+apply (erule conjunct2)
+done
+
+lemma cont2cont_defl_approx:
+  assumes cont: "cont f" and below: "\<And>x y. f x\<cdot>y \<sqsubseteq> y"
+  shows "cont (\<lambda>x. defl_approx i (f x))"
+unfolding defl_approx_def
+using finite_deflation_approx assms
+by (rule cont2cont_eventual_iterate_oo)
+
+definition
   fd_take :: "nat \<Rightarrow> 'a fin_defl \<Rightarrow> 'a fin_defl"
 where
-  "fd_take i d = Abs_fin_defl (eventual (\<lambda>n. iterate n\<cdot>(approx i oo Rep_fin_defl d)))"
+  "fd_take i d = Abs_fin_defl (defl_approx i (Rep_fin_defl d))"
 
 lemma Rep_fin_defl_fd_take:
-  "Rep_fin_defl (fd_take i d) =
-    eventual (\<lambda>n. iterate n\<cdot>(approx i oo Rep_fin_defl d))"
+  "Rep_fin_defl (fd_take i d) = defl_approx i (Rep_fin_defl d)"
 unfolding fd_take_def
 apply (rule Abs_fin_defl_inverse [unfolded mem_Collect_eq])
-apply (rule pre_deflation.finite_deflation_d)
-apply (rule pre_deflation_d_f)
-apply (rule finite_deflation_approx)
-apply (rule Rep_fin_defl.below)
+apply (rule finite_deflation_defl_approx)
+apply (rule deflation_Rep_fin_defl)
 done
 
 lemma fd_take_fixed_iff:
   "Rep_fin_defl (fd_take i d)\<cdot>x = x \<longleftrightarrow>
     approx i\<cdot>x = x \<and> Rep_fin_defl d\<cdot>x = x"
 unfolding Rep_fin_defl_fd_take
-by (rule eventual_iterate_oo_fixed_iff
-    [OF finite_deflation_approx Rep_fin_defl.below])
+apply (rule defl_approx_fixed_iff)
+apply (rule deflation_Rep_fin_defl)
+done
 
 lemma fd_take_below: "fd_take n d \<sqsubseteq> d"
 apply (rule fin_defl_belowI)
@@ -463,6 +612,41 @@
 interpretation cast: deflation "cast\<cdot>d"
 by (rule deflation_cast)
 
+lemma cast_approx: "cast\<cdot>(approx n\<cdot>A) = defl_approx n (cast\<cdot>A)"
+apply (rule alg_defl.principal_induct)
+apply (rule adm_eq)
+apply simp
+apply (simp add: cont2cont_defl_approx cast.below)
+apply (simp only: approx_alg_defl_principal)
+apply (simp only: cast_alg_defl_principal)
+apply (simp only: Rep_fin_defl_fd_take)
+done
+
+lemma cast_approx_fixed_iff:
+  "cast\<cdot>(approx i\<cdot>A)\<cdot>x = x \<longleftrightarrow> approx i\<cdot>x = x \<and> cast\<cdot>A\<cdot>x = x"
+apply (simp only: cast_approx)
+apply (rule defl_approx_fixed_iff)
+apply (rule deflation_cast)
+done
+
+lemma defl_approx_cast: "defl_approx i (cast\<cdot>A) = cast\<cdot>(approx i\<cdot>A)"
+by (rule cast_approx [symmetric])
+
+lemma cast_below_imp_below: "cast\<cdot>A \<sqsubseteq> cast\<cdot>B \<Longrightarrow> A \<sqsubseteq> B"
+apply (rule profinite_below_ext)
+apply (drule_tac i=i in defl_approx_below)
+apply (rule deflation_cast)
+apply (rule deflation_cast)
+apply (simp only: defl_approx_cast)
+apply (cut_tac x="approx i\<cdot>A" in alg_defl.compact_imp_principal)
+apply (rule compact_approx)
+apply (cut_tac x="approx i\<cdot>B" in alg_defl.compact_imp_principal)
+apply (rule compact_approx)
+apply clarsimp
+apply (simp add: cast_alg_defl_principal)
+apply (simp add: below_fin_defl_def)
+done
+
 lemma "cast\<cdot>(\<Squnion>i. alg_defl_principal (Abs_fin_defl (approx i)))\<cdot>x = x"
 apply (subst contlub_cfun_arg)
 apply (rule chainI)
--- a/src/HOLCF/Tools/domain/domain_extender.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOLCF/Tools/domain/domain_extender.ML	Tue May 19 14:13:37 2009 +0100
@@ -6,11 +6,13 @@
 
 signature DOMAIN_EXTENDER =
 sig
-  val add_domain_cmd: string -> (string list * binding * mixfix *
-    (binding * (bool * binding option * string) list * mixfix) list) list
+  val add_domain_cmd: string ->
+    ((string * string option) list * binding * mixfix *
+      (binding * (bool * binding option * string) list * mixfix) list) list
     -> theory -> theory
-  val add_domain: string -> (string list * binding * mixfix *
-    (binding * (bool * binding option * typ) list * mixfix) list) list
+  val add_domain: string ->
+    ((string * string option) list * binding * mixfix *
+      (binding * (bool * binding option * typ) list * mixfix) list) list
     -> theory -> theory
 end;
 
@@ -22,7 +24,7 @@
 (* ----- general testing and preprocessing of constructor list -------------- *)
 fun check_and_sort_domain
   (dtnvs : (string * typ list) list)
-  (cons'' : ((binding * (bool * binding option * typ) list * mixfix) list) list)
+  (cons'' : (binding * (bool * binding option * typ) list * mixfix) list list)
   (sg : theory)
   : ((string * typ list) *
       (binding * (bool * binding option * typ) list * mixfix) list) list =
@@ -73,9 +75,15 @@
 					     quote (string_of_typ sg t) ^ 
 					    " with different arguments"))
         |   analyse indirect (TVar _) = Imposs "extender:analyse";
-	fun check_pcpo T = if pcpo_type sg T then T
-          else error("Constructor argument type is not of sort pcpo: "^string_of_typ sg T);
-	val analyse_con = upd_second (map (upd_third (check_pcpo o analyse false)));
+        fun check_pcpo lazy T =
+          let val ok = if lazy then cpo_type else pcpo_type
+          in if ok sg T then T else error
+            ("Constructor argument type is not of sort pcpo: " ^
+              string_of_typ sg T)
+          end;
+        fun analyse_arg (lazy, sel, T) =
+          (lazy, sel, check_pcpo lazy (analyse false T));
+        fun analyse_con (b, args, mx) = (b, map analyse_arg args, mx);
       in ((dname,distinct_typevars), map analyse_con cons') end; 
   in ListPair.map analyse_equation (dtnvs,cons'')
   end; (* let *)
@@ -85,12 +93,16 @@
 fun gen_add_domain
   (prep_typ : theory -> 'a -> typ)
   (comp_dnam : string)
-  (eqs''' : (string list * binding * mixfix *
+  (eqs''' : ((string * string option) list * binding * mixfix *
               (binding * (bool * binding option * 'a) list * mixfix) list) list)
   (thy''' : theory) =
   let
+    fun readS (SOME s) = Syntax.read_sort_global thy''' s
+      | readS NONE = Sign.defaultS thy''';
+    fun readTFree (a, s) = TFree (a, readS s);
+
     val dtnvs = map (fn (vs,dname:binding,mx,_) => 
-                  (dname, map (Syntax.read_typ_global thy''') vs, mx)) eqs''';
+                  (dname, map readTFree vs, mx)) eqs''';
     val cons''' = map (fn (_,_,_,cons) => cons) eqs''';
     fun thy_type  (dname,tvars,mx) = (dname, length tvars, mx);
     fun thy_arity (dname,tvars,mx) = (Sign.full_name thy''' dname, map (snd o dest_TFree) tvars, pcpoS);
@@ -148,10 +160,10 @@
 val cons_decl =
   P.binding -- Scan.repeat dest_decl -- P.opt_mixfix;
 
-val type_var' =
-  (P.type_ident ^^ Scan.optional (P.$$$ "::" ^^ P.!!! P.sort) "");
+val type_var' : (string * string option) parser =
+  (P.type_ident -- Scan.option (P.$$$ "::" |-- P.!!! P.sort));
 
-val type_args' =
+val type_args' : (string * string option) list parser =
   type_var' >> single ||
   P.$$$ "(" |-- P.!!! (P.list1 type_var' --| P.$$$ ")") ||
   Scan.succeed [];
@@ -164,11 +176,12 @@
   Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") --
   P.and_list1 domain_decl;
 
-fun mk_domain (opt_name : string option, doms : (((string list * binding) * mixfix) *
+fun mk_domain (opt_name : string option,
+  doms : ((((string * string option) list * binding) * mixfix) *
     ((binding * (bool * binding option * string) list) * mixfix) list) list ) =
   let
     val names = map (fn (((_, t), _), _) => Binding.name_of t) doms;
-    val specs : (string list * binding * mixfix *
+    val specs : ((string * string option) list * binding * mixfix *
       (binding * (bool * binding option * string) list * mixfix) list) list =
         map (fn (((vs, t), mx), cons) =>
           (vs, t, mx, map (fn ((c, ds), mx) => (c, ds, mx)) cons)) doms;
--- a/src/HOLCF/Tools/domain/domain_library.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOLCF/Tools/domain/domain_library.ML	Tue May 19 14:13:37 2009 +0100
@@ -52,6 +52,7 @@
 signature DOMAIN_LIBRARY =
 sig
   val Imposs : string -> 'a;
+  val cpo_type : theory -> typ -> bool;
   val pcpo_type : theory -> typ -> bool;
   val string_of_typ : theory -> typ -> string;
 
@@ -190,6 +191,7 @@
       | index_vnames([],occupied) = [];
 in index_vnames(map nonreserved ids, [("O",0),("o",0)]) end;
 
+fun cpo_type sg t = Sign.of_sort sg (Sign.certify_typ sg t, @{sort cpo});
 fun pcpo_type sg t = Sign.of_sort sg (Sign.certify_typ sg t, @{sort pcpo});
 fun string_of_typ sg = Syntax.string_of_typ_global sg o Sign.certify_typ sg;
 
--- a/src/HOLCF/Tools/domain/domain_theorems.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOLCF/Tools/domain/domain_theorems.ML	Tue May 19 14:13:37 2009 +0100
@@ -12,6 +12,8 @@
 sig
   val theorems: Domain_Library.eq * Domain_Library.eq list -> theory -> thm list * theory;
   val comp_theorems: bstring * Domain_Library.eq list -> theory -> thm list * theory;
+  val quiet_mode: bool ref;
+  val trace_domain: bool ref;
 end;
 
 structure Domain_Theorems :> DOMAIN_THEOREMS =
--- a/src/HOLCF/Tools/fixrec_package.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/HOLCF/Tools/fixrec_package.ML	Tue May 19 14:13:37 2009 +0100
@@ -195,7 +195,7 @@
     val unfold_thms = unfolds names tuple_unfold_thm;
     fun mk_note (n, thm) = ((Binding.name n, []), [thm]);
     val (thmss, lthy'') = lthy'
-      |> fold_map (LocalTheory.note Thm.theoremK o mk_note)
+      |> fold_map (LocalTheory.note Thm.generatedK o mk_note)
         ((all_names ^ "_induct", tuple_induct_thm) :: unfold_thms);
   in
     (lthy'', names, fixdef_thms, map snd unfold_thms)
@@ -373,7 +373,7 @@
       val simps2 : (Attrib.binding * thm list) list =
         map (apsnd (fn thm => [thm])) (List.concat simps);
       val (_, lthy'') = lthy'
-        |> fold_map (LocalTheory.note Thm.theoremK) (simps1 @ simps2);
+        |> fold_map (LocalTheory.note Thm.generatedK) (simps1 @ simps2);
     in
       lthy''
     end
--- a/src/Provers/quantifier1.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/Provers/quantifier1.ML	Tue May 19 14:13:37 2009 +0100
@@ -21,11 +21,21 @@
         "!x. x=t --> P(x)" is covered by the congreunce rule for -->;
         "!x. t=x --> P(x)" must be taken care of by an ordinary rewrite rule.
         As must be "? x. t=x & P(x)".
-
         
      And similarly for the bounded quantifiers.
 
 Gries etc call this the "1 point rules"
+
+The above also works for !x1..xn. and ?x1..xn by moving the defined
+qunatifier inside first, but not for nested bounded quantifiers.
+
+For set comprehensions the basic permutations
+      ... & x = t & ...  ->  x = t & (... & ...)
+      ... & t = x & ...  ->  t = x & (... & ...)
+are also exported.
+
+To avoid looping, NONE is returned if the term cannot be rearranged,
+esp if x=t/t=x sits at the front already.
 *)
 
 signature QUANTIFIER1_DATA =
@@ -61,6 +71,7 @@
   val rearrange_ex:  theory -> simpset -> term -> thm option
   val rearrange_ball: (simpset -> tactic) -> theory -> simpset -> term -> thm option
   val rearrange_bex:  (simpset -> tactic) -> theory -> simpset -> term -> thm option
+  val rearrange_Coll: tactic -> theory -> simpset -> term -> thm option
 end;
 
 functor Quantifier1Fun(Data: QUANTIFIER1_DATA): QUANTIFIER1 =
@@ -70,29 +81,28 @@
 
 (* FIXME: only test! *)
 fun def xs eq =
-  let val n = length xs
-  in case dest_eq eq of
-      SOME(c,s,t) =>
-        s = Bound n andalso not(loose_bvar1(t,n)) orelse
-        t = Bound n andalso not(loose_bvar1(s,n))
-    | NONE => false
-  end;
+  (case dest_eq eq of
+     SOME(c,s,t) =>
+       let val n = length xs
+       in s = Bound n andalso not(loose_bvar1(t,n)) orelse
+          t = Bound n andalso not(loose_bvar1(s,n)) end
+   | NONE => false);
 
-fun extract_conj xs t = case dest_conj t of NONE => NONE
+fun extract_conj fst xs t = case dest_conj t of NONE => NONE
     | SOME(conj,P,Q) =>
-        (if def xs P then SOME(xs,P,Q) else
+        (if def xs P then (if fst then NONE else SOME(xs,P,Q)) else
          if def xs Q then SOME(xs,Q,P) else
-         (case extract_conj xs P of
+         (case extract_conj false xs P of
             SOME(xs,eq,P') => SOME(xs,eq, conj $ P' $ Q)
-          | NONE => (case extract_conj xs Q of
+          | NONE => (case extract_conj false xs Q of
                        SOME(xs,eq,Q') => SOME(xs,eq,conj $ P $ Q')
                      | NONE => NONE)));
 
-fun extract_imp xs t = case dest_imp t of NONE => NONE
-    | SOME(imp,P,Q) => if def xs P then SOME(xs,P,Q)
-                       else (case extract_conj xs P of
+fun extract_imp fst xs t = case dest_imp t of NONE => NONE
+    | SOME(imp,P,Q) => if def xs P then (if fst then NONE else SOME(xs,P,Q))
+                       else (case extract_conj false xs P of
                                SOME(xs,eq,P') => SOME(xs, eq, imp $ P' $ Q)
-                             | NONE => (case extract_imp xs Q of
+                             | NONE => (case extract_imp false xs Q of
                                           NONE => NONE
                                         | SOME(xs,eq,Q') =>
                                             SOME(xs,eq,imp$P$Q')));
@@ -100,8 +110,8 @@
 fun extract_quant extract q =
   let fun exqu xs ((qC as Const(qa,_)) $ Abs(x,T,Q)) =
             if qa = q then exqu ((qC,x,T)::xs) Q else NONE
-        | exqu xs P = extract xs P
-  in exqu end;
+        | exqu xs P = extract (null xs) xs P
+  in exqu [] end;
 
 fun prove_conv tac thy tu =
   Goal.prove (ProofContext.init thy) [] [] (Logic.mk_equals tu)
@@ -147,7 +157,7 @@
   in quant xs (qC $ Abs(x,T,Q)) end;
 
 fun rearrange_all thy _ (F as (all as Const(q,_)) $ Abs(x,T, P)) =
-     (case extract_quant extract_imp q [] P of
+     (case extract_quant extract_imp q P of
         NONE => NONE
       | SOME(xs,eq,Q) =>
           let val R = quantify all x T xs (imp $ eq $ Q)
@@ -155,7 +165,7 @@
   | rearrange_all _ _ _ = NONE;
 
 fun rearrange_ball tac thy ss (F as Ball $ A $ Abs(x,T,P)) =
-     (case extract_imp [] P of
+     (case extract_imp true [] P of
         NONE => NONE
       | SOME(xs,eq,Q) => if not(null xs) then NONE else
           let val R = imp $ eq $ Q
@@ -163,7 +173,7 @@
   | rearrange_ball _ _ _ _ = NONE;
 
 fun rearrange_ex thy _ (F as (ex as Const(q,_)) $ Abs(x,T,P)) =
-     (case extract_quant extract_conj q [] P of
+     (case extract_quant extract_conj q P of
         NONE => NONE
       | SOME(xs,eq,Q) =>
           let val R = quantify ex x T xs (conj $ eq $ Q)
@@ -171,10 +181,17 @@
   | rearrange_ex _ _ _ = NONE;
 
 fun rearrange_bex tac thy ss (F as Bex $ A $ Abs(x,T,P)) =
-     (case extract_conj [] P of
+     (case extract_conj true [] P of
         NONE => NONE
       | SOME(xs,eq,Q) => if not(null xs) then NONE else
           SOME(prove_conv (tac ss) thy (F,Bex $ A $ Abs(x,T,conj$eq$Q))))
   | rearrange_bex _ _ _ _ = NONE;
 
+fun rearrange_Coll tac thy _ (F as Coll $ Abs(x,T,P)) =
+     (case extract_conj true [] P of
+        NONE => NONE
+      | SOME(_,eq,Q) =>
+          let val R = Coll $ Abs(x,T, conj $ eq $ Q)
+          in SOME(prove_conv tac thy (F,R)) end);
+
 end;
--- a/src/Pure/IsaMakefile	Tue May 19 14:13:23 2009 +0100
+++ b/src/Pure/IsaMakefile	Tue May 19 14:13:37 2009 +0100
@@ -56,7 +56,7 @@
   General/table.ML General/url.ML General/xml.ML General/yxml.ML	\
   Isar/ROOT.ML Isar/args.ML Isar/attrib.ML Isar/auto_bind.ML		\
   Isar/calculation.ML Isar/class.ML Isar/class_target.ML Isar/code.ML	\
-  Isar/code_unit.ML Isar/constdefs.ML Isar/context_rules.ML		\
+  Isar/constdefs.ML Isar/context_rules.ML		\
   Isar/element.ML Isar/expression.ML Isar/isar_cmd.ML			\
   Isar/isar_document.ML Isar/isar_syn.ML Isar/local_defs.ML		\
   Isar/local_syntax.ML Isar/local_theory.ML Isar/locale.ML		\
--- a/src/Pure/Isar/ROOT.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/Pure/Isar/ROOT.ML	Tue May 19 14:13:37 2009 +0100
@@ -61,7 +61,6 @@
 use "../simplifier.ML";
 
 (*executable theory content*)
-use "code_unit.ML";
 use "code.ML";
 
 (*specifications*)
--- a/src/Pure/Isar/attrib.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/Pure/Isar/attrib.ML	Tue May 19 14:13:37 2009 +0100
@@ -123,7 +123,7 @@
 
 fun attribute thy = attribute_i thy o intern_src thy;
 
-fun eval_thms ctxt args = ProofContext.note_thmss Thm.theoremK
+fun eval_thms ctxt args = ProofContext.note_thmss Thm.generatedK
     [(Thm.empty_binding, args |> map (fn (a, atts) =>
         (ProofContext.get_fact ctxt a, map (attribute (ProofContext.theory_of ctxt)) atts)))] ctxt
   |> fst |> maps snd;
--- a/src/Pure/Isar/code.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/Pure/Isar/code.ML	Tue May 19 14:13:37 2009 +0100
@@ -7,6 +7,55 @@
 
 signature CODE =
 sig
+  (*constructor sets*)
+  val constrset_of_consts: theory -> (string * typ) list
+    -> string * ((string * sort) list * (string * typ list) list)
+
+  (*typ instantiations*)
+  val typscheme: theory -> string * typ -> (string * sort) list * typ
+  val inst_thm: theory -> sort Vartab.table -> thm -> thm
+
+  (*constants*)
+  val string_of_typ: theory -> typ -> string
+  val string_of_const: theory -> string -> string
+  val no_args: theory -> string -> int
+  val check_const: theory -> term -> string
+  val read_bare_const: theory -> string -> string * typ
+  val read_const: theory -> string -> string
+
+  (*constant aliasses*)
+  val add_const_alias: thm -> theory -> theory
+  val triv_classes: theory -> class list
+  val resubst_alias: theory -> string -> string
+
+  (*code equations*)
+  val mk_eqn: theory -> (string -> bool) -> thm * bool -> thm * bool
+  val mk_eqn_liberal: theory -> (string -> bool) -> thm -> (thm * bool) option
+  val assert_eqn: theory -> thm * bool -> thm * bool
+  val assert_eqns_const: theory -> string
+    -> (thm * bool) list -> (thm * bool) list
+  val const_typ_eqn: thm -> string * typ
+  val const_eqn: theory -> thm -> string
+  val typscheme_eqn: theory -> thm -> (string * sort) list * typ
+  val expand_eta: theory -> int -> thm -> thm
+  val rewrite_eqn: simpset -> thm -> thm
+  val rewrite_head: thm list -> thm -> thm
+  val norm_args: theory -> thm list -> thm list 
+  val norm_varnames: theory -> thm list -> thm list
+
+  (*case certificates*)
+  val case_cert: thm -> string * (int * string list)
+
+  (*infrastructure*)
+  val add_attribute: string * attribute parser -> theory -> theory
+  val purge_data: theory -> theory
+
+  (*executable content*)
+  val add_datatype: (string * typ) list -> theory -> theory
+  val add_datatype_cmd: string list -> theory -> theory
+  val type_interpretation:
+    (string * ((string * sort) list * (string * typ list) list)
+      -> theory -> theory) -> theory -> theory
   val add_eqn: thm -> theory -> theory
   val add_nbe_eqn: thm -> theory -> theory
   val add_default_eqn: thm -> theory -> theory
@@ -15,27 +64,16 @@
   val del_eqn: thm -> theory -> theory
   val del_eqns: string -> theory -> theory
   val add_eqnl: string * (thm * bool) list lazy -> theory -> theory
-  val add_datatype: (string * typ) list -> theory -> theory
-  val add_datatype_cmd: string list -> theory -> theory
-  val type_interpretation:
-    (string * ((string * sort) list * (string * typ list) list)
-      -> theory -> theory) -> theory -> theory
   val add_case: thm -> theory -> theory
   val add_undefined: string -> theory -> theory
-  val purge_data: theory -> theory
 
-  val these_eqns: theory -> string -> (thm * bool) list
+  (*data retrieval*)
   val get_datatype: theory -> string -> ((string * sort) list * (string * typ list) list)
   val get_datatype_of_constr: theory -> string -> string option
+  val default_typscheme: theory -> string -> (string * sort) list * typ
+  val these_eqns: theory -> string -> (thm * bool) list
   val get_case_scheme: theory -> string -> (int * (int * string list)) option
   val is_undefined: theory -> string -> bool
-  val default_typscheme: theory -> string -> (string * sort) list * typ
-  val assert_eqn: theory -> thm * bool -> thm * bool
-  val assert_eqns_const: theory -> string
-    -> (thm * bool) list -> (thm * bool) list
-
-  val add_attribute: string * attribute parser -> theory -> theory
-
   val print_codesetup: theory -> unit
 end;
 
@@ -70,6 +108,400 @@
 structure Code : PRIVATE_CODE =
 struct
 
+(* auxiliary *)
+
+fun string_of_typ thy = setmp show_sorts true (Syntax.string_of_typ_global thy);
+fun string_of_const thy c = case AxClass.inst_of_param thy c
+ of SOME (c, tyco) => Sign.extern_const thy c ^ " " ^ enclose "[" "]" (Sign.extern_type thy tyco)
+  | NONE => Sign.extern_const thy c;
+
+fun no_args thy = length o fst o strip_type o Sign.the_const_type thy;
+
+
+(* utilities *)
+
+fun typscheme thy (c, ty) =
+  let
+    val ty' = Logic.unvarifyT ty;
+    fun dest (TFree (v, sort)) = (v, sort)
+      | dest ty = error ("Illegal type parameter in type scheme: " ^ Syntax.string_of_typ_global thy ty);
+    val vs = map dest (Sign.const_typargs thy (c, ty'));
+  in (vs, Type.strip_sorts ty') end;
+
+fun inst_thm thy tvars' thm =
+  let
+    val tvars = (Term.add_tvars o Thm.prop_of) thm [];
+    val inter_sort = Sorts.inter_sort (Sign.classes_of thy);
+    fun mk_inst (tvar as (v, sort)) = case Vartab.lookup tvars' v
+     of SOME sort' => SOME (pairself (Thm.ctyp_of thy o TVar)
+          (tvar, (v, inter_sort (sort, sort'))))
+      | NONE => NONE;
+    val insts = map_filter mk_inst tvars;
+  in Thm.instantiate (insts, []) thm end;
+
+fun expand_eta thy k thm =
+  let
+    val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm;
+    val (head, args) = strip_comb lhs;
+    val l = if k = ~1
+      then (length o fst o strip_abs) rhs
+      else Int.max (0, k - length args);
+    val used = Name.make_context (map (fst o fst) (Term.add_vars lhs []));
+    fun get_name _ 0 = pair []
+      | get_name (Abs (v, ty, t)) k =
+          Name.variants [v]
+          ##>> get_name t (k - 1)
+          #>> (fn ([v'], vs') => (v', ty) :: vs')
+      | get_name t k = 
+          let
+            val (tys, _) = (strip_type o fastype_of) t
+          in case tys
+           of [] => raise TERM ("expand_eta", [t])
+            | ty :: _ =>
+                Name.variants [""]
+                #-> (fn [v] => get_name (t $ Var ((v, 0), ty)) (k - 1)
+                #>> (fn vs' => (v, ty) :: vs'))
+          end;
+    val (vs, _) = get_name rhs l used;
+    fun expand (v, ty) thm = Drule.fun_cong_rule thm
+      (Thm.cterm_of thy (Var ((v, 0), ty)));
+  in
+    thm
+    |> fold expand vs
+    |> Conv.fconv_rule Drule.beta_eta_conversion
+  end;
+
+fun eqn_conv conv =
+  let
+    fun lhs_conv ct = if can Thm.dest_comb ct
+      then (Conv.combination_conv lhs_conv conv) ct
+      else Conv.all_conv ct;
+  in Conv.combination_conv (Conv.arg_conv lhs_conv) conv end;
+
+fun head_conv conv =
+  let
+    fun lhs_conv ct = if can Thm.dest_comb ct
+      then (Conv.fun_conv lhs_conv) ct
+      else conv ct;
+  in Conv.fun_conv (Conv.arg_conv lhs_conv) end;
+
+val rewrite_eqn = Conv.fconv_rule o eqn_conv o Simplifier.rewrite;
+val rewrite_head = Conv.fconv_rule o head_conv o MetaSimplifier.rewrite false;
+
+fun norm_args thy thms =
+  let
+    val num_args_of = length o snd o strip_comb o fst o Logic.dest_equals;
+    val k = fold (curry Int.max o num_args_of o Thm.prop_of) thms 0;
+  in
+    thms
+    |> map (expand_eta thy k)
+    |> map (Conv.fconv_rule Drule.beta_eta_conversion)
+  end;
+
+fun canonical_tvars thy thm =
+  let
+    val ctyp = Thm.ctyp_of thy;
+    val purify_tvar = unprefix "'" #> Name.desymbolize false #> prefix "'";
+    fun tvars_subst_for thm = (fold_types o fold_atyps)
+      (fn TVar (v_i as (v, _), sort) => let
+            val v' = purify_tvar v
+          in if v = v' then I
+          else insert (op =) (v_i, (v', sort)) end
+        | _ => I) (prop_of thm) [];
+    fun mk_inst (v_i, (v', sort)) (maxidx, acc) =
+      let
+        val ty = TVar (v_i, sort)
+      in
+        (maxidx + 1, (ctyp ty, ctyp (TVar ((v', maxidx), sort))) :: acc)
+      end;
+    val maxidx = Thm.maxidx_of thm + 1;
+    val (_, inst) = fold mk_inst (tvars_subst_for thm) (maxidx + 1, []);
+  in Thm.instantiate (inst, []) thm end;
+
+fun canonical_vars thy thm =
+  let
+    val cterm = Thm.cterm_of thy;
+    val purify_var = Name.desymbolize false;
+    fun vars_subst_for thm = fold_aterms
+      (fn Var (v_i as (v, _), ty) => let
+            val v' = purify_var v
+          in if v = v' then I
+          else insert (op =) (v_i, (v', ty)) end
+        | _ => I) (prop_of thm) [];
+    fun mk_inst (v_i as (v, i), (v', ty)) (maxidx, acc) =
+      let
+        val t = Var (v_i, ty)
+      in
+        (maxidx + 1, (cterm t, cterm (Var ((v', maxidx), ty))) :: acc)
+      end;
+    val maxidx = Thm.maxidx_of thm + 1;
+    val (_, inst) = fold mk_inst (vars_subst_for thm) (maxidx + 1, []);
+  in Thm.instantiate ([], inst) thm end;
+
+fun canonical_absvars thm =
+  let
+    val t = Thm.plain_prop_of thm;
+    val purify_var = Name.desymbolize false;
+    val t' = Term.map_abs_vars purify_var t;
+  in Thm.rename_boundvars t t' thm end;
+
+fun norm_varnames thy thms =
+  let
+    fun burrow_thms f [] = []
+      | burrow_thms f thms =
+          thms
+          |> Conjunction.intr_balanced
+          |> f
+          |> Conjunction.elim_balanced (length thms)
+  in
+    thms
+    |> map (canonical_vars thy)
+    |> map canonical_absvars
+    |> map Drule.zero_var_indexes
+    |> burrow_thms (canonical_tvars thy)
+    |> Drule.zero_var_indexes_list
+  end;
+
+
+(* const aliasses *)
+
+structure ConstAlias = TheoryDataFun
+(
+  type T = ((string * string) * thm) list * class list;
+  val empty = ([], []);
+  val copy = I;
+  val extend = I;
+  fun merge _ ((alias1, classes1), (alias2, classes2)) : T =
+    (Library.merge (eq_snd Thm.eq_thm_prop) (alias1, alias2),
+      Library.merge (op =) (classes1, classes2));
+);
+
+fun add_const_alias thm thy =
+  let
+    val lhs_rhs = case try Logic.dest_equals (Thm.prop_of thm)
+     of SOME lhs_rhs => lhs_rhs
+      | _ => error ("Not an equation: " ^ Display.string_of_thm thm);
+    val c_c' = case try (pairself (AxClass.unoverload_const thy o dest_Const)) lhs_rhs
+     of SOME c_c' => c_c'
+      | _ => error ("Not an equation with two constants: " ^ Display.string_of_thm thm);
+    val some_class = the_list (AxClass.class_of_param thy (snd c_c'));
+  in thy |>
+    ConstAlias.map (fn (alias, classes) =>
+      ((c_c', thm) :: alias, fold (insert (op =)) some_class classes))
+  end;
+
+fun resubst_alias thy =
+  let
+    val alias = fst (ConstAlias.get thy);
+    val subst_inst_param = Option.map fst o AxClass.inst_of_param thy;
+    fun subst_alias c =
+      get_first (fn ((c', c''), _) => if c = c'' then SOME c' else NONE) alias;
+  in
+    perhaps subst_inst_param
+    #> perhaps subst_alias
+  end;
+
+val triv_classes = snd o ConstAlias.get;
+
+
+(* reading constants as terms *)
+
+fun check_bare_const thy t = case try dest_Const t
+ of SOME c_ty => c_ty
+  | NONE => error ("Not a constant: " ^ Syntax.string_of_term_global thy t);
+
+fun check_const thy = AxClass.unoverload_const thy o check_bare_const thy;
+
+fun read_bare_const thy = check_bare_const thy o Syntax.read_term_global thy;
+
+fun read_const thy = AxClass.unoverload_const thy o read_bare_const thy;
+
+
+(* constructor sets *)
+
+fun constrset_of_consts thy cs =
+  let
+    val _ = map (fn (c, _) => if (is_some o AxClass.class_of_param thy) c
+      then error ("Is a class parameter: " ^ string_of_const thy c) else ()) cs;
+    fun no_constr (c, ty) = error ("Not a datatype constructor: " ^ string_of_const thy c
+      ^ " :: " ^ string_of_typ thy ty);
+    fun last_typ c_ty ty =
+      let
+        val frees = OldTerm.typ_tfrees ty;
+        val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) ty
+          handle TYPE _ => no_constr c_ty
+        val _ = if has_duplicates (eq_fst (op =)) vs then no_constr c_ty else ();
+        val _ = if length frees <> length vs then no_constr c_ty else ();
+      in (tyco, vs) end;
+    fun ty_sorts (c, ty) =
+      let
+        val ty_decl = (Logic.unvarifyT o Sign.the_const_type thy) c;
+        val (tyco, _) = last_typ (c, ty) ty_decl;
+        val (_, vs) = last_typ (c, ty) ty;
+      in ((tyco, map snd vs), (c, (map fst vs, ty))) end;
+    fun add ((tyco', sorts'), c) ((tyco, sorts), cs) =
+      let
+        val _ = if tyco' <> tyco
+          then error "Different type constructors in constructor set"
+          else ();
+        val sorts'' = map2 (curry (Sorts.inter_sort (Sign.classes_of thy))) sorts' sorts
+      in ((tyco, sorts), c :: cs) end;
+    fun inst vs' (c, (vs, ty)) =
+      let
+        val the_v = the o AList.lookup (op =) (vs ~~ vs');
+        val ty' = map_atyps (fn TFree (v, _) => TFree (the_v v)) ty;
+      in (c, (fst o strip_type) ty') end;
+    val c' :: cs' = map ty_sorts cs;
+    val ((tyco, sorts), cs'') = fold add cs' (apsnd single c');
+    val vs = Name.names Name.context Name.aT sorts;
+    val cs''' = map (inst vs) cs'';
+  in (tyco, (vs, rev cs''')) end;
+
+
+(* code equations *)
+
+exception BAD_THM of string;
+fun bad_thm msg = raise BAD_THM msg;
+fun error_thm f thm = f thm handle BAD_THM msg => error msg;
+fun try_thm f thm = SOME (f thm) handle BAD_THM _ => NONE;
+
+fun is_linear thm =
+  let val (_, args) = (strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm
+  in not (has_duplicates (op =) ((fold o fold_aterms)
+    (fn Var (v, _) => cons v | _ => I) args [])) end;
+
+fun gen_assert_eqn thy is_constr_head is_constr_pat (thm, proper) =
+  let
+    val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm
+      handle TERM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm thm)
+           | THM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm thm);
+    fun vars_of t = fold_aterms (fn Var (v, _) => insert (op =) v
+      | Free _ => bad_thm ("Illegal free variable in equation\n"
+          ^ Display.string_of_thm thm)
+      | _ => I) t [];
+    fun tvars_of t = fold_term_types (fn _ =>
+      fold_atyps (fn TVar (v, _) => insert (op =) v
+        | TFree _ => bad_thm 
+      ("Illegal free type variable in equation\n" ^ Display.string_of_thm thm))) t [];
+    val lhs_vs = vars_of lhs;
+    val rhs_vs = vars_of rhs;
+    val lhs_tvs = tvars_of lhs;
+    val rhs_tvs = tvars_of rhs;
+    val _ = if null (subtract (op =) lhs_vs rhs_vs)
+      then ()
+      else bad_thm ("Free variables on right hand side of equation\n"
+        ^ Display.string_of_thm thm);
+    val _ = if null (subtract (op =) lhs_tvs rhs_tvs)
+      then ()
+      else bad_thm ("Free type variables on right hand side of equation\n"
+        ^ Display.string_of_thm thm)    val (head, args) = (strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm;
+    val (c, ty) = case head
+     of Const (c_ty as (_, ty)) => (AxClass.unoverload_const thy c_ty, ty)
+      | _ => bad_thm ("Equation not headed by constant\n" ^ Display.string_of_thm thm);
+    fun check _ (Abs _) = bad_thm
+          ("Abstraction on left hand side of equation\n"
+            ^ Display.string_of_thm thm)
+      | check 0 (Var _) = ()
+      | check _ (Var _) = bad_thm
+          ("Variable with application on left hand side of equation\n"
+            ^ Display.string_of_thm thm)
+      | check n (t1 $ t2) = (check (n+1) t1; check 0 t2)
+      | check n (Const (c_ty as (c, ty))) = if n = (length o fst o strip_type) ty
+          then if not proper orelse is_constr_pat (AxClass.unoverload_const thy c_ty)
+            then ()
+            else bad_thm (quote c ^ " is not a constructor, on left hand side of equation\n"
+              ^ Display.string_of_thm thm)
+          else bad_thm
+            ("Partially applied constant " ^ quote c ^ " on left hand side of equation\n"
+               ^ Display.string_of_thm thm);
+    val _ = map (check 0) args;
+    val _ = if not proper orelse is_linear thm then ()
+      else bad_thm ("Duplicate variables on left hand side of equation\n"
+        ^ Display.string_of_thm thm);
+    val _ = if (is_none o AxClass.class_of_param thy) c
+      then ()
+      else bad_thm ("Polymorphic constant as head in equation\n"
+        ^ Display.string_of_thm thm)
+    val _ = if not (is_constr_head c)
+      then ()
+      else bad_thm ("Constructor as head in equation\n"
+        ^ Display.string_of_thm thm)
+    val ty_decl = Sign.the_const_type thy c;
+    val _ = if Sign.typ_equiv thy (Type.strip_sorts ty_decl, Type.strip_sorts ty)
+      then () else bad_thm ("Type\n" ^ string_of_typ thy ty
+           ^ "\nof equation\n"
+           ^ Display.string_of_thm thm
+           ^ "\nis incompatible with declared function type\n"
+           ^ string_of_typ thy ty_decl)
+  in (thm, proper) end;
+
+fun assert_eqn thy is_constr = error_thm (gen_assert_eqn thy is_constr is_constr);
+
+val const_typ_eqn = dest_Const o fst o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of;
+
+
+(*those following are permissive wrt. to overloaded constants!*)
+
+fun mk_eqn thy is_constr_head = error_thm (gen_assert_eqn thy is_constr_head (K true)) o
+  apfst (LocalDefs.meta_rewrite_rule (ProofContext.init thy));
+
+fun mk_eqn_liberal thy is_constr_head = Option.map (fn (thm, _) => (thm, is_linear thm))
+  o try_thm (gen_assert_eqn thy is_constr_head (K true))
+  o rpair false o LocalDefs.meta_rewrite_rule (ProofContext.init thy);
+
+fun const_typ_eqn_unoverload thy thm =
+  let
+    val (c, ty) = const_typ_eqn thm;
+    val c' = AxClass.unoverload_const thy (c, ty);
+  in (c', ty) end;
+
+fun typscheme_eqn thy = typscheme thy o const_typ_eqn_unoverload thy;
+fun const_eqn thy = fst o const_typ_eqn_unoverload thy;
+
+
+(* case cerificates *)
+
+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;
+    val _ = case head of Free _ => true
+      | Var _ => true
+      | _ => raise TERM ("case_cert", []);
+    val ([(case_var, _)], case_expr) = Term.strip_abs_eta 1 raw_case_expr;
+    val (Const (case_const, _), raw_params) = strip_comb case_expr;
+    val n = find_index (fn Free (v, _) => v = case_var | _ => false) raw_params;
+    val _ = if n = ~1 then raise TERM ("case_cert", []) else ();
+    val params = map (fst o dest_Var) (nth_drop n raw_params);
+    fun dest_case t =
+      let
+        val (head' $ t_co, rhs) = Logic.dest_equals t;
+        val _ = if head' = head then () else raise TERM ("case_cert", []);
+        val (Const (co, _), args) = strip_comb t_co;
+        val (Var (param, _), args') = strip_comb rhs;
+        val _ = if args' = args then () else raise TERM ("case_cert", []);
+      in (param, co) end;
+    fun analyze_cases cases =
+      let
+        val co_list = fold (AList.update (op =) o dest_case) cases [];
+      in map (the o AList.lookup (op =) co_list) params end;
+    fun analyze_let t =
+      let
+        val (head' $ arg, Var (param', _) $ arg') = Logic.dest_equals t;
+        val _ = if head' = head then () else raise TERM ("case_cert", []);
+        val _ = if arg' = arg then () else raise TERM ("case_cert", []);
+        val _ = if [param'] = params then () else raise TERM ("case_cert", []);
+      in [] end;
+    fun analyze (cases as [let_case]) =
+          (analyze_cases cases handle Bind => analyze_let let_case)
+      | analyze cases = analyze_cases cases;
+  in (case_const, (n, analyze cases)) end;
+
+fun case_cert thm = case_certificate thm
+  handle Bind => error "bad case certificate"
+       | TERM _ => error "bad case certificate";
+
+
 (** code attributes **)
 
 structure CodeAttr = TheoryDataFun (
@@ -116,7 +548,8 @@
 
 fun add_drop_redundant thy (thm, proper) thms =
   let
-    val args_of = snd o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of;
+    val args_of = snd o strip_comb o map_types Type.strip_sorts
+      o fst o Logic.dest_equals o Thm.plain_prop_of;
     val args = args_of thm;
     val incr_idx = Logic.incr_indexes ([], Thm.maxidx_of thm + 1);
     fun matches_args args' = length args <= length args' andalso
@@ -332,16 +765,16 @@
           (Pretty.block o Pretty.breaks) (
             Pretty.str s
             :: Pretty.str "="
-            :: separate (Pretty.str "|") (map (fn (c, []) => Pretty.str (Code_Unit.string_of_const thy c)
+            :: separate (Pretty.str "|") (map (fn (c, []) => Pretty.str (string_of_const thy c)
                  | (c, tys) =>
                      (Pretty.block o Pretty.breaks)
-                        (Pretty.str (Code_Unit.string_of_const thy c)
+                        (Pretty.str (string_of_const thy c)
                           :: Pretty.str "of"
                           :: map (Pretty.quote o Syntax.pretty_typ_global thy) tys)) cos)
           );
     val eqns = the_eqns exec
       |> Symtab.dest
-      |> (map o apfst) (Code_Unit.string_of_const thy)
+      |> (map o apfst) (string_of_const thy)
       |> (map o apsnd) (snd o fst)
       |> sort (string_ord o pairself fst);
     val dtyps = the_dtyps exec
@@ -377,13 +810,13 @@
             val max' = Thm.maxidx_of thm' + 1;
           in (thm', max') end;
         val (thms', maxidx) = fold_map incr_thm thms 0;
-        val ty1 :: tys = map (snd o Code_Unit.const_typ_eqn) thms';
+        val ty1 :: tys = map (snd o const_typ_eqn) thms';
         fun unify ty env = Sign.typ_unify thy (ty1, ty) env
           handle Type.TUNIFY =>
             error ("Type unificaton failed, while unifying code equations\n"
             ^ (cat_lines o map Display.string_of_thm) thms
             ^ "\nwith types\n"
-            ^ (cat_lines o map (Code_Unit.string_of_typ thy)) (ty1 :: tys));
+            ^ (cat_lines o map (string_of_typ thy)) (ty1 :: tys));
         val (env, _) = fold unify tys (Vartab.empty, maxidx)
         val instT = Vartab.fold (fn (x_i, (sort, ty)) =>
           cons (Thm.ctyp_of thy (TVar (x_i, sort)), Thm.ctyp_of thy ty)) env [];
@@ -408,13 +841,13 @@
 
 fun is_constr thy = is_some o get_datatype_of_constr thy;
 
-fun assert_eqn thy = Code_Unit.assert_eqn thy (is_constr thy);
+val assert_eqn = fn thy => assert_eqn thy (is_constr thy);
 
 fun assert_eqns_const thy c eqns =
   let
-    fun cert (eqn as (thm, _)) = if c = Code_Unit.const_eqn thy thm
+    fun cert (eqn as (thm, _)) = if c = const_eqn thy thm
       then eqn else error ("Wrong head of code equation,\nexpected constant "
-        ^ Code_Unit.string_of_const thy c ^ "\n" ^ Display.string_of_thm thm)
+        ^ string_of_const thy c ^ "\n" ^ Display.string_of_thm thm)
   in map (cert o assert_eqn thy) eqns end;
 
 fun change_eqns delete c f = (map_exec_purge (SOME [c]) o map_eqns
@@ -422,19 +855,19 @@
     o apfst) (fn (_, eqns) => (true, f eqns));
 
 fun gen_add_eqn default (eqn as (thm, _)) thy =
-  let val c = Code_Unit.const_eqn thy thm
+  let val c = const_eqn thy thm
   in change_eqns false c (add_thm thy default eqn) thy end;
 
 fun add_eqn thm thy =
-  gen_add_eqn false (Code_Unit.mk_eqn thy (is_constr thy) (thm, true)) thy;
+  gen_add_eqn false (mk_eqn thy (is_constr thy) (thm, true)) thy;
 
 fun add_default_eqn thm thy =
-  case Code_Unit.mk_eqn_liberal thy (is_constr thy) thm
+  case mk_eqn_liberal thy (is_constr thy) thm
    of SOME eqn => gen_add_eqn true eqn thy
     | NONE => thy;
 
 fun add_nbe_eqn thm thy =
-  gen_add_eqn false (Code_Unit.mk_eqn thy (is_constr thy) (thm, false)) thy;
+  gen_add_eqn false (mk_eqn thy (is_constr thy) (thm, false)) thy;
 
 fun add_eqnl (c, lthms) thy =
   let
@@ -445,8 +878,8 @@
   (fn thm => Context.mapping (add_default_eqn thm) I);
 val add_default_eqn_attrib = Attrib.internal (K add_default_eqn_attribute);
 
-fun del_eqn thm thy = case Code_Unit.mk_eqn_liberal thy (is_constr thy) thm
- of SOME (thm, _) => change_eqns true (Code_Unit.const_eqn thy thm) (del_thm thm) thy
+fun del_eqn thm thy = case mk_eqn_liberal thy (is_constr thy) thm
+ of SOME (thm, _) => change_eqns true (const_eqn thy thm) (del_thm thm) thy
   | NONE => thy;
 
 fun del_eqns c = change_eqns true c (K (false, Lazy.value []));
@@ -460,7 +893,7 @@
 fun add_datatype raw_cs thy =
   let
     val cs = map (fn c_ty as (_, ty) => (AxClass.unoverload_const thy c_ty, ty)) raw_cs;
-    val (tyco, vs_cos) = Code_Unit.constrset_of_consts thy cs;
+    val (tyco, vs_cos) = constrset_of_consts thy cs;
     val old_cs = (map fst o snd o get_datatype thy) tyco;
     fun drop_outdated_cases cases = fold Symtab.delete_safe
       (Symtab.fold (fn (c, (_, (_, cos))) =>
@@ -480,12 +913,12 @@
 
 fun add_datatype_cmd raw_cs thy =
   let
-    val cs = map (Code_Unit.read_bare_const thy) raw_cs;
+    val cs = map (read_bare_const thy) raw_cs;
   in add_datatype cs thy end;
 
 fun add_case thm thy =
   let
-    val (c, (k, case_pats)) = Code_Unit.case_cert thm;
+    val (c, (k, case_pats)) = case_cert thm;
     val _ = case filter_out (is_constr thy) case_pats
      of [] => ()
       | cs => error ("Non-constructor(s) in case certificate: " ^ commas (map quote cs));
@@ -516,7 +949,7 @@
 
 fun default_typscheme thy c =
   let
-    fun the_const_typscheme c = (curry (Code_Unit.typscheme thy) c o snd o dest_Const
+    fun the_const_typscheme c = (curry (typscheme thy) c o snd o dest_Const
       o TermSubst.zero_var_indexes o curry Const "" o Sign.the_const_type thy) c;
     fun strip_sorts (vs, ty) = (map (fn (v, _) => (v, [])) vs, ty);
   in case AxClass.class_of_param thy c
@@ -524,7 +957,7 @@
     | NONE => if is_constr thy c
         then strip_sorts (the_const_typscheme c)
         else case get_eqns thy c
-         of (thm, _) :: _ => (Code_Unit.typscheme_eqn thy o Drule.zero_var_indexes) thm
+         of (thm, _) :: _ => (typscheme_eqn thy o Drule.zero_var_indexes) thm
           | [] => strip_sorts (the_const_typscheme c) end;
 
 end; (*struct*)
--- a/src/Pure/Isar/code_unit.ML	Tue May 19 14:13:23 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,447 +0,0 @@
-(*  Title:      Pure/Isar/code_unit.ML
-    Author:     Florian Haftmann, TU Muenchen
-
-Basic notions of code generation.  Auxiliary.
-*)
-
-signature CODE_UNIT =
-sig
-  (*typ instantiations*)
-  val typscheme: theory -> string * typ -> (string * sort) list * typ
-  val inst_thm: theory -> sort Vartab.table -> thm -> thm
-  val constrain_thm: theory -> sort -> thm -> thm
-
-  (*constant aliasses*)
-  val add_const_alias: thm -> theory -> theory
-  val triv_classes: theory -> class list
-  val resubst_alias: theory -> string -> string
-
-  (*constants*)
-  val string_of_typ: theory -> typ -> string
-  val string_of_const: theory -> string -> string
-  val no_args: theory -> string -> int
-  val check_const: theory -> term -> string
-  val read_bare_const: theory -> string -> string * typ
-  val read_const: theory -> string -> string
-
-  (*constructor sets*)
-  val constrset_of_consts: theory -> (string * typ) list
-    -> string * ((string * sort) list * (string * typ list) list)
-
-  (*code equations*)
-  val mk_eqn: theory -> (string -> bool) -> thm * bool -> thm * bool
-  val mk_eqn_liberal: theory -> (string -> bool) -> thm -> (thm * bool) option
-  val assert_eqn: theory -> (string -> bool) -> thm * bool -> thm * bool
-  val const_eqn: theory -> thm -> string
-    val const_typ_eqn: thm -> string * typ
-  val typscheme_eqn: theory -> thm -> (string * sort) list * typ
-  val expand_eta: theory -> int -> thm -> thm
-  val rewrite_eqn: simpset -> thm -> thm
-  val rewrite_head: thm list -> thm -> thm
-  val norm_args: theory -> thm list -> thm list 
-  val norm_varnames: theory -> thm list -> thm list
-
-  (*case certificates*)
-  val case_cert: thm -> string * (int * string list)
-end;
-
-structure Code_Unit: CODE_UNIT =
-struct
-
-
-(* auxiliary *)
-
-fun string_of_typ thy = setmp show_sorts true (Syntax.string_of_typ_global thy);
-fun string_of_const thy c = case AxClass.inst_of_param thy c
- of SOME (c, tyco) => Sign.extern_const thy c ^ " " ^ enclose "[" "]" (Sign.extern_type thy tyco)
-  | NONE => Sign.extern_const thy c;
-
-fun no_args thy = length o fst o strip_type o Sign.the_const_type thy;
-
-
-(* utilities *)
-
-fun typscheme thy (c, ty) =
-  let
-    val ty' = Logic.unvarifyT ty;
-    fun dest (TFree (v, sort)) = (v, sort)
-      | dest ty = error ("Illegal type parameter in type scheme: " ^ Syntax.string_of_typ_global thy ty);
-    val vs = map dest (Sign.const_typargs thy (c, ty'));
-  in (vs, Type.strip_sorts ty') end;
-
-fun inst_thm thy tvars' thm =
-  let
-    val tvars = (Term.add_tvars o Thm.prop_of) thm [];
-    val inter_sort = Sorts.inter_sort (Sign.classes_of thy);
-    fun mk_inst (tvar as (v, sort)) = case Vartab.lookup tvars' v
-     of SOME sort' => SOME (pairself (Thm.ctyp_of thy o TVar)
-          (tvar, (v, inter_sort (sort, sort'))))
-      | NONE => NONE;
-    val insts = map_filter mk_inst tvars;
-  in Thm.instantiate (insts, []) thm end;
-
-fun constrain_thm thy sort thm =
-  let
-    val constrain = curry (Sorts.inter_sort (Sign.classes_of thy)) sort
-    val tvars = (Term.add_tvars o Thm.prop_of) thm [];
-    fun mk_inst (tvar as (v, sort)) = pairself (Thm.ctyp_of thy o TVar o pair v)
-      (sort, constrain sort)
-    val insts = map mk_inst tvars;
-  in Thm.instantiate (insts, []) thm end;
-
-fun expand_eta thy k thm =
-  let
-    val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm;
-    val (head, args) = strip_comb lhs;
-    val l = if k = ~1
-      then (length o fst o strip_abs) rhs
-      else Int.max (0, k - length args);
-    val used = Name.make_context (map (fst o fst) (Term.add_vars lhs []));
-    fun get_name _ 0 = pair []
-      | get_name (Abs (v, ty, t)) k =
-          Name.variants [v]
-          ##>> get_name t (k - 1)
-          #>> (fn ([v'], vs') => (v', ty) :: vs')
-      | get_name t k = 
-          let
-            val (tys, _) = (strip_type o fastype_of) t
-          in case tys
-           of [] => raise TERM ("expand_eta", [t])
-            | ty :: _ =>
-                Name.variants [""]
-                #-> (fn [v] => get_name (t $ Var ((v, 0), ty)) (k - 1)
-                #>> (fn vs' => (v, ty) :: vs'))
-          end;
-    val (vs, _) = get_name rhs l used;
-    fun expand (v, ty) thm = Drule.fun_cong_rule thm
-      (Thm.cterm_of thy (Var ((v, 0), ty)));
-  in
-    thm
-    |> fold expand vs
-    |> Conv.fconv_rule Drule.beta_eta_conversion
-  end;
-
-fun eqn_conv conv =
-  let
-    fun lhs_conv ct = if can Thm.dest_comb ct
-      then (Conv.combination_conv lhs_conv conv) ct
-      else Conv.all_conv ct;
-  in Conv.combination_conv (Conv.arg_conv lhs_conv) conv end;
-
-fun head_conv conv =
-  let
-    fun lhs_conv ct = if can Thm.dest_comb ct
-      then (Conv.fun_conv lhs_conv) ct
-      else conv ct;
-  in Conv.fun_conv (Conv.arg_conv lhs_conv) end;
-
-val rewrite_eqn = Conv.fconv_rule o eqn_conv o Simplifier.rewrite;
-val rewrite_head = Conv.fconv_rule o head_conv o MetaSimplifier.rewrite false;
-
-fun norm_args thy thms =
-  let
-    val num_args_of = length o snd o strip_comb o fst o Logic.dest_equals;
-    val k = fold (curry Int.max o num_args_of o Thm.prop_of) thms 0;
-  in
-    thms
-    |> map (expand_eta thy k)
-    |> map (Conv.fconv_rule Drule.beta_eta_conversion)
-  end;
-
-fun canonical_tvars thy thm =
-  let
-    val ctyp = Thm.ctyp_of thy;
-    val purify_tvar = unprefix "'" #> Name.desymbolize false #> prefix "'";
-    fun tvars_subst_for thm = (fold_types o fold_atyps)
-      (fn TVar (v_i as (v, _), sort) => let
-            val v' = purify_tvar v
-          in if v = v' then I
-          else insert (op =) (v_i, (v', sort)) end
-        | _ => I) (prop_of thm) [];
-    fun mk_inst (v_i, (v', sort)) (maxidx, acc) =
-      let
-        val ty = TVar (v_i, sort)
-      in
-        (maxidx + 1, (ctyp ty, ctyp (TVar ((v', maxidx), sort))) :: acc)
-      end;
-    val maxidx = Thm.maxidx_of thm + 1;
-    val (_, inst) = fold mk_inst (tvars_subst_for thm) (maxidx + 1, []);
-  in Thm.instantiate (inst, []) thm end;
-
-fun canonical_vars thy thm =
-  let
-    val cterm = Thm.cterm_of thy;
-    val purify_var = Name.desymbolize false;
-    fun vars_subst_for thm = fold_aterms
-      (fn Var (v_i as (v, _), ty) => let
-            val v' = purify_var v
-          in if v = v' then I
-          else insert (op =) (v_i, (v', ty)) end
-        | _ => I) (prop_of thm) [];
-    fun mk_inst (v_i as (v, i), (v', ty)) (maxidx, acc) =
-      let
-        val t = Var (v_i, ty)
-      in
-        (maxidx + 1, (cterm t, cterm (Var ((v', maxidx), ty))) :: acc)
-      end;
-    val maxidx = Thm.maxidx_of thm + 1;
-    val (_, inst) = fold mk_inst (vars_subst_for thm) (maxidx + 1, []);
-  in Thm.instantiate ([], inst) thm end;
-
-fun canonical_absvars thm =
-  let
-    val t = Thm.plain_prop_of thm;
-    val purify_var = Name.desymbolize false;
-    val t' = Term.map_abs_vars purify_var t;
-  in Thm.rename_boundvars t t' thm end;
-
-fun norm_varnames thy thms =
-  let
-    fun burrow_thms f [] = []
-      | burrow_thms f thms =
-          thms
-          |> Conjunction.intr_balanced
-          |> f
-          |> Conjunction.elim_balanced (length thms)
-  in
-    thms
-    |> map (canonical_vars thy)
-    |> map canonical_absvars
-    |> map Drule.zero_var_indexes
-    |> burrow_thms (canonical_tvars thy)
-    |> Drule.zero_var_indexes_list
-  end;
-
-
-(* const aliasses *)
-
-structure ConstAlias = TheoryDataFun
-(
-  type T = ((string * string) * thm) list * class list;
-  val empty = ([], []);
-  val copy = I;
-  val extend = I;
-  fun merge _ ((alias1, classes1), (alias2, classes2)) : T =
-    (Library.merge (eq_snd Thm.eq_thm_prop) (alias1, alias2),
-      Library.merge (op =) (classes1, classes2));
-);
-
-fun add_const_alias thm thy =
-  let
-    val lhs_rhs = case try Logic.dest_equals (Thm.prop_of thm)
-     of SOME lhs_rhs => lhs_rhs
-      | _ => error ("Not an equation: " ^ Display.string_of_thm thm);
-    val c_c' = case try (pairself (AxClass.unoverload_const thy o dest_Const)) lhs_rhs
-     of SOME c_c' => c_c'
-      | _ => error ("Not an equation with two constants: " ^ Display.string_of_thm thm);
-    val some_class = the_list (AxClass.class_of_param thy (snd c_c'));
-  in thy |>
-    ConstAlias.map (fn (alias, classes) =>
-      ((c_c', thm) :: alias, fold (insert (op =)) some_class classes))
-  end;
-
-fun resubst_alias thy =
-  let
-    val alias = fst (ConstAlias.get thy);
-    val subst_inst_param = Option.map fst o AxClass.inst_of_param thy;
-    fun subst_alias c =
-      get_first (fn ((c', c''), _) => if c = c'' then SOME c' else NONE) alias;
-  in
-    perhaps subst_inst_param
-    #> perhaps subst_alias
-  end;
-
-val triv_classes = snd o ConstAlias.get;
-
-
-(* reading constants as terms *)
-
-fun check_bare_const thy t = case try dest_Const t
- of SOME c_ty => c_ty
-  | NONE => error ("Not a constant: " ^ Syntax.string_of_term_global thy t);
-
-fun check_const thy = AxClass.unoverload_const thy o check_bare_const thy;
-
-fun read_bare_const thy = check_bare_const thy o Syntax.read_term_global thy;
-
-fun read_const thy = AxClass.unoverload_const thy o read_bare_const thy;
-
-
-(* constructor sets *)
-
-fun constrset_of_consts thy cs =
-  let
-    val _ = map (fn (c, _) => if (is_some o AxClass.class_of_param thy) c
-      then error ("Is a class parameter: " ^ string_of_const thy c) else ()) cs;
-    fun no_constr (c, ty) = error ("Not a datatype constructor: " ^ string_of_const thy c
-      ^ " :: " ^ string_of_typ thy ty);
-    fun last_typ c_ty ty =
-      let
-        val frees = OldTerm.typ_tfrees ty;
-        val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) ty
-          handle TYPE _ => no_constr c_ty
-        val _ = if has_duplicates (eq_fst (op =)) vs then no_constr c_ty else ();
-        val _ = if length frees <> length vs then no_constr c_ty else ();
-      in (tyco, vs) end;
-    fun ty_sorts (c, ty) =
-      let
-        val ty_decl = (Logic.unvarifyT o Sign.the_const_type thy) c;
-        val (tyco, _) = last_typ (c, ty) ty_decl;
-        val (_, vs) = last_typ (c, ty) ty;
-      in ((tyco, map snd vs), (c, (map fst vs, ty))) end;
-    fun add ((tyco', sorts'), c) ((tyco, sorts), cs) =
-      let
-        val _ = if tyco' <> tyco
-          then error "Different type constructors in constructor set"
-          else ();
-        val sorts'' = map2 (curry (Sorts.inter_sort (Sign.classes_of thy))) sorts' sorts
-      in ((tyco, sorts), c :: cs) end;
-    fun inst vs' (c, (vs, ty)) =
-      let
-        val the_v = the o AList.lookup (op =) (vs ~~ vs');
-        val ty' = map_atyps (fn TFree (v, _) => TFree (the_v v)) ty;
-      in (c, (fst o strip_type) ty') end;
-    val c' :: cs' = map ty_sorts cs;
-    val ((tyco, sorts), cs'') = fold add cs' (apsnd single c');
-    val vs = Name.names Name.context Name.aT sorts;
-    val cs''' = map (inst vs) cs'';
-  in (tyco, (vs, rev cs''')) end;
-
-
-(* code equations *)
-
-exception BAD_THM of string;
-fun bad_thm msg = raise BAD_THM msg;
-fun error_thm f thm = f thm handle BAD_THM msg => error msg;
-fun try_thm f thm = SOME (f thm) handle BAD_THM _ => NONE;
-
-fun is_linear thm =
-  let val (_, args) = (strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm
-  in not (has_duplicates (op =) ((fold o fold_aterms)
-    (fn Var (v, _) => cons v | _ => I) args [])) end;
-
-fun gen_assert_eqn thy is_constr_head is_constr_pat (thm, proper) =
-  let
-    val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm
-      handle TERM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm thm)
-           | THM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm thm);
-    fun vars_of t = fold_aterms (fn Var (v, _) => insert (op =) v
-      | Free _ => bad_thm ("Illegal free variable in equation\n"
-          ^ Display.string_of_thm thm)
-      | _ => I) t [];
-    fun tvars_of t = fold_term_types (fn _ =>
-      fold_atyps (fn TVar (v, _) => insert (op =) v
-        | TFree _ => bad_thm 
-      ("Illegal free type variable in equation\n" ^ Display.string_of_thm thm))) t [];
-    val lhs_vs = vars_of lhs;
-    val rhs_vs = vars_of rhs;
-    val lhs_tvs = tvars_of lhs;
-    val rhs_tvs = tvars_of rhs;
-    val _ = if null (subtract (op =) lhs_vs rhs_vs)
-      then ()
-      else bad_thm ("Free variables on right hand side of equation\n"
-        ^ Display.string_of_thm thm);
-    val _ = if null (subtract (op =) lhs_tvs rhs_tvs)
-      then ()
-      else bad_thm ("Free type variables on right hand side of equation\n"
-        ^ Display.string_of_thm thm)    val (head, args) = (strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm;
-    val (c, ty) = case head
-     of Const (c_ty as (_, ty)) => (AxClass.unoverload_const thy c_ty, ty)
-      | _ => bad_thm ("Equation not headed by constant\n" ^ Display.string_of_thm thm);
-    fun check _ (Abs _) = bad_thm
-          ("Abstraction on left hand side of equation\n"
-            ^ Display.string_of_thm thm)
-      | check 0 (Var _) = ()
-      | check _ (Var _) = bad_thm
-          ("Variable with application on left hand side of equation\n"
-            ^ Display.string_of_thm thm)
-      | check n (t1 $ t2) = (check (n+1) t1; check 0 t2)
-      | check n (Const (c_ty as (c, ty))) = if n = (length o fst o strip_type) ty
-          then if not proper orelse is_constr_pat (AxClass.unoverload_const thy c_ty)
-            then ()
-            else bad_thm (quote c ^ " is not a constructor, on left hand side of equation\n"
-              ^ Display.string_of_thm thm)
-          else bad_thm
-            ("Partially applied constant " ^ quote c ^ " on left hand side of equation\n"
-               ^ Display.string_of_thm thm);
-    val _ = map (check 0) args;
-    val _ = if not proper orelse is_linear thm then ()
-      else bad_thm ("Duplicate variables on left hand side of equation\n"
-        ^ Display.string_of_thm thm);
-    val _ = if (is_none o AxClass.class_of_param thy) c
-      then ()
-      else bad_thm ("Polymorphic constant as head in equation\n"
-        ^ Display.string_of_thm thm)
-    val _ = if not (is_constr_head c)
-      then ()
-      else bad_thm ("Constructor as head in equation\n"
-        ^ Display.string_of_thm thm)
-    val ty_decl = Sign.the_const_type thy c;
-    val _ = if Sign.typ_equiv thy (Type.strip_sorts ty_decl, Type.strip_sorts ty)
-      then () else bad_thm ("Type\n" ^ string_of_typ thy ty
-           ^ "\nof equation\n"
-           ^ Display.string_of_thm thm
-           ^ "\nis incompatible with declared function type\n"
-           ^ string_of_typ thy ty_decl)
-  in (thm, proper) end;
-
-fun assert_eqn thy is_constr = error_thm (gen_assert_eqn thy is_constr is_constr);
-
-val const_typ_eqn = dest_Const o fst o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of;
-
-fun typscheme_eqn thy = typscheme thy o const_typ_eqn;
-
-(*these are permissive wrt. to overloaded constants!*)
-fun mk_eqn thy is_constr_head = error_thm (gen_assert_eqn thy is_constr_head (K true)) o
-  apfst (LocalDefs.meta_rewrite_rule (ProofContext.init thy));
-
-fun mk_eqn_liberal thy is_constr_head = Option.map (fn (thm, _) => (thm, is_linear thm))
-  o try_thm (gen_assert_eqn thy is_constr_head (K true))
-  o rpair false o LocalDefs.meta_rewrite_rule (ProofContext.init thy);
-
-fun const_eqn thy = AxClass.unoverload_const thy o const_typ_eqn;
-
-
-(* case cerificates *)
-
-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;
-    val _ = case head of Free _ => true
-      | Var _ => true
-      | _ => raise TERM ("case_cert", []);
-    val ([(case_var, _)], case_expr) = Term.strip_abs_eta 1 raw_case_expr;
-    val (Const (case_const, _), raw_params) = strip_comb case_expr;
-    val n = find_index (fn Free (v, _) => v = case_var | _ => false) raw_params;
-    val _ = if n = ~1 then raise TERM ("case_cert", []) else ();
-    val params = map (fst o dest_Var) (nth_drop n raw_params);
-    fun dest_case t =
-      let
-        val (head' $ t_co, rhs) = Logic.dest_equals t;
-        val _ = if head' = head then () else raise TERM ("case_cert", []);
-        val (Const (co, _), args) = strip_comb t_co;
-        val (Var (param, _), args') = strip_comb rhs;
-        val _ = if args' = args then () else raise TERM ("case_cert", []);
-      in (param, co) end;
-    fun analyze_cases cases =
-      let
-        val co_list = fold (AList.update (op =) o dest_case) cases [];
-      in map (the o AList.lookup (op =) co_list) params end;
-    fun analyze_let t =
-      let
-        val (head' $ arg, Var (param', _) $ arg') = Logic.dest_equals t;
-        val _ = if head' = head then () else raise TERM ("case_cert", []);
-        val _ = if arg' = arg then () else raise TERM ("case_cert", []);
-        val _ = if [param'] = params then () else raise TERM ("case_cert", []);
-      in [] end;
-    fun analyze (cases as [let_case]) =
-          (analyze_cases cases handle Bind => analyze_let let_case)
-      | analyze cases = analyze_cases cases;
-  in (case_const, (n, analyze cases)) end;
-
-fun case_cert thm = case_certificate thm
-  handle Bind => error "bad case certificate"
-       | TERM _ => error "bad case certificate";
-
-end;
--- a/src/Pure/Thy/thm_deps.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/Pure/Thy/thm_deps.ML	Tue May 19 14:13:37 2009 +0100
@@ -66,7 +66,7 @@
       case Thm.get_group thm of
         NONE => I | SOME grp => Symtab.update (grp, ())) thms Symtab.empty;
     val (thms', _) = fold (fn p as (s, thm) => fn q as (thms, grps') =>
-      if member (op =) [Thm.theoremK, Thm.lemmaK, Thm.corollaryK, Thm.axiomK] (Thm.get_kind thm)
+      if member (op =) [Thm.theoremK, Thm.generatedK, Thm.lemmaK, Thm.corollaryK, Thm.axiomK] (Thm.get_kind thm)
         andalso is_unused p
       then
         (case Thm.get_group thm of
--- a/src/Pure/codegen.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/Pure/codegen.ML	Tue May 19 14:13:37 2009 +0100
@@ -75,7 +75,7 @@
   val mk_type: bool -> typ -> Pretty.T
   val mk_term_of: codegr -> string -> bool -> typ -> Pretty.T
   val mk_gen: codegr -> string -> bool -> string list -> string -> typ -> Pretty.T
-  val test_fn: (int -> (string * term) list option) ref
+  val test_fn: (int -> term list option) ref
   val test_term: Proof.context -> term -> int -> term list option
   val eval_result: (unit -> term) ref
   val eval_term: theory -> term -> term
@@ -329,7 +329,7 @@
   end;
 
 val assoc_const_i = gen_assoc_const (K I);
-val assoc_const = gen_assoc_const Code_Unit.read_bare_const;
+val assoc_const = gen_assoc_const Code.read_bare_const;
 
 
 (**** associate types with target language types ****)
@@ -871,39 +871,35 @@
               [mk_gen gr module true xs a T, mk_type true T]) Ts) @
          (if member (op =) xs s then [str a] else []))));
 
-val test_fn : (int -> (string * term) list option) ref = ref (fn _ => NONE);
+val test_fn : (int -> term list option) ref = ref (fn _ => NONE);
 
 fun test_term ctxt t =
   let
     val thy = ProofContext.theory_of ctxt;
     val (code, gr) = setmp mode ["term_of", "test"]
       (generate_code_i thy [] "Generated") [("testf", t)];
-    val frees = Name.names Name.context "a" ((map snd o fst o strip_abs) t);
-    val frees' = frees ~~
-      map (fn i => "arg" ^ string_of_int i) (1 upto length frees);
+    val Ts = map snd (fst (strip_abs t));
+    val args = map_index (fn (i, T) => ("arg" ^ string_of_int i, T)) Ts;
     val s = "structure TestTerm =\nstruct\n\n" ^
       cat_lines (map snd code) ^
       "\nopen Generated;\n\n" ^ string_of
         (Pretty.block [str "val () = Codegen.test_fn :=",
           Pretty.brk 1, str ("(fn i =>"), Pretty.brk 1,
-          mk_let (map (fn ((s, T), s') =>
-              (mk_tuple [str s', str (s' ^ "_t")],
+          mk_let (map (fn (s, T) =>
+              (mk_tuple [str s, str (s ^ "_t")],
                Pretty.block [mk_gen gr "Generated" false [] "" T, Pretty.brk 1,
-                 str "(i + 1)"])) frees')
+                 str "i"])) args)
             (Pretty.block [str "if ",
-              mk_app false (str "testf") (map (str o snd) frees'),
+              mk_app false (str "testf") (map (str o fst) args),
               Pretty.brk 1, str "then NONE",
               Pretty.brk 1, str "else ",
               Pretty.block [str "SOME ", Pretty.block (str "[" ::
-                flat (separate [str ",", Pretty.brk 1]
-                  (map (fn ((s, T), s') => [Pretty.block
-                    [str ("(" ^ quote (Symbol.escape s) ^ ","), Pretty.brk 1,
-                     str (s' ^ "_t ())")]]) frees')) @
+                Pretty.commas (map (fn (s, _) => str (s ^ "_t ()")) args) @
                   [str "]"])]]),
           str ");"]) ^
       "\n\nend;\n";
     val _ = ML_Context.eval_in (SOME ctxt) false Position.none s;
-  in ! test_fn #> (Option.map o map) snd end;
+  in ! test_fn end;
 
 
 
--- a/src/Pure/more_thm.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/Pure/more_thm.ML	Tue May 19 14:13:37 2009 +0100
@@ -70,6 +70,7 @@
   val assumptionK: string
   val definitionK: string
   val theoremK: string
+  val generatedK : string
   val lemmaK: string
   val corollaryK: string
   val internalK: string
@@ -388,6 +389,7 @@
 val assumptionK = "assumption";
 val definitionK = "definition";
 val theoremK = "theorem";
+val generatedK = "generatedK"
 val lemmaK = "lemma";
 val corollaryK = "corollary";
 val internalK = Markup.internalK;
--- a/src/Tools/code/code_haskell.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/Tools/code/code_haskell.ML	Tue May 19 14:13:37 2009 +0100
@@ -480,7 +480,7 @@
 
 fun add_monad target' raw_c_bind thy =
   let
-    val c_bind = Code_Unit.read_const thy raw_c_bind;
+    val c_bind = Code.read_const thy raw_c_bind;
   in if target = target' then
     thy
     |> Code_Target.add_syntax_const target c_bind
--- a/src/Tools/code/code_ml.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/Tools/code/code_ml.ML	Tue May 19 14:13:37 2009 +0100
@@ -996,7 +996,7 @@
     val (ml_code, target_names) = eval_code_of NONE thy naming program (consts' @ tycos');
     val (consts'', tycos'') = chop (length consts') target_names;
     val consts_map = map2 (fn const => fn NONE =>
-        error ("Constant " ^ (quote o Code_Unit.string_of_const thy) const
+        error ("Constant " ^ (quote o Code.string_of_const thy) const
           ^ "\nhas a user-defined serialization")
       | SOME const'' => (const, const'')) consts consts''
     val tycos_map = map2 (fn tyco => fn NONE =>
@@ -1050,7 +1050,7 @@
 
 fun ml_code_antiq raw_const {struct_name, background} =
   let
-    val const = Code_Unit.check_const (ProofContext.theory_of background) raw_const;
+    val const = Code.check_const (ProofContext.theory_of background) raw_const;
     val is_first = is_first_occ background;
     val background' = register_const const background;
   in (print_code struct_name is_first (print_const const), background') end;
@@ -1059,7 +1059,7 @@
   let
     val thy = ProofContext.theory_of background;
     val tyco = Sign.intern_type thy raw_tyco;
-    val constrs = map (Code_Unit.check_const thy) raw_constrs;
+    val constrs = map (Code.check_const thy) raw_constrs;
     val constrs' = (map fst o snd o Code.get_datatype thy) tyco;
     val _ = if gen_eq_set (op =) (constrs, constrs') then ()
       else error ("Type " ^ quote tyco ^ ": given constructors diverge from real constructors")
--- a/src/Tools/code/code_preproc.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/Tools/code/code_preproc.ML	Tue May 19 14:13:37 2009 +0100
@@ -117,11 +117,11 @@
   in
     eqns
     |> apply_functrans thy c functrans
-    |> (map o apfst) (Code_Unit.rewrite_eqn pre)
+    |> (map o apfst) (Code.rewrite_eqn pre)
     |> (map o apfst) (AxClass.unoverload thy)
     |> map (Code.assert_eqn thy)
-    |> burrow_fst (Code_Unit.norm_args thy)
-    |> burrow_fst (Code_Unit.norm_varnames thy)
+    |> burrow_fst (Code.norm_args thy)
+    |> burrow_fst (Code.norm_varnames thy)
   end;
 
 fun preprocess_conv thy ct =
@@ -186,7 +186,7 @@
 
 fun pretty thy eqngr =
   AList.make (snd o Graph.get_node eqngr) (Graph.keys eqngr)
-  |> (map o apfst) (Code_Unit.string_of_const thy)
+  |> (map o apfst) (Code.string_of_const thy)
   |> sort (string_ord o pairself fst)
   |> map (fn (s, thms) =>
        (Pretty.block o Pretty.fbreaks) (
@@ -216,7 +216,7 @@
 fun tyscm_rhss_of thy c eqns =
   let
     val tyscm = case eqns of [] => Code.default_typscheme thy c
-      | ((thm, _) :: _) => Code_Unit.typscheme_eqn thy thm;
+      | ((thm, _) :: _) => Code.typscheme_eqn thy thm;
     val rhss = consts_of thy eqns;
   in (tyscm, rhss) end;
 
@@ -394,7 +394,7 @@
     val inst_tab = Vartab.empty |> fold (fn (v, sort) =>
       Vartab.update ((v, 0), sort)) lhs;
     val eqns = proto_eqns
-      |> (map o apfst) (Code_Unit.inst_thm thy inst_tab);
+      |> (map o apfst) (Code.inst_thm thy inst_tab);
     val (tyscm, rhss') = tyscm_rhss_of thy c eqns;
     val eqngr' = Graph.new_node (c, (tyscm, eqns)) eqngr;
   in (map (pair c) rhss' @ rhss, eqngr') end;
--- a/src/Tools/code/code_target.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/Tools/code/code_target.ML	Tue May 19 14:13:37 2009 +0100
@@ -286,7 +286,7 @@
 fun gen_add_syntax_const prep_const target raw_c raw_syn thy =
   let
     val c = prep_const thy raw_c;
-    fun check_args (syntax as (n, _)) = if n > Code_Unit.no_args thy c
+    fun check_args (syntax as (n, _)) = if n > Code.no_args thy c
       then error ("Too many arguments in syntax for constant " ^ quote c)
       else syntax;
   in case raw_syn
@@ -319,8 +319,8 @@
       | add (name, NONE) incls = Symtab.delete name incls;
   in map_includes target (add args) thy end;
 
-val add_include = gen_add_include Code_Unit.check_const;
-val add_include_cmd = gen_add_include Code_Unit.read_const;
+val add_include = gen_add_include Code.check_const;
+val add_include_cmd = gen_add_include Code.read_const;
 
 fun add_module_alias target (thyname, modlname) =
   let
@@ -363,11 +363,11 @@
 val allow_abort = gen_allow_abort (K I);
 val add_reserved = add_reserved;
 
-val add_syntax_class_cmd = gen_add_syntax_class read_class Code_Unit.read_const;
+val add_syntax_class_cmd = gen_add_syntax_class read_class Code.read_const;
 val add_syntax_inst_cmd = gen_add_syntax_inst read_class read_tyco;
 val add_syntax_tyco_cmd = gen_add_syntax_tyco read_tyco;
-val add_syntax_const_cmd = gen_add_syntax_const Code_Unit.read_const;
-val allow_abort_cmd = gen_allow_abort Code_Unit.read_const;
+val add_syntax_const_cmd = gen_add_syntax_const Code.read_const;
+val allow_abort_cmd = gen_allow_abort Code.read_const;
 
 fun the_literals thy =
   let
@@ -387,15 +387,15 @@
 local
 
 fun labelled_name thy program name = case Graph.get_node program name
- of Code_Thingol.Fun (c, _) => quote (Code_Unit.string_of_const thy c)
+ of Code_Thingol.Fun (c, _) => quote (Code.string_of_const thy c)
   | Code_Thingol.Datatype (tyco, _) => "type " ^ quote (Sign.extern_type thy tyco)
-  | Code_Thingol.Datatypecons (c, _) => quote (Code_Unit.string_of_const thy c)
+  | Code_Thingol.Datatypecons (c, _) => quote (Code.string_of_const thy c)
   | Code_Thingol.Class (class, _) => "class " ^ quote (Sign.extern_class thy class)
   | Code_Thingol.Classrel (sub, super) => let
         val Code_Thingol.Class (sub, _) = Graph.get_node program sub
         val Code_Thingol.Class (super, _) = Graph.get_node program super
       in quote (Sign.extern_class thy sub ^ " < " ^ Sign.extern_class thy super) end
-  | Code_Thingol.Classparam (c, _) => quote (Code_Unit.string_of_const thy c)
+  | Code_Thingol.Classparam (c, _) => quote (Code.string_of_const thy c)
   | Code_Thingol.Classinst ((class, (tyco, _)), _) => let
         val Code_Thingol.Class (class, _) = Graph.get_node program class
         val Code_Thingol.Datatype (tyco, _) = Graph.get_node program tyco
--- a/src/Tools/code/code_thingol.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/Tools/code/code_thingol.ML	Tue May 19 14:13:37 2009 +0100
@@ -498,7 +498,7 @@
       let
         val thms = if null (Term.add_tfreesT ty []) orelse (null o fst o strip_type) ty
           then raw_thms
-          else (map o apfst) (Code_Unit.expand_eta thy 1) raw_thms;
+          else (map o apfst) (Code.expand_eta thy 1) raw_thms;
       in
         fold_map (translate_tyvar_sort thy algbr funcgr) vs
         ##>> translate_typ thy algbr funcgr ty
@@ -634,7 +634,7 @@
         Term.strip_abs_eta 1 (the_single ts_clause)
       in [(true, (Free v_ty, body))] end
       else map (uncurry mk_clause)
-        (AList.make (Code_Unit.no_args thy) case_pats ~~ ts_clause);
+        (AList.make (Code.no_args thy) case_pats ~~ ts_clause);
     fun retermify ty (_, (IVar x, body)) =
           (x, ty) `|-> body
       | retermify _ (_, (pat, body)) =
@@ -812,7 +812,7 @@
     fun read_const_expr "*" = ([], consts_of NONE)
       | read_const_expr s = if String.isSuffix ".*" s
           then ([], consts_of (SOME (unsuffix ".*" s)))
-          else ([Code_Unit.read_const thy s], []);
+          else ([Code.read_const thy s], []);
   in pairself flat o split_list o map read_const_expr end;
 
 fun code_depgr thy consts =
@@ -839,7 +839,7 @@
       |> map (the o Symtab.lookup mapping)
       |> distinct (op =);
     val conn = [] |> fold (fn consts => cons (consts, succs consts)) constss;
-    fun namify consts = map (Code_Unit.string_of_const thy) consts
+    fun namify consts = map (Code.string_of_const thy) consts
       |> commas;
     val prgr = map (fn (consts, constss) =>
       { name = namify consts, ID = namify consts, dir = "", unfold = true,
--- a/src/Tools/nbe.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/Tools/nbe.ML	Tue May 19 14:13:37 2009 +0100
@@ -436,7 +436,7 @@
   let
     fun subst_const f = map_aterms (fn t as Term.Const (c, ty) => Term.Const (f c, ty)
       | t => t);
-    val resubst_triv_consts = subst_const (Code_Unit.resubst_alias thy);
+    val resubst_triv_consts = subst_const (Code.resubst_alias thy);
     val ty' = typ_of_itype program vs0 ty;
     fun type_infer t =
       singleton (TypeInfer.infer_types (Syntax.pp_global thy) (Sign.tsig_of thy) I
@@ -459,7 +459,7 @@
 (* evaluation oracle *)
 
 fun add_triv_classes thy = curry (Sorts.inter_sort (Sign.classes_of thy))
-  (Code_Unit.triv_classes thy);
+  (Code.triv_classes thy);
 
 fun mk_equals thy lhs raw_rhs =
   let
--- a/src/Tools/quickcheck.ML	Tue May 19 14:13:23 2009 +0100
+++ b/src/Tools/quickcheck.ML	Tue May 19 14:13:37 2009 +0100
@@ -94,7 +94,7 @@
       error "Term to be tested contains type variables";
     val _ = null (Term.add_vars t []) orelse
       error "Term to be tested contains schematic variables";
-    val frees = map dest_Free (OldTerm.term_frees t);
+    val frees = Term.add_frees t [];
   in (map fst frees, list_abs_free (frees, t)) end
 
 fun test_term ctxt quiet generator_name size i t =
@@ -104,12 +104,12 @@
      of NONE => if quiet then mk_testers ctxt t' else mk_testers_strict ctxt t'
       | SOME name => [mk_tester_select name ctxt t'];
     fun iterate f 0 = NONE
-      | iterate f k = case f () handle Match => (if quiet then ()
+      | iterate f j = case f () handle Match => (if quiet then ()
              else warning "Exception Match raised during quickcheck"; NONE)
-          of NONE => iterate f (k - 1) | SOME q => SOME q;
+          of NONE => iterate f (j - 1) | SOME q => SOME q;
     fun with_testers k [] = NONE
       | with_testers k (tester :: testers) =
-          case iterate (fn () => tester k) i
+          case iterate (fn () => tester (k - 1)) i
            of NONE => with_testers k testers
             | SOME q => SOME q;
     fun with_size k = if k > size then NONE