merged
authorhaftmann
Wed, 23 Sep 2009 14:00:43 +0200
changeset 32701 5059a733a4b8
parent 32657 5f13912245ff (diff)
parent 32700 e743ca6e97e7 (current diff)
child 32702 457135bdd596
merged
src/HOL/Inductive.thy
src/HOL/ex/predicate_compile.ML
--- a/src/HOL/Code_Eval.thy	Wed Sep 23 11:34:21 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,271 +0,0 @@
-(*  Title:      HOL/Code_Eval.thy
-    Author:     Florian Haftmann, TU Muenchen
-*)
-
-header {* Term evaluation using the generic code generator *}
-
-theory Code_Eval
-imports Plain Typerep Code_Numeral
-begin
-
-subsection {* Term representation *}
-
-subsubsection {* Terms and class @{text term_of} *}
-
-datatype "term" = dummy_term
-
-definition Const :: "String.literal \<Rightarrow> typerep \<Rightarrow> term" where
-  "Const _ _ = dummy_term"
-
-definition App :: "term \<Rightarrow> term \<Rightarrow> term" where
-  "App _ _ = dummy_term"
-
-code_datatype Const App
-
-class term_of = typerep +
-  fixes term_of :: "'a \<Rightarrow> term"
-
-lemma term_of_anything: "term_of x \<equiv> t"
-  by (rule eq_reflection) (cases "term_of x", cases t, simp)
-
-definition valapp :: "('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)
-  \<Rightarrow> 'a \<times> (unit \<Rightarrow> term) \<Rightarrow> 'b \<times> (unit \<Rightarrow> term)" where
-  "valapp f x = (fst f (fst x), \<lambda>u. App (snd f ()) (snd x ()))"
-
-lemma valapp_code [code, code_unfold]:
-  "valapp (f, tf) (x, tx) = (f x, \<lambda>u. App (tf ()) (tx ()))"
-  by (simp only: valapp_def fst_conv snd_conv)
-
-
-subsubsection {* @{text term_of} instances *}
-
-instantiation "fun" :: (typerep, typerep) term_of
-begin
-
-definition
-  "term_of (f \<Colon> 'a \<Rightarrow> 'b) = Const (STR ''dummy_pattern'') (Typerep.Typerep (STR ''fun'')
-     [Typerep.typerep TYPE('a), Typerep.typerep TYPE('b)])"
-
-instance ..
-
-end
-
-setup {*
-let
-  fun add_term_of tyco raw_vs thy =
-    let
-      val vs = map (fn (v, _) => (v, @{sort typerep})) raw_vs;
-      val ty = Type (tyco, map TFree vs);
-      val lhs = Const (@{const_name term_of}, ty --> @{typ term})
-        $ Free ("x", ty);
-      val rhs = @{term "undefined \<Colon> term"};
-      val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
-      fun triv_name_of t = (fst o dest_Free o fst o strip_comb o fst
-        o HOLogic.dest_eq o HOLogic.dest_Trueprop) t ^ "_triv";
-    in
-      thy
-      |> TheoryTarget.instantiation ([tyco], vs, @{sort term_of})
-      |> `(fn lthy => Syntax.check_term lthy eq)
-      |-> (fn eq => Specification.definition (NONE, ((Binding.name (triv_name_of eq), []), eq)))
-      |> snd
-      |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
-    end;
-  fun ensure_term_of (tyco, (raw_vs, _)) thy =
-    let
-      val need_inst = not (can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort term_of})
-        andalso can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort typerep};
-    in if need_inst then add_term_of tyco raw_vs thy else thy end;
-in
-  Code.type_interpretation ensure_term_of
-end
-*}
-
-setup {*
-let
-  fun mk_term_of_eq thy ty vs tyco (c, tys) =
-    let
-      val t = list_comb (Const (c, tys ---> ty),
-        map Free (Name.names Name.context "a" tys));
-      val (arg, rhs) = pairself (Thm.cterm_of thy o map_types Logic.unvarifyT o Logic.varify)
-        (t, (map_aterms (fn t as Free (v, ty) => HOLogic.mk_term_of ty t | t => t) o HOLogic.reflect_term) t)
-      val cty = Thm.ctyp_of thy ty;
-    in
-      @{thm term_of_anything}
-      |> Drule.instantiate' [SOME cty] [SOME arg, SOME rhs]
-      |> Thm.varifyT
-    end;
-  fun add_term_of_code tyco raw_vs raw_cs thy =
-    let
-      val algebra = Sign.classes_of thy;
-      val vs = map (fn (v, sort) =>
-        (v, curry (Sorts.inter_sort algebra) @{sort typerep} sort)) raw_vs;
-      val ty = Type (tyco, map TFree vs);
-      val cs = (map o apsnd o map o map_atyps)
-        (fn TFree (v, _) => TFree (v, (the o AList.lookup (op =) vs) v)) raw_cs;
-      val const = AxClass.param_of_inst thy (@{const_name term_of}, tyco);
-      val eqs = map (mk_term_of_eq thy ty vs tyco) cs;
-   in
-      thy
-      |> Code.del_eqns const
-      |> fold Code.add_eqn eqs
-    end;
-  fun ensure_term_of_code (tyco, (raw_vs, cs)) thy =
-    let
-      val has_inst = can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort term_of};
-    in if has_inst then add_term_of_code tyco raw_vs cs thy else thy end;
-in
-  Code.type_interpretation ensure_term_of_code
-end
-*}
-
-
-subsubsection {* Code generator setup *}
-
-lemmas [code del] = term.recs term.cases term.size
-lemma [code, code del]: "eq_class.eq (t1\<Colon>term) t2 \<longleftrightarrow> eq_class.eq t1 t2" ..
-
-lemma [code, code del]: "(term_of \<Colon> typerep \<Rightarrow> term) = term_of" ..
-lemma [code, code del]: "(term_of \<Colon> term \<Rightarrow> term) = term_of" ..
-lemma [code, code del]: "(term_of \<Colon> String.literal \<Rightarrow> term) = term_of" ..
-lemma [code, code del]:
-  "(Code_Eval.term_of \<Colon> 'a::{type, term_of} Predicate.pred \<Rightarrow> Code_Eval.term) = Code_Eval.term_of" ..
-lemma [code, code del]:
-  "(Code_Eval.term_of \<Colon> 'a::{type, term_of} Predicate.seq \<Rightarrow> Code_Eval.term) = Code_Eval.term_of" ..
-
-lemma term_of_char [unfolded typerep_fun_def typerep_char_def typerep_nibble_def, code]: "Code_Eval.term_of c =
-    (let (n, m) = nibble_pair_of_char c
-  in Code_Eval.App (Code_Eval.App (Code_Eval.Const (STR ''String.char.Char'') (TYPEREP(nibble \<Rightarrow> nibble \<Rightarrow> char)))
-    (Code_Eval.term_of n)) (Code_Eval.term_of m))"
-  by (subst term_of_anything) rule 
-
-code_type "term"
-  (Eval "Term.term")
-
-code_const Const and App
-  (Eval "Term.Const/ ((_), (_))" and "Term.$/ ((_), (_))")
-
-code_const "term_of \<Colon> String.literal \<Rightarrow> term"
-  (Eval "HOLogic.mk'_message'_string")
-
-code_reserved Eval HOLogic
-
-
-subsubsection {* Syntax *}
-
-definition termify :: "'a \<Rightarrow> term" where
-  [code del]: "termify x = dummy_term"
-
-abbreviation valtermify :: "'a \<Rightarrow> 'a \<times> (unit \<Rightarrow> term)" where
-  "valtermify x \<equiv> (x, \<lambda>u. termify x)"
-
-setup {*
-let
-  fun map_default f xs =
-    let val ys = map f xs
-    in if exists is_some ys
-      then SOME (map2 the_default xs ys)
-      else NONE
-    end;
-  fun subst_termify_app (Const (@{const_name termify}, T), [t]) =
-        if not (Term.has_abs t)
-        then if fold_aterms (fn Const _ => I | _ => K false) t true
-          then SOME (HOLogic.reflect_term t)
-          else error "Cannot termify expression containing variables"
-        else error "Cannot termify expression containing abstraction"
-    | subst_termify_app (t, ts) = case map_default subst_termify ts
-       of SOME ts' => SOME (list_comb (t, ts'))
-        | NONE => NONE
-  and subst_termify (Abs (v, T, t)) = (case subst_termify t
-       of SOME t' => SOME (Abs (v, T, t'))
-        | NONE => NONE)
-    | subst_termify t = subst_termify_app (strip_comb t) 
-  fun check_termify ts ctxt = map_default subst_termify ts
-    |> Option.map (rpair ctxt)
-in
-  Context.theory_map (Syntax.add_term_check 0 "termify" check_termify)
-end;
-*}
-
-locale term_syntax
-begin
-
-notation App (infixl "<\<cdot>>" 70)
-  and valapp (infixl "{\<cdot>}" 70)
-
-end
-
-interpretation term_syntax .
-
-no_notation App (infixl "<\<cdot>>" 70)
-  and valapp (infixl "{\<cdot>}" 70)
-
-
-subsection {* Numeric types *}
-
-definition term_of_num :: "'a\<Colon>{semiring_div} \<Rightarrow> 'a\<Colon>{semiring_div} \<Rightarrow> term" where
-  "term_of_num two = (\<lambda>_. dummy_term)"
-
-lemma (in term_syntax) term_of_num_code [code]:
-  "term_of_num two k = (if k = 0 then termify Int.Pls
-    else (if k mod two = 0
-      then termify Int.Bit0 <\<cdot>> term_of_num two (k div two)
-      else termify Int.Bit1 <\<cdot>> term_of_num two (k div two)))"
-  by (auto simp add: term_of_anything Const_def App_def term_of_num_def Let_def)
-
-lemma (in term_syntax) term_of_nat_code [code]:
-  "term_of (n::nat) = termify (number_of :: int \<Rightarrow> nat) <\<cdot>> term_of_num (2::nat) n"
-  by (simp only: term_of_anything)
-
-lemma (in term_syntax) term_of_int_code [code]:
-  "term_of (k::int) = (if k = 0 then termify (0 :: int)
-    else if k > 0 then termify (number_of :: int \<Rightarrow> int) <\<cdot>> term_of_num (2::int) k
-      else termify (uminus :: int \<Rightarrow> int) <\<cdot>> (termify (number_of :: int \<Rightarrow> int) <\<cdot>> term_of_num (2::int) (- k)))"
-  by (simp only: term_of_anything)
-
-lemma (in term_syntax) term_of_code_numeral_code [code]:
-  "term_of (k::code_numeral) = termify (number_of :: int \<Rightarrow> code_numeral) <\<cdot>> term_of_num (2::code_numeral) k"
-  by (simp only: term_of_anything)
-
-subsection {* Obfuscate *}
-
-print_translation {*
-let
-  val term = Const ("<TERM>", dummyT);
-  fun tr1' [_, _] = term;
-  fun tr2' [] = term;
-in
-  [(@{const_syntax Const}, tr1'),
-    (@{const_syntax App}, tr1'),
-    (@{const_syntax dummy_term}, tr2')]
-end
-*}
-
-hide const dummy_term App valapp
-hide (open) const Const termify valtermify term_of term_of_num
-
-
-subsection {* Evaluation setup *}
-
-ML {*
-signature EVAL =
-sig
-  val eval_ref: (unit -> term) option ref
-  val eval_term: theory -> term -> term
-end;
-
-structure Eval : EVAL =
-struct
-
-val eval_ref = ref (NONE : (unit -> term) option);
-
-fun eval_term thy t =
-  Code_ML.eval NONE ("Eval.eval_ref", eval_ref) I thy (HOLogic.mk_term_of (fastype_of t) t) [];
-
-end;
-*}
-
-setup {*
-  Value.add_evaluator ("code", Eval.eval_term o ProofContext.theory_of)
-*}
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Code_Evaluation.thy	Wed Sep 23 14:00:43 2009 +0200
@@ -0,0 +1,271 @@
+(*  Title:      HOL/Code_Evaluation.thy
+    Author:     Florian Haftmann, TU Muenchen
+*)
+
+header {* Term evaluation using the generic code generator *}
+
+theory Code_Evaluation
+imports Plain Typerep Code_Numeral
+begin
+
+subsection {* Term representation *}
+
+subsubsection {* Terms and class @{text term_of} *}
+
+datatype "term" = dummy_term
+
+definition Const :: "String.literal \<Rightarrow> typerep \<Rightarrow> term" where
+  "Const _ _ = dummy_term"
+
+definition App :: "term \<Rightarrow> term \<Rightarrow> term" where
+  "App _ _ = dummy_term"
+
+code_datatype Const App
+
+class term_of = typerep +
+  fixes term_of :: "'a \<Rightarrow> term"
+
+lemma term_of_anything: "term_of x \<equiv> t"
+  by (rule eq_reflection) (cases "term_of x", cases t, simp)
+
+definition valapp :: "('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)
+  \<Rightarrow> 'a \<times> (unit \<Rightarrow> term) \<Rightarrow> 'b \<times> (unit \<Rightarrow> term)" where
+  "valapp f x = (fst f (fst x), \<lambda>u. App (snd f ()) (snd x ()))"
+
+lemma valapp_code [code, code_unfold]:
+  "valapp (f, tf) (x, tx) = (f x, \<lambda>u. App (tf ()) (tx ()))"
+  by (simp only: valapp_def fst_conv snd_conv)
+
+
+subsubsection {* @{text term_of} instances *}
+
+instantiation "fun" :: (typerep, typerep) term_of
+begin
+
+definition
+  "term_of (f \<Colon> 'a \<Rightarrow> 'b) = Const (STR ''dummy_pattern'') (Typerep.Typerep (STR ''fun'')
+     [Typerep.typerep TYPE('a), Typerep.typerep TYPE('b)])"
+
+instance ..
+
+end
+
+setup {*
+let
+  fun add_term_of tyco raw_vs thy =
+    let
+      val vs = map (fn (v, _) => (v, @{sort typerep})) raw_vs;
+      val ty = Type (tyco, map TFree vs);
+      val lhs = Const (@{const_name term_of}, ty --> @{typ term})
+        $ Free ("x", ty);
+      val rhs = @{term "undefined \<Colon> term"};
+      val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
+      fun triv_name_of t = (fst o dest_Free o fst o strip_comb o fst
+        o HOLogic.dest_eq o HOLogic.dest_Trueprop) t ^ "_triv";
+    in
+      thy
+      |> TheoryTarget.instantiation ([tyco], vs, @{sort term_of})
+      |> `(fn lthy => Syntax.check_term lthy eq)
+      |-> (fn eq => Specification.definition (NONE, ((Binding.name (triv_name_of eq), []), eq)))
+      |> snd
+      |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
+    end;
+  fun ensure_term_of (tyco, (raw_vs, _)) thy =
+    let
+      val need_inst = not (can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort term_of})
+        andalso can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort typerep};
+    in if need_inst then add_term_of tyco raw_vs thy else thy end;
+in
+  Code.type_interpretation ensure_term_of
+end
+*}
+
+setup {*
+let
+  fun mk_term_of_eq thy ty vs tyco (c, tys) =
+    let
+      val t = list_comb (Const (c, tys ---> ty),
+        map Free (Name.names Name.context "a" tys));
+      val (arg, rhs) = pairself (Thm.cterm_of thy o map_types Logic.unvarifyT o Logic.varify)
+        (t, (map_aterms (fn t as Free (v, ty) => HOLogic.mk_term_of ty t | t => t) o HOLogic.reflect_term) t)
+      val cty = Thm.ctyp_of thy ty;
+    in
+      @{thm term_of_anything}
+      |> Drule.instantiate' [SOME cty] [SOME arg, SOME rhs]
+      |> Thm.varifyT
+    end;
+  fun add_term_of_code tyco raw_vs raw_cs thy =
+    let
+      val algebra = Sign.classes_of thy;
+      val vs = map (fn (v, sort) =>
+        (v, curry (Sorts.inter_sort algebra) @{sort typerep} sort)) raw_vs;
+      val ty = Type (tyco, map TFree vs);
+      val cs = (map o apsnd o map o map_atyps)
+        (fn TFree (v, _) => TFree (v, (the o AList.lookup (op =) vs) v)) raw_cs;
+      val const = AxClass.param_of_inst thy (@{const_name term_of}, tyco);
+      val eqs = map (mk_term_of_eq thy ty vs tyco) cs;
+   in
+      thy
+      |> Code.del_eqns const
+      |> fold Code.add_eqn eqs
+    end;
+  fun ensure_term_of_code (tyco, (raw_vs, cs)) thy =
+    let
+      val has_inst = can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort term_of};
+    in if has_inst then add_term_of_code tyco raw_vs cs thy else thy end;
+in
+  Code.type_interpretation ensure_term_of_code
+end
+*}
+
+
+subsubsection {* Code generator setup *}
+
+lemmas [code del] = term.recs term.cases term.size
+lemma [code, code del]: "eq_class.eq (t1\<Colon>term) t2 \<longleftrightarrow> eq_class.eq t1 t2" ..
+
+lemma [code, code del]: "(term_of \<Colon> typerep \<Rightarrow> term) = term_of" ..
+lemma [code, code del]: "(term_of \<Colon> term \<Rightarrow> term) = term_of" ..
+lemma [code, code del]: "(term_of \<Colon> String.literal \<Rightarrow> term) = term_of" ..
+lemma [code, code del]:
+  "(Code_Evaluation.term_of \<Colon> 'a::{type, term_of} Predicate.pred \<Rightarrow> Code_Evaluation.term) = Code_Evaluation.term_of" ..
+lemma [code, code del]:
+  "(Code_Evaluation.term_of \<Colon> 'a::{type, term_of} Predicate.seq \<Rightarrow> Code_Evaluation.term) = Code_Evaluation.term_of" ..
+
+lemma term_of_char [unfolded typerep_fun_def typerep_char_def typerep_nibble_def, code]: "Code_Evaluation.term_of c =
+    (let (n, m) = nibble_pair_of_char c
+  in Code_Evaluation.App (Code_Evaluation.App (Code_Evaluation.Const (STR ''String.char.Char'') (TYPEREP(nibble \<Rightarrow> nibble \<Rightarrow> char)))
+    (Code_Evaluation.term_of n)) (Code_Evaluation.term_of m))"
+  by (subst term_of_anything) rule 
+
+code_type "term"
+  (Eval "Term.term")
+
+code_const Const and App
+  (Eval "Term.Const/ ((_), (_))" and "Term.$/ ((_), (_))")
+
+code_const "term_of \<Colon> String.literal \<Rightarrow> term"
+  (Eval "HOLogic.mk'_message'_string")
+
+code_reserved Eval HOLogic
+
+
+subsubsection {* Syntax *}
+
+definition termify :: "'a \<Rightarrow> term" where
+  [code del]: "termify x = dummy_term"
+
+abbreviation valtermify :: "'a \<Rightarrow> 'a \<times> (unit \<Rightarrow> term)" where
+  "valtermify x \<equiv> (x, \<lambda>u. termify x)"
+
+setup {*
+let
+  fun map_default f xs =
+    let val ys = map f xs
+    in if exists is_some ys
+      then SOME (map2 the_default xs ys)
+      else NONE
+    end;
+  fun subst_termify_app (Const (@{const_name termify}, T), [t]) =
+        if not (Term.has_abs t)
+        then if fold_aterms (fn Const _ => I | _ => K false) t true
+          then SOME (HOLogic.reflect_term t)
+          else error "Cannot termify expression containing variables"
+        else error "Cannot termify expression containing abstraction"
+    | subst_termify_app (t, ts) = case map_default subst_termify ts
+       of SOME ts' => SOME (list_comb (t, ts'))
+        | NONE => NONE
+  and subst_termify (Abs (v, T, t)) = (case subst_termify t
+       of SOME t' => SOME (Abs (v, T, t'))
+        | NONE => NONE)
+    | subst_termify t = subst_termify_app (strip_comb t) 
+  fun check_termify ts ctxt = map_default subst_termify ts
+    |> Option.map (rpair ctxt)
+in
+  Context.theory_map (Syntax.add_term_check 0 "termify" check_termify)
+end;
+*}
+
+locale term_syntax
+begin
+
+notation App (infixl "<\<cdot>>" 70)
+  and valapp (infixl "{\<cdot>}" 70)
+
+end
+
+interpretation term_syntax .
+
+no_notation App (infixl "<\<cdot>>" 70)
+  and valapp (infixl "{\<cdot>}" 70)
+
+
+subsection {* Numeric types *}
+
+definition term_of_num :: "'a\<Colon>{semiring_div} \<Rightarrow> 'a\<Colon>{semiring_div} \<Rightarrow> term" where
+  "term_of_num two = (\<lambda>_. dummy_term)"
+
+lemma (in term_syntax) term_of_num_code [code]:
+  "term_of_num two k = (if k = 0 then termify Int.Pls
+    else (if k mod two = 0
+      then termify Int.Bit0 <\<cdot>> term_of_num two (k div two)
+      else termify Int.Bit1 <\<cdot>> term_of_num two (k div two)))"
+  by (auto simp add: term_of_anything Const_def App_def term_of_num_def Let_def)
+
+lemma (in term_syntax) term_of_nat_code [code]:
+  "term_of (n::nat) = termify (number_of :: int \<Rightarrow> nat) <\<cdot>> term_of_num (2::nat) n"
+  by (simp only: term_of_anything)
+
+lemma (in term_syntax) term_of_int_code [code]:
+  "term_of (k::int) = (if k = 0 then termify (0 :: int)
+    else if k > 0 then termify (number_of :: int \<Rightarrow> int) <\<cdot>> term_of_num (2::int) k
+      else termify (uminus :: int \<Rightarrow> int) <\<cdot>> (termify (number_of :: int \<Rightarrow> int) <\<cdot>> term_of_num (2::int) (- k)))"
+  by (simp only: term_of_anything)
+
+lemma (in term_syntax) term_of_code_numeral_code [code]:
+  "term_of (k::code_numeral) = termify (number_of :: int \<Rightarrow> code_numeral) <\<cdot>> term_of_num (2::code_numeral) k"
+  by (simp only: term_of_anything)
+
+subsection {* Obfuscate *}
+
+print_translation {*
+let
+  val term = Const ("<TERM>", dummyT);
+  fun tr1' [_, _] = term;
+  fun tr2' [] = term;
+in
+  [(@{const_syntax Const}, tr1'),
+    (@{const_syntax App}, tr1'),
+    (@{const_syntax dummy_term}, tr2')]
+end
+*}
+
+hide const dummy_term App valapp
+hide (open) const Const termify valtermify term_of term_of_num
+
+
+subsection {* Evaluation setup *}
+
+ML {*
+signature EVAL =
+sig
+  val eval_ref: (unit -> term) option ref
+  val eval_term: theory -> term -> term
+end;
+
+structure Eval : EVAL =
+struct
+
+val eval_ref = ref (NONE : (unit -> term) option);
+
+fun eval_term thy t =
+  Code_ML.eval NONE ("Eval.eval_ref", eval_ref) I thy (HOLogic.mk_term_of (fastype_of t) t) [];
+
+end;
+*}
+
+setup {*
+  Value.add_evaluator ("code", Eval.eval_term o ProofContext.theory_of)
+*}
+
+end
--- a/src/HOL/Decision_Procs/Approximation.thy	Wed Sep 23 11:34:21 2009 +0200
+++ b/src/HOL/Decision_Procs/Approximation.thy	Wed Sep 23 14:00:43 2009 +0200
@@ -3246,12 +3246,13 @@
         = map (` (variable_of_bound o prop_of)) prems
 
       fun add_deps (name, bnds)
-        = Graph.add_deps_acyclic
-            (name, remove (op =) name (Term.add_free_names (prop_of bnds) []))
+        = Graph.add_deps_acyclic (name,
+            remove (op =) name (Term.add_free_names (prop_of bnds) []))
+
       val order = Graph.empty
                   |> fold Graph.new_node variable_bounds
                   |> fold add_deps variable_bounds
-                  |> Graph.topological_order |> rev
+                  |> Graph.strong_conn |> map the_single |> rev
                   |> map_filter (AList.lookup (op =) variable_bounds)
 
       fun prepend_prem th tac
@@ -3338,7 +3339,7 @@
                       etac @{thm meta_eqE},
                       rtac @{thm impI}] i)
       THEN Subgoal.FOCUS (fn {prems, ...} => reorder_bounds_tac prems i) @{context} i
-      THEN TRY (filter_prems_tac (K false) i)
+      THEN DETERM (TRY (filter_prems_tac (K false) i))
       THEN DETERM (Reflection.genreify_tac ctxt form_equations NONE i)
       THEN rewrite_interpret_form_tac ctxt prec splitting taylor i
       THEN gen_eval_tac eval_oracle ctxt i))
@@ -3350,7 +3351,7 @@
 
   fun mk_approx' prec t = (@{const "approx'"}
                          $ HOLogic.mk_number @{typ nat} prec
-                         $ t $ @{term "[] :: (float * float) list"})
+                         $ t $ @{term "[] :: (float * float) option list"})
 
   fun dest_result (Const (@{const_name "Some"}, _) $
                    ((Const (@{const_name "Pair"}, _)) $
--- a/src/HOL/Decision_Procs/ex/Approximation_Ex.thy	Wed Sep 23 11:34:21 2009 +0200
+++ b/src/HOL/Decision_Procs/ex/Approximation_Ex.thy	Wed Sep 23 14:00:43 2009 +0200
@@ -72,7 +72,9 @@
   shows "g / v * tan (35 * d) \<in> { 3 * d .. 3.1 * d }"
   using assms by (approximation 80)
 
-lemma "\<phi> \<in> { 0 .. 1 :: real } \<longrightarrow> \<phi> ^ 2 \<le> \<phi>"
-  by (approximation 30 splitting: \<phi>=1 taylor: \<phi> = 3)
+lemma "x \<in> { 0 .. 1 :: real } \<longrightarrow> x ^ 2 \<le> x"
+  by (approximation 30 splitting: x=1 taylor: x = 3)
+
+value [approximate] "10"
 
 end
--- a/src/HOL/Inductive.thy	Wed Sep 23 11:34:21 2009 +0200
+++ b/src/HOL/Inductive.thy	Wed Sep 23 14:00:43 2009 +0200
@@ -267,26 +267,6 @@
   Ball_def Bex_def
   induct_rulify_fallback
 
-ML {*
-val def_lfp_unfold = @{thm def_lfp_unfold}
-val def_gfp_unfold = @{thm def_gfp_unfold}
-val def_lfp_induct = @{thm def_lfp_induct}
-val def_coinduct = @{thm def_coinduct}
-val inf_bool_eq = @{thm inf_bool_eq} RS @{thm eq_reflection}
-val inf_fun_eq = @{thm inf_fun_eq} RS @{thm eq_reflection}
-val sup_bool_eq = @{thm sup_bool_eq} RS @{thm eq_reflection}
-val sup_fun_eq = @{thm sup_fun_eq} RS @{thm eq_reflection}
-val le_boolI = @{thm le_boolI}
-val le_boolI' = @{thm le_boolI'}
-val le_funI = @{thm le_funI}
-val le_boolE = @{thm le_boolE}
-val le_funE = @{thm le_funE}
-val le_boolD = @{thm le_boolD}
-val le_funD = @{thm le_funD}
-val le_bool_def = @{thm le_bool_def} RS @{thm eq_reflection}
-val le_fun_def = @{thm le_fun_def} RS @{thm eq_reflection}
-*}
-
 use "Tools/inductive.ML"
 setup Inductive.setup
 
--- a/src/HOL/IsaMakefile	Wed Sep 23 11:34:21 2009 +0200
+++ b/src/HOL/IsaMakefile	Wed Sep 23 14:00:43 2009 +0200
@@ -210,7 +210,7 @@
 
 MAIN_DEPENDENCIES = $(PLAIN_DEPENDENCIES) \
   ATP_Linkup.thy \
-  Code_Eval.thy \
+  Code_Evaluation.thy \
   Code_Numeral.thy \
   Equiv_Relations.thy \
   Groebner_Basis.thy \
--- a/src/HOL/Library/Code_Char.thy	Wed Sep 23 11:34:21 2009 +0200
+++ b/src/HOL/Library/Code_Char.thy	Wed Sep 23 14:00:43 2009 +0200
@@ -5,7 +5,7 @@
 header {* Code generation of pretty characters (and strings) *}
 
 theory Code_Char
-imports List Code_Eval Main
+imports List Code_Evaluation Main
 begin
 
 code_type char
@@ -32,7 +32,7 @@
   (OCaml "!((_ : char) = _)")
   (Haskell infixl 4 "==")
 
-code_const "Code_Eval.term_of \<Colon> char \<Rightarrow> term"
+code_const "Code_Evaluation.term_of \<Colon> char \<Rightarrow> term"
   (Eval "HOLogic.mk'_char/ (IntInf.fromInt/ (Char.ord/ _))")
 
 end
--- a/src/HOL/Library/Code_Integer.thy	Wed Sep 23 11:34:21 2009 +0200
+++ b/src/HOL/Library/Code_Integer.thy	Wed Sep 23 14:00:43 2009 +0200
@@ -100,7 +100,7 @@
 
 text {* Evaluation *}
 
-code_const "Code_Eval.term_of \<Colon> int \<Rightarrow> term"
+code_const "Code_Evaluation.term_of \<Colon> int \<Rightarrow> term"
   (Eval "HOLogic.mk'_number/ HOLogic.intT")
 
 end
\ No newline at end of file
--- a/src/HOL/Library/Efficient_Nat.thy	Wed Sep 23 11:34:21 2009 +0200
+++ b/src/HOL/Library/Efficient_Nat.thy	Wed Sep 23 14:00:43 2009 +0200
@@ -415,9 +415,9 @@
 text {* Evaluation *}
 
 lemma [code, code del]:
-  "(Code_Eval.term_of \<Colon> nat \<Rightarrow> term) = Code_Eval.term_of" ..
+  "(Code_Evaluation.term_of \<Colon> nat \<Rightarrow> term) = Code_Evaluation.term_of" ..
 
-code_const "Code_Eval.term_of \<Colon> nat \<Rightarrow> term"
+code_const "Code_Evaluation.term_of \<Colon> nat \<Rightarrow> term"
   (SML "HOLogic.mk'_number/ HOLogic.natT")
 
 
--- a/src/HOL/Library/Fin_Fun.thy	Wed Sep 23 11:34:21 2009 +0200
+++ b/src/HOL/Library/Fin_Fun.thy	Wed Sep 23 14:00:43 2009 +0200
@@ -311,17 +311,17 @@
 notation scomp (infixl "o\<rightarrow>" 60)
 
 definition (in term_syntax) valtermify_finfun_const ::
-  "'b\<Colon>typerep \<times> (unit \<Rightarrow> Code_Eval.term) \<Rightarrow> ('a\<Colon>typerep \<Rightarrow>\<^isub>f 'b) \<times> (unit \<Rightarrow> Code_Eval.term)" where
-  "valtermify_finfun_const y = Code_Eval.valtermify finfun_const {\<cdot>} y"
+  "'b\<Colon>typerep \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> ('a\<Colon>typerep \<Rightarrow>\<^isub>f 'b) \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
+  "valtermify_finfun_const y = Code_Evaluation.valtermify finfun_const {\<cdot>} y"
 
 definition (in term_syntax) valtermify_finfun_update_code ::
-  "'a\<Colon>typerep \<times> (unit \<Rightarrow> Code_Eval.term) \<Rightarrow> 'b\<Colon>typerep \<times> (unit \<Rightarrow> Code_Eval.term) \<Rightarrow> ('a \<Rightarrow>\<^isub>f 'b) \<times> (unit \<Rightarrow> Code_Eval.term) \<Rightarrow> ('a \<Rightarrow>\<^isub>f 'b) \<times> (unit \<Rightarrow> Code_Eval.term)" where
-  "valtermify_finfun_update_code x y f = Code_Eval.valtermify finfun_update_code {\<cdot>} f {\<cdot>} x {\<cdot>} y"
+  "'a\<Colon>typerep \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> 'b\<Colon>typerep \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> ('a \<Rightarrow>\<^isub>f 'b) \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> ('a \<Rightarrow>\<^isub>f 'b) \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
+  "valtermify_finfun_update_code x y f = Code_Evaluation.valtermify finfun_update_code {\<cdot>} f {\<cdot>} x {\<cdot>} y"
 
 instantiation finfun :: (random, random) random
 begin
 
-primrec random_finfun_aux :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> Random.seed \<Rightarrow> ('a \<Rightarrow>\<^isub>f 'b \<times> (unit \<Rightarrow> Code_Eval.term)) \<times> Random.seed" where
+primrec random_finfun_aux :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> Random.seed \<Rightarrow> ('a \<Rightarrow>\<^isub>f 'b \<times> (unit \<Rightarrow> Code_Evaluation.term)) \<times> Random.seed" where
     "random_finfun_aux 0 j = Quickcheck.collapse (Random.select_weight
        [(1, Quickcheck.random j o\<rightarrow> (\<lambda>y. Pair (valtermify_finfun_const y)))])"
   | "random_finfun_aux (Suc_code_numeral i) j = Quickcheck.collapse (Random.select_weight
--- a/src/HOL/Library/Nested_Environment.thy	Wed Sep 23 11:34:21 2009 +0200
+++ b/src/HOL/Library/Nested_Environment.thy	Wed Sep 23 14:00:43 2009 +0200
@@ -567,6 +567,6 @@
 qed simp_all
 
 lemma [code, code del]:
-  "(Code_Eval.term_of :: ('a::{term_of, type}, 'b::{term_of, type}, 'c::{term_of, type}) env \<Rightarrow> term) = Code_Eval.term_of" ..
+  "(Code_Evaluation.term_of :: ('a::{term_of, type}, 'b::{term_of, type}, 'c::{term_of, type}) env \<Rightarrow> term) = Code_Evaluation.term_of" ..
 
 end
--- a/src/HOL/Lim.thy	Wed Sep 23 11:34:21 2009 +0200
+++ b/src/HOL/Lim.thy	Wed Sep 23 14:00:43 2009 +0200
@@ -84,6 +84,8 @@
 lemma LIM_const [simp]: "(%x. k) -- x --> k"
 by (simp add: LIM_def)
 
+lemma LIM_cong_limit: "\<lbrakk> f -- x --> L ; K = L \<rbrakk> \<Longrightarrow> f -- x --> K" by simp
+
 lemma LIM_add:
   fixes f g :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
   assumes f: "f -- a --> L" and g: "g -- a --> M"
--- a/src/HOL/Quickcheck.thy	Wed Sep 23 11:34:21 2009 +0200
+++ b/src/HOL/Quickcheck.thy	Wed Sep 23 14:00:43 2009 +0200
@@ -3,7 +3,7 @@
 header {* A simple counterexample generator *}
 
 theory Quickcheck
-imports Random Code_Eval
+imports Random Code_Evaluation
 uses ("Tools/quickcheck_generators.ML")
 begin
 
@@ -24,7 +24,7 @@
 
 definition
   "random i = Random.range 2 o\<rightarrow>
-    (\<lambda>k. Pair (if k = 0 then Code_Eval.valtermify False else Code_Eval.valtermify True))"
+    (\<lambda>k. Pair (if k = 0 then Code_Evaluation.valtermify False else Code_Evaluation.valtermify True))"
 
 instance ..
 
@@ -34,7 +34,7 @@
 begin
 
 definition random_itself :: "code_numeral \<Rightarrow> Random.seed \<Rightarrow> ('a itself \<times> (unit \<Rightarrow> term)) \<times> Random.seed" where
-  "random_itself _ = Pair (Code_Eval.valtermify TYPE('a))"
+  "random_itself _ = Pair (Code_Evaluation.valtermify TYPE('a))"
 
 instance ..
 
@@ -44,7 +44,7 @@
 begin
 
 definition
-  "random _ = Random.select chars o\<rightarrow> (\<lambda>c. Pair (c, \<lambda>u. Code_Eval.term_of c))"
+  "random _ = Random.select chars o\<rightarrow> (\<lambda>c. Pair (c, \<lambda>u. Code_Evaluation.term_of c))"
 
 instance ..
 
@@ -54,7 +54,7 @@
 begin
 
 definition 
-  "random _ = Pair (STR '''', \<lambda>u. Code_Eval.term_of (STR ''''))"
+  "random _ = Pair (STR '''', \<lambda>u. Code_Evaluation.term_of (STR ''''))"
 
 instance ..
 
@@ -63,10 +63,10 @@
 instantiation nat :: random
 begin
 
-definition random_nat :: "code_numeral \<Rightarrow> Random.seed \<Rightarrow> (nat \<times> (unit \<Rightarrow> Code_Eval.term)) \<times> Random.seed" where
+definition random_nat :: "code_numeral \<Rightarrow> Random.seed \<Rightarrow> (nat \<times> (unit \<Rightarrow> Code_Evaluation.term)) \<times> Random.seed" where
   "random_nat i = Random.range (i + 1) o\<rightarrow> (\<lambda>k. Pair (
      let n = Code_Numeral.nat_of k
-     in (n, \<lambda>_. Code_Eval.term_of n)))"
+     in (n, \<lambda>_. Code_Evaluation.term_of n)))"
 
 instance ..
 
@@ -78,7 +78,7 @@
 definition
   "random i = Random.range (2 * i + 1) o\<rightarrow> (\<lambda>k. Pair (
      let j = (if k \<ge> i then Code_Numeral.int_of (k - i) else - Code_Numeral.int_of (i - k))
-     in (j, \<lambda>_. Code_Eval.term_of j)))"
+     in (j, \<lambda>_. Code_Evaluation.term_of j)))"
 
 instance ..
 
@@ -95,7 +95,7 @@
 
 definition random_fun_lift :: "(Random.seed \<Rightarrow> ('b \<times> (unit \<Rightarrow> term)) \<times> Random.seed)
   \<Rightarrow> Random.seed \<Rightarrow> (('a\<Colon>term_of \<Rightarrow> 'b\<Colon>typerep) \<times> (unit \<Rightarrow> term)) \<times> Random.seed" where
-  "random_fun_lift f = random_fun_aux TYPEREP('a) TYPEREP('b) (op =) Code_Eval.term_of f Random.split_seed"
+  "random_fun_lift f = random_fun_aux TYPEREP('a) TYPEREP('b) (op =) Code_Evaluation.term_of f Random.split_seed"
 
 instantiation "fun" :: ("{eq, term_of}", random) random
 begin
--- a/src/HOL/Rational.thy	Wed Sep 23 11:34:21 2009 +0200
+++ b/src/HOL/Rational.thy	Wed Sep 23 14:00:43 2009 +0200
@@ -1002,8 +1002,8 @@
   by simp
 
 definition (in term_syntax)
-  valterm_fract :: "int \<times> (unit \<Rightarrow> Code_Eval.term) \<Rightarrow> int \<times> (unit \<Rightarrow> Code_Eval.term) \<Rightarrow> rat \<times> (unit \<Rightarrow> Code_Eval.term)" where
-  [code_unfold]: "valterm_fract k l = Code_Eval.valtermify Fract {\<cdot>} k {\<cdot>} l"
+  valterm_fract :: "int \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> int \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> rat \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
+  [code_unfold]: "valterm_fract k l = Code_Evaluation.valtermify Fract {\<cdot>} k {\<cdot>} l"
 
 notation fcomp (infixl "o>" 60)
 notation scomp (infixl "o\<rightarrow>" 60)
@@ -1014,7 +1014,7 @@
 definition
   "Quickcheck.random i = Quickcheck.random i o\<rightarrow> (\<lambda>num. Random.range i o\<rightarrow> (\<lambda>denom. Pair (
      let j = Code_Numeral.int_of (denom + 1)
-     in valterm_fract num (j, \<lambda>u. Code_Eval.term_of j))))"
+     in valterm_fract num (j, \<lambda>u. Code_Evaluation.term_of j))))"
 
 instance ..
 
--- a/src/HOL/RealDef.thy	Wed Sep 23 11:34:21 2009 +0200
+++ b/src/HOL/RealDef.thy	Wed Sep 23 14:00:43 2009 +0200
@@ -1128,8 +1128,8 @@
   by (simp add: of_rat_divide)
 
 definition (in term_syntax)
-  valterm_ratreal :: "rat \<times> (unit \<Rightarrow> Code_Eval.term) \<Rightarrow> real \<times> (unit \<Rightarrow> Code_Eval.term)" where
-  [code_unfold]: "valterm_ratreal k = Code_Eval.valtermify Ratreal {\<cdot>} k"
+  valterm_ratreal :: "rat \<times> (unit \<Rightarrow> Code_Evaluation.term) \<Rightarrow> real \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
+  [code_unfold]: "valterm_ratreal k = Code_Evaluation.valtermify Ratreal {\<cdot>} k"
 
 notation fcomp (infixl "o>" 60)
 notation scomp (infixl "o\<rightarrow>" 60)
--- a/src/HOL/Tools/hologic.ML	Wed Sep 23 11:34:21 2009 +0200
+++ b/src/HOL/Tools/hologic.ML	Wed Sep 23 14:00:43 2009 +0200
@@ -613,17 +613,17 @@
   | mk_typerep (T as TFree _) = Const ("Typerep.typerep_class.typerep",
       Term.itselfT T --> typerepT) $ Logic.mk_type T;
 
-val termT = Type ("Code_Eval.term", []);
+val termT = Type ("Code_Evaluation.term", []);
 
-fun term_of_const T = Const ("Code_Eval.term_of_class.term_of", T --> termT);
+fun term_of_const T = Const ("Code_Evaluation.term_of_class.term_of", T --> termT);
 
 fun mk_term_of T t = term_of_const T $ t;
 
 fun reflect_term (Const (c, T)) =
-      Const ("Code_Eval.Const", literalT --> typerepT --> termT)
+      Const ("Code_Evaluation.Const", literalT --> typerepT --> termT)
         $ mk_literal c $ mk_typerep T
   | reflect_term (t1 $ t2) =
-      Const ("Code_Eval.App", termT --> termT --> termT)
+      Const ("Code_Evaluation.App", termT --> termT --> termT)
         $ reflect_term t1 $ reflect_term t2
   | reflect_term (Abs (v, _, t)) = Abs (v, termT, reflect_term t)
   | reflect_term t = t;
@@ -631,7 +631,7 @@
 fun mk_valtermify_app c vs T =
   let
     fun termifyT T = mk_prodT (T, unitT --> termT);
-    fun valapp T T' = Const ("Code_Eval.valapp",
+    fun valapp T T' = Const ("Code_Evaluation.valapp",
       termifyT (T --> T') --> termifyT T --> termifyT T');
     fun mk_fTs [] _ = []
       | mk_fTs (_ :: Ts) T = (Ts ---> T) :: mk_fTs Ts T;
--- a/src/HOL/Tools/inductive.ML	Wed Sep 23 11:34:21 2009 +0200
+++ b/src/HOL/Tools/inductive.ML	Wed Sep 23 14:00:43 2009 +0200
@@ -103,7 +103,10 @@
       "(P & True) = P" "(True & P) = P"
     by (fact simp_thms)+};
 
-val simp_thms'' = inf_fun_eq :: inf_bool_eq :: simp_thms';
+val simp_thms'' = map mk_meta_eq [@{thm inf_fun_eq}, @{thm inf_bool_eq}] @ simp_thms';
+
+val simp_thms''' = map mk_meta_eq
+  [@{thm le_fun_def}, @{thm le_bool_def}, @{thm sup_fun_eq}, @{thm sup_bool_eq}];
 
 
 (** context data **)
@@ -171,15 +174,15 @@
       (case concl of
           (_ $ (_ $ (Const ("Not", _) $ _) $ _)) => []
         | _ => [thm' RS (thm' RS eq_to_mono2)]);
-    fun dest_less_concl thm = dest_less_concl (thm RS le_funD)
-      handle THM _ => thm RS le_boolD
+    fun dest_less_concl thm = dest_less_concl (thm RS @{thm le_funD})
+      handle THM _ => thm RS @{thm le_boolD}
   in
     case concl of
       Const ("==", _) $ _ $ _ => eq2mono (thm RS meta_eq_to_obj_eq)
     | _ $ (Const ("op =", _) $ _ $ _) => eq2mono thm
     | _ $ (Const (@{const_name HOL.less_eq}, _) $ _ $ _) =>
       [dest_less_concl (Seq.hd (REPEAT (FIRSTGOAL
-         (resolve_tac [le_funI, le_boolI'])) thm))]
+         (resolve_tac [@{thm le_funI}, @{thm le_boolI'}])) thm))]
     | _ => [thm]
   end handle THM _ =>
     error ("Bad monotonicity theorem:\n" ^ Display.string_of_thm_without_context thm);
@@ -323,11 +326,11 @@
     (HOLogic.mk_Trueprop
       (Const (@{const_name Orderings.mono}, (predT --> predT) --> HOLogic.boolT) $ fp_fun))
     (fn _ => EVERY [rtac @{thm monoI} 1,
-      REPEAT (resolve_tac [le_funI, le_boolI'] 1),
+      REPEAT (resolve_tac [@{thm le_funI}, @{thm le_boolI'}] 1),
       REPEAT (FIRST
         [atac 1,
          resolve_tac (List.concat (map mk_mono monos) @ get_monos ctxt) 1,
-         etac le_funE 1, dtac le_boolD 1])]));
+         etac @{thm le_funE} 1, dtac @{thm le_boolD} 1])]));
 
 
 (* prove introduction rules *)
@@ -338,7 +341,7 @@
 
     val unfold = funpow k (fn th => th RS fun_cong)
       (mono RS (fp_def RS
-        (if coind then def_gfp_unfold else def_lfp_unfold)));
+        (if coind then @{thm def_gfp_unfold} else @{thm def_lfp_unfold})));
 
     fun select_disj 1 1 = []
       | select_disj _ 1 = [rtac disjI1]
@@ -553,13 +556,13 @@
     val ind_concl = HOLogic.mk_Trueprop
       (HOLogic.mk_binrel "HOL.ord_class.less_eq" (rec_const, ind_pred));
 
-    val raw_fp_induct = (mono RS (fp_def RS def_lfp_induct));
+    val raw_fp_induct = (mono RS (fp_def RS @{thm def_lfp_induct}));
 
     val induct = SkipProof.prove ctxt'' [] ind_prems ind_concl
       (fn {prems, ...} => EVERY
         [rewrite_goals_tac [inductive_conj_def],
          DETERM (rtac raw_fp_induct 1),
-         REPEAT (resolve_tac [le_funI, le_boolI] 1),
+         REPEAT (resolve_tac [@{thm le_funI}, @{thm le_boolI}] 1),
          rewrite_goals_tac simp_thms'',
          (*This disjE separates out the introduction rules*)
          REPEAT (FIRSTGOAL (eresolve_tac [disjE, exE, FalseE])),
@@ -577,7 +580,7 @@
         [rewrite_goals_tac rec_preds_defs,
          REPEAT (EVERY
            [REPEAT (resolve_tac [conjI, impI] 1),
-            REPEAT (eresolve_tac [le_funE, le_boolE] 1),
+            REPEAT (eresolve_tac [@{thm le_funE}, @{thm le_boolE}] 1),
             atac 1,
             rewrite_goals_tac simp_thms',
             atac 1])])
@@ -763,8 +766,8 @@
            (snd (Variable.add_fixes (map (fst o dest_Free) params) ctxt1)) ctxt1)
            (rotate_prems ~1 (ObjectLogic.rulify
              (fold_rule rec_preds_defs
-               (rewrite_rule [le_fun_def, le_bool_def, sup_fun_eq, sup_bool_eq]
-                (mono RS (fp_def RS def_coinduct))))))
+               (rewrite_rule simp_thms'''
+                (mono RS (fp_def RS @{thm def_coinduct}))))))
        else
          prove_indrule quiet_mode cs argTs bs xs rec_const params intr_ts mono fp_def
            rec_preds_defs ctxt1);
--- a/src/HOL/Tools/quickcheck_generators.ML	Wed Sep 23 11:34:21 2009 +0200
+++ b/src/HOL/Tools/quickcheck_generators.ML	Wed Sep 23 14:00:43 2009 +0200
@@ -76,7 +76,7 @@
     val lhs = HOLogic.mk_random T size;
     val rhs = HOLogic.mk_ST [((HOLogic.mk_random T' size, @{typ Random.seed}), SOME (v, Tm'))]
       (HOLogic.mk_return Tm @{typ Random.seed}
-      (mk_const "Code_Eval.valapp" [T', T]
+      (mk_const "Code_Evaluation.valapp" [T', T]
         $ HOLogic.mk_prod (t_constr, Abs ("u", @{typ unit}, HOLogic.reflect_term t_constr)) $ t_v))
       @{typ Random.seed} (SOME Tm, @{typ Random.seed});
     val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
--- a/src/HOL/ex/predicate_compile.ML	Wed Sep 23 11:34:21 2009 +0200
+++ b/src/HOL/ex/predicate_compile.ML	Wed Sep 23 14:00:43 2009 +0200
@@ -169,7 +169,7 @@
   end;
 
 fun dest_randomT (Type ("fun", [@{typ Random.seed},
-  Type ("*", [Type ("*", [T, @{typ "unit => Code_Eval.term"}]) ,@{typ Random.seed}])])) = T
+  Type ("*", [Type ("*", [T, @{typ "unit => Code_Evaluation.term"}]) ,@{typ Random.seed}])])) = T
   | dest_randomT T = raise TYPE ("dest_randomT", [T], [])
 
 (* destruction of intro rules *)
@@ -707,7 +707,7 @@
 end;
 
 (* termify_code:
-val termT = Type ("Code_Eval.term", []);
+val termT = Type ("Code_Evaluation.term", []);
 fun termifyT T = HOLogic.mk_prodT (T, HOLogic.unitT --> termT)
 *)
 (*
@@ -1198,7 +1198,7 @@
       val t1' = mk_valtermify_term t1
       val t2' = mk_valtermify_term t2
     in
-      Const ("Code_Eval.valapp", termifyT T --> termifyT T1 --> termifyT T2) $ t1' $ t2'
+      Const ("Code_Evaluation.valapp", termifyT T --> termifyT T1 --> termifyT T2) $ t1' $ t2'
     end
   | mk_valtermify_term _ = error "Not a valid term for mk_valtermify_term"
 *)
--- a/src/Pure/envir.ML	Wed Sep 23 11:34:21 2009 +0200
+++ b/src/Pure/envir.ML	Wed Sep 23 14:00:43 2009 +0200
@@ -282,12 +282,9 @@
   let
     val funerr = "fastype: expected function type";
     fun fast Ts (f $ u) =
-          (case fast Ts f of
+          (case Type.devar tyenv (fast Ts f) of
             Type ("fun", [_, T]) => T
-          | TVar v =>
-              (case Type.lookup tyenv v of
-                SOME (Type ("fun", [_, T])) => T
-              | _ => raise TERM (funerr, [f $ u]))
+          | TVar v => raise TERM (funerr, [f $ u])
           | _ => raise TERM (funerr, [f $ u]))
       | fast Ts (Const (_, T)) = T
       | fast Ts (Free (_, T)) = T
--- a/src/Pure/type.ML	Wed Sep 23 11:34:21 2009 +0200
+++ b/src/Pure/type.ML	Wed Sep 23 14:00:43 2009 +0200
@@ -55,6 +55,7 @@
   exception TYPE_MATCH
   type tyenv = (sort * typ) Vartab.table
   val lookup: tyenv -> indexname * sort -> typ option
+  val devar: tyenv -> typ -> typ
   val typ_match: tsig -> typ * typ -> tyenv -> tyenv
   val typ_instance: tsig -> typ * typ -> bool
   val raw_match: typ * typ -> tyenv -> tyenv