merged
authornipkow
Fri, 20 Feb 2009 20:51:06 +0100
changeset 30028 f2421131a83c
parent 30026 be13af70c27c (diff)
parent 30027 ab40c5e007e0 (current diff)
child 30030 00d04a562df1
merged
--- a/src/HOL/IsaMakefile	Fri Feb 20 20:50:49 2009 +0100
+++ b/src/HOL/IsaMakefile	Fri Feb 20 20:51:06 2009 +0100
@@ -339,6 +339,8 @@
   Library/Random.thy	Library/Quickcheck.thy	\
   Library/Poly_Deriv.thy \
   Library/Polynomial.thy \
+  Library/Product_plus.thy \
+  Library/Product_Vector.thy \
   Library/Enum.thy Library/Float.thy $(SRC)/Tools/float.ML $(SRC)/HOL/Tools/float_arith.ML \
   Library/reify_data.ML Library/reflection.ML
 	@cd Library; $(ISABELLE_TOOL) usedir $(OUT)/HOL Library
--- a/src/HOL/Library/Library.thy	Fri Feb 20 20:50:49 2009 +0100
+++ b/src/HOL/Library/Library.thy	Fri Feb 20 20:51:06 2009 +0100
@@ -41,6 +41,7 @@
   Poly_Deriv
   Polynomial
   Primes
+  Product_Vector
   Quickcheck
   Quicksort
   Quotient
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Product_Vector.thy	Fri Feb 20 20:51:06 2009 +0100
@@ -0,0 +1,273 @@
+(*  Title:      HOL/Library/Product_Vector.thy
+    Author:     Brian Huffman
+*)
+
+header {* Cartesian Products as Vector Spaces *}
+
+theory Product_Vector
+imports Inner_Product Product_plus
+begin
+
+subsection {* Product is a real vector space *}
+
+instantiation "*" :: (real_vector, real_vector) real_vector
+begin
+
+definition scaleR_prod_def:
+  "scaleR r A = (scaleR r (fst A), scaleR r (snd A))"
+
+lemma fst_scaleR [simp]: "fst (scaleR r A) = scaleR r (fst A)"
+  unfolding scaleR_prod_def by simp
+
+lemma snd_scaleR [simp]: "snd (scaleR r A) = scaleR r (snd A)"
+  unfolding scaleR_prod_def by simp
+
+lemma scaleR_Pair [simp]: "scaleR r (a, b) = (scaleR r a, scaleR r b)"
+  unfolding scaleR_prod_def by simp
+
+instance proof
+  fix a b :: real and x y :: "'a \<times> 'b"
+  show "scaleR a (x + y) = scaleR a x + scaleR a y"
+    by (simp add: expand_prod_eq scaleR_right_distrib)
+  show "scaleR (a + b) x = scaleR a x + scaleR b x"
+    by (simp add: expand_prod_eq scaleR_left_distrib)
+  show "scaleR a (scaleR b x) = scaleR (a * b) x"
+    by (simp add: expand_prod_eq)
+  show "scaleR 1 x = x"
+    by (simp add: expand_prod_eq)
+qed
+
+end
+
+subsection {* Product is a normed vector space *}
+
+instantiation
+  "*" :: (real_normed_vector, real_normed_vector) real_normed_vector
+begin
+
+definition norm_prod_def:
+  "norm x = sqrt ((norm (fst x))\<twosuperior> + (norm (snd x))\<twosuperior>)"
+
+definition sgn_prod_def:
+  "sgn (x::'a \<times> 'b) = scaleR (inverse (norm x)) x"
+
+lemma norm_Pair: "norm (a, b) = sqrt ((norm a)\<twosuperior> + (norm b)\<twosuperior>)"
+  unfolding norm_prod_def by simp
+
+instance proof
+  fix r :: real and x y :: "'a \<times> 'b"
+  show "0 \<le> norm x"
+    unfolding norm_prod_def by simp
+  show "norm x = 0 \<longleftrightarrow> x = 0"
+    unfolding norm_prod_def
+    by (simp add: expand_prod_eq)
+  show "norm (x + y) \<le> norm x + norm y"
+    unfolding norm_prod_def
+    apply (rule order_trans [OF _ real_sqrt_sum_squares_triangle_ineq])
+    apply (simp add: add_mono power_mono norm_triangle_ineq)
+    done
+  show "norm (scaleR r x) = \<bar>r\<bar> * norm x"
+    unfolding norm_prod_def
+    apply (simp add: norm_scaleR power_mult_distrib)
+    apply (simp add: right_distrib [symmetric])
+    apply (simp add: real_sqrt_mult_distrib)
+    done
+  show "sgn x = scaleR (inverse (norm x)) x"
+    by (rule sgn_prod_def)
+qed
+
+end
+
+subsection {* Product is an inner product space *}
+
+instantiation "*" :: (real_inner, real_inner) real_inner
+begin
+
+definition inner_prod_def:
+  "inner x y = inner (fst x) (fst y) + inner (snd x) (snd y)"
+
+lemma inner_Pair [simp]: "inner (a, b) (c, d) = inner a c + inner b d"
+  unfolding inner_prod_def by simp
+
+instance proof
+  fix r :: real
+  fix x y z :: "'a::real_inner * 'b::real_inner"
+  show "inner x y = inner y x"
+    unfolding inner_prod_def
+    by (simp add: inner_commute)
+  show "inner (x + y) z = inner x z + inner y z"
+    unfolding inner_prod_def
+    by (simp add: inner_left_distrib)
+  show "inner (scaleR r x) y = r * inner x y"
+    unfolding inner_prod_def
+    by (simp add: inner_scaleR_left right_distrib)
+  show "0 \<le> inner x x"
+    unfolding inner_prod_def
+    by (intro add_nonneg_nonneg inner_ge_zero)
+  show "inner x x = 0 \<longleftrightarrow> x = 0"
+    unfolding inner_prod_def expand_prod_eq
+    by (simp add: add_nonneg_eq_0_iff)
+  show "norm x = sqrt (inner x x)"
+    unfolding norm_prod_def inner_prod_def
+    by (simp add: power2_norm_eq_inner)
+qed
+
+end
+
+subsection {* Pair operations are linear and continuous *}
+
+interpretation fst!: bounded_linear fst
+  apply (unfold_locales)
+  apply (rule fst_add)
+  apply (rule fst_scaleR)
+  apply (rule_tac x="1" in exI, simp add: norm_Pair)
+  done
+
+interpretation snd!: bounded_linear snd
+  apply (unfold_locales)
+  apply (rule snd_add)
+  apply (rule snd_scaleR)
+  apply (rule_tac x="1" in exI, simp add: norm_Pair)
+  done
+
+text {* TODO: move to NthRoot *}
+lemma sqrt_add_le_add_sqrt:
+  assumes x: "0 \<le> x" and y: "0 \<le> y"
+  shows "sqrt (x + y) \<le> sqrt x + sqrt y"
+apply (rule power2_le_imp_le)
+apply (simp add: real_sum_squared_expand add_nonneg_nonneg x y)
+apply (simp add: mult_nonneg_nonneg x y)
+apply (simp add: add_nonneg_nonneg x y)
+done
+
+lemma bounded_linear_Pair:
+  assumes f: "bounded_linear f"
+  assumes g: "bounded_linear g"
+  shows "bounded_linear (\<lambda>x. (f x, g x))"
+proof
+  interpret f: bounded_linear f by fact
+  interpret g: bounded_linear g by fact
+  fix x y and r :: real
+  show "(f (x + y), g (x + y)) = (f x, g x) + (f y, g y)"
+    by (simp add: f.add g.add)
+  show "(f (r *\<^sub>R x), g (r *\<^sub>R x)) = r *\<^sub>R (f x, g x)"
+    by (simp add: f.scaleR g.scaleR)
+  obtain Kf where "0 < Kf" and norm_f: "\<And>x. norm (f x) \<le> norm x * Kf"
+    using f.pos_bounded by fast
+  obtain Kg where "0 < Kg" and norm_g: "\<And>x. norm (g x) \<le> norm x * Kg"
+    using g.pos_bounded by fast
+  have "\<forall>x. norm (f x, g x) \<le> norm x * (Kf + Kg)"
+    apply (rule allI)
+    apply (simp add: norm_Pair)
+    apply (rule order_trans [OF sqrt_add_le_add_sqrt], simp, simp)
+    apply (simp add: right_distrib)
+    apply (rule add_mono [OF norm_f norm_g])
+    done
+  then show "\<exists>K. \<forall>x. norm (f x, g x) \<le> norm x * K" ..
+qed
+
+text {*
+  TODO: The next three proofs are nearly identical to each other.
+  Is there a good way to factor out the common parts?
+*}
+
+lemma LIMSEQ_Pair:
+  assumes "X ----> a" and "Y ----> b"
+  shows "(\<lambda>n. (X n, Y n)) ----> (a, b)"
+proof (rule LIMSEQ_I)
+  fix r :: real assume "0 < r"
+  then have "0 < r / sqrt 2" (is "0 < ?s")
+    by (simp add: divide_pos_pos)
+  obtain M where M: "\<forall>n\<ge>M. norm (X n - a) < ?s"
+    using LIMSEQ_D [OF `X ----> a` `0 < ?s`] ..
+  obtain N where N: "\<forall>n\<ge>N. norm (Y n - b) < ?s"
+    using LIMSEQ_D [OF `Y ----> b` `0 < ?s`] ..
+  have "\<forall>n\<ge>max M N. norm ((X n, Y n) - (a, b)) < r"
+    using M N by (simp add: real_sqrt_sum_squares_less norm_Pair)
+  then show "\<exists>n0. \<forall>n\<ge>n0. norm ((X n, Y n) - (a, b)) < r" ..
+qed
+
+lemma Cauchy_Pair:
+  assumes "Cauchy X" and "Cauchy Y"
+  shows "Cauchy (\<lambda>n. (X n, Y n))"
+proof (rule CauchyI)
+  fix r :: real assume "0 < r"
+  then have "0 < r / sqrt 2" (is "0 < ?s")
+    by (simp add: divide_pos_pos)
+  obtain M where M: "\<forall>m\<ge>M. \<forall>n\<ge>M. norm (X m - X n) < ?s"
+    using CauchyD [OF `Cauchy X` `0 < ?s`] ..
+  obtain N where N: "\<forall>m\<ge>N. \<forall>n\<ge>N. norm (Y m - Y n) < ?s"
+    using CauchyD [OF `Cauchy Y` `0 < ?s`] ..
+  have "\<forall>m\<ge>max M N. \<forall>n\<ge>max M N. norm ((X m, Y m) - (X n, Y n)) < r"
+    using M N by (simp add: real_sqrt_sum_squares_less norm_Pair)
+  then show "\<exists>n0. \<forall>m\<ge>n0. \<forall>n\<ge>n0. norm ((X m, Y m) - (X n, Y n)) < r" ..
+qed
+
+lemma LIM_Pair:
+  assumes "f -- x --> a" and "g -- x --> b"
+  shows "(\<lambda>x. (f x, g x)) -- x --> (a, b)"
+proof (rule LIM_I)
+  fix r :: real assume "0 < r"
+  then have "0 < r / sqrt 2" (is "0 < ?e")
+    by (simp add: divide_pos_pos)
+  obtain s where s: "0 < s"
+    "\<forall>z. z \<noteq> x \<and> norm (z - x) < s \<longrightarrow> norm (f z - a) < ?e"
+    using LIM_D [OF `f -- x --> a` `0 < ?e`] by fast
+  obtain t where t: "0 < t"
+    "\<forall>z. z \<noteq> x \<and> norm (z - x) < t \<longrightarrow> norm (g z - b) < ?e"
+    using LIM_D [OF `g -- x --> b` `0 < ?e`] by fast
+  have "0 < min s t \<and>
+    (\<forall>z. z \<noteq> x \<and> norm (z - x) < min s t \<longrightarrow> norm ((f z, g z) - (a, b)) < r)"
+    using s t by (simp add: real_sqrt_sum_squares_less norm_Pair)
+  then show
+    "\<exists>s>0. \<forall>z. z \<noteq> x \<and> norm (z - x) < s \<longrightarrow> norm ((f z, g z) - (a, b)) < r" ..
+qed
+
+lemma isCont_Pair [simp]:
+  "\<lbrakk>isCont f x; isCont g x\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. (f x, g x)) x"
+  unfolding isCont_def by (rule LIM_Pair)
+
+
+subsection {* Product is a complete vector space *}
+
+instance "*" :: (banach, banach) banach
+proof
+  fix X :: "nat \<Rightarrow> 'a \<times> 'b" assume "Cauchy X"
+  have 1: "(\<lambda>n. fst (X n)) ----> lim (\<lambda>n. fst (X n))"
+    using fst.Cauchy [OF `Cauchy X`]
+    by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff)
+  have 2: "(\<lambda>n. snd (X n)) ----> lim (\<lambda>n. snd (X n))"
+    using snd.Cauchy [OF `Cauchy X`]
+    by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff)
+  have "X ----> (lim (\<lambda>n. fst (X n)), lim (\<lambda>n. snd (X n)))"
+    using LIMSEQ_Pair [OF 1 2] by simp
+  then show "convergent X"
+    by (rule convergentI)
+qed
+
+subsection {* Frechet derivatives involving pairs *}
+
+lemma FDERIV_Pair:
+  assumes f: "FDERIV f x :> f'" and g: "FDERIV g x :> g'"
+  shows "FDERIV (\<lambda>x. (f x, g x)) x :> (\<lambda>h. (f' h, g' h))"
+apply (rule FDERIV_I)
+apply (rule bounded_linear_Pair)
+apply (rule FDERIV_bounded_linear [OF f])
+apply (rule FDERIV_bounded_linear [OF g])
+apply (simp add: norm_Pair)
+apply (rule real_LIM_sandwich_zero)
+apply (rule LIM_add_zero)
+apply (rule FDERIV_D [OF f])
+apply (rule FDERIV_D [OF g])
+apply (rename_tac h)
+apply (simp add: divide_nonneg_pos)
+apply (rename_tac h)
+apply (subst add_divide_distrib [symmetric])
+apply (rule divide_right_mono [OF _ norm_ge_zero])
+apply (rule order_trans [OF sqrt_add_le_add_sqrt])
+apply simp
+apply simp
+apply simp
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Product_plus.thy	Fri Feb 20 20:51:06 2009 +0100
@@ -0,0 +1,115 @@
+(*  Title:      HOL/Library/Product_plus.thy
+    Author:     Brian Huffman
+*)
+
+header {* Additive group operations on product types *}
+
+theory Product_plus
+imports Main
+begin
+
+subsection {* Operations *}
+
+instantiation "*" :: (zero, zero) zero
+begin
+
+definition zero_prod_def: "0 = (0, 0)"
+
+instance ..
+end
+
+instantiation "*" :: (plus, plus) plus
+begin
+
+definition plus_prod_def:
+  "x + y = (fst x + fst y, snd x + snd y)"
+
+instance ..
+end
+
+instantiation "*" :: (minus, minus) minus
+begin
+
+definition minus_prod_def:
+  "x - y = (fst x - fst y, snd x - snd y)"
+
+instance ..
+end
+
+instantiation "*" :: (uminus, uminus) uminus
+begin
+
+definition uminus_prod_def:
+  "- x = (- fst x, - snd x)"
+
+instance ..
+end
+
+lemma fst_zero [simp]: "fst 0 = 0"
+  unfolding zero_prod_def by simp
+
+lemma snd_zero [simp]: "snd 0 = 0"
+  unfolding zero_prod_def by simp
+
+lemma fst_add [simp]: "fst (x + y) = fst x + fst y"
+  unfolding plus_prod_def by simp
+
+lemma snd_add [simp]: "snd (x + y) = snd x + snd y"
+  unfolding plus_prod_def by simp
+
+lemma fst_diff [simp]: "fst (x - y) = fst x - fst y"
+  unfolding minus_prod_def by simp
+
+lemma snd_diff [simp]: "snd (x - y) = snd x - snd y"
+  unfolding minus_prod_def by simp
+
+lemma fst_uminus [simp]: "fst (- x) = - fst x"
+  unfolding uminus_prod_def by simp
+
+lemma snd_uminus [simp]: "snd (- x) = - snd x"
+  unfolding uminus_prod_def by simp
+
+lemma add_Pair [simp]: "(a, b) + (c, d) = (a + c, b + d)"
+  unfolding plus_prod_def by simp
+
+lemma diff_Pair [simp]: "(a, b) - (c, d) = (a - c, b - d)"
+  unfolding minus_prod_def by simp
+
+lemma uminus_Pair [simp, code]: "- (a, b) = (- a, - b)"
+  unfolding uminus_prod_def by simp
+
+lemmas expand_prod_eq = Pair_fst_snd_eq
+
+
+subsection {* Class instances *}
+
+instance "*" :: (semigroup_add, semigroup_add) semigroup_add
+  by default (simp add: expand_prod_eq add_assoc)
+
+instance "*" :: (ab_semigroup_add, ab_semigroup_add) ab_semigroup_add
+  by default (simp add: expand_prod_eq add_commute)
+
+instance "*" :: (monoid_add, monoid_add) monoid_add
+  by default (simp_all add: expand_prod_eq)
+
+instance "*" :: (comm_monoid_add, comm_monoid_add) comm_monoid_add
+  by default (simp add: expand_prod_eq)
+
+instance "*" ::
+  (cancel_semigroup_add, cancel_semigroup_add) cancel_semigroup_add
+    by default (simp_all add: expand_prod_eq)
+
+instance "*" ::
+  (cancel_ab_semigroup_add, cancel_ab_semigroup_add) cancel_ab_semigroup_add
+    by default (simp add: expand_prod_eq)
+
+instance "*" ::
+  (cancel_comm_monoid_add, cancel_comm_monoid_add) cancel_comm_monoid_add ..
+
+instance "*" :: (group_add, group_add) group_add
+  by default (simp_all add: expand_prod_eq diff_minus)
+
+instance "*" :: (ab_group_add, ab_group_add) ab_group_add
+  by default (simp_all add: expand_prod_eq)
+
+end
--- a/src/HOL/Tools/datatype_package.ML	Fri Feb 20 20:50:49 2009 +0100
+++ b/src/HOL/Tools/datatype_package.ML	Fri Feb 20 20:51:06 2009 +0100
@@ -659,7 +659,7 @@
       | pretty_constr (co, [ty']) =
           (Pretty.block o Pretty.breaks)
             [Syntax.pretty_term ctxt (Const (co, ty' --> ty)),
-              Syntax.pretty_typ ctxt ty']
+              pretty_typ_br ty']
       | pretty_constr (co, tys) =
           (Pretty.block o Pretty.breaks)
             (Syntax.pretty_term ctxt (Const (co, tys ---> ty)) ::
--- a/src/HOL/ex/Eval_Examples.thy	Fri Feb 20 20:50:49 2009 +0100
+++ b/src/HOL/ex/Eval_Examples.thy	Fri Feb 20 20:51:06 2009 +0100
@@ -1,6 +1,4 @@
-(*  ID:         $Id$
-    Author:     Florian Haftmann, TU Muenchen
-*)
+(* Author: Florian Haftmann, TU Muenchen *)
 
 header {* Small examples for evaluation mechanisms *}
 
--- a/src/Pure/Isar/code.ML	Fri Feb 20 20:50:49 2009 +0100
+++ b/src/Pure/Isar/code.ML	Fri Feb 20 20:51:06 2009 +0100
@@ -35,7 +35,7 @@
   val these_raw_eqns: theory -> string -> (thm * bool) list
   val get_datatype: theory -> string -> ((string * sort) list * (string * typ list) list)
   val get_datatype_of_constr: theory -> string -> string option
-  val get_case_scheme: theory -> string -> (int * string list) option
+  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
 
@@ -111,7 +111,7 @@
 
 (** logical and syntactical specification of executable code **)
 
-(* defining equations *)
+(* code equations *)
 
 type eqns = bool * (thm * bool) list lazy;
   (*default flag, theorems with linear flag (perhaps lazy)*)
@@ -136,7 +136,7 @@
       Pattern.matchess thy (args, (map incr_idx o curry Library.take (length args)) args');
     fun drop (thm', linear') = if (linear orelse not linear')
       andalso matches_args (args_of thm') then 
-        (warning ("Code generator: dropping redundant defining equation\n" ^ Display.string_of_thm thm'); true)
+        (warning ("Code generator: dropping redundant code equation\n" ^ Display.string_of_thm thm'); true)
       else false;
   in (thm, linear) :: filter_out drop thms end;
 
@@ -409,7 +409,7 @@
   in
     (Pretty.writeln o Pretty.chunks) [
       Pretty.block (
-        Pretty.str "defining equations:"
+        Pretty.str "code equations:"
         :: Pretty.fbrk
         :: (Pretty.fbreaks o map pretty_eqn) eqns
       ),
@@ -452,7 +452,7 @@
         val ty1 :: tys = map (snd o Code_Unit.const_typ_eqn) thms';
         fun unify ty env = Sign.typ_unify thy (ty1, ty) env
           handle Type.TUNIFY =>
-            error ("Type unificaton failed, while unifying defining equations\n"
+            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));
@@ -463,7 +463,7 @@
 
 fun check_linear (eqn as (thm, linear)) =
   if linear then eqn else Code_Unit.bad_thm
-    ("Duplicate variables on left hand side of defining equation:\n"
+    ("Duplicate variables on left hand side of code equation:\n"
       ^ Display.string_of_thm thm);
 
 fun mk_eqn thy linear =
@@ -525,22 +525,13 @@
        then SOME tyco else NONE
     | _ => NONE;
 
-fun get_constr_typ thy c =
-  case get_datatype_of_constr thy c
-   of SOME tyco => let
-          val (vs, cos) = get_datatype thy tyco;
-          val SOME tys = AList.lookup (op =) cos c;
-          val ty = tys ---> Type (tyco, map TFree vs);
-        in SOME (Logic.varifyT ty) end
-    | NONE => NONE;
-
 fun recheck_eqn thy = Code_Unit.error_thm
   (Code_Unit.assert_linear (is_some o get_datatype_of_constr thy) o apfst (Code_Unit.assert_eqn thy));
 
 fun recheck_eqns_const thy c eqns =
   let
     fun cert (eqn as (thm, _)) = if c = Code_Unit.const_eqn thm
-      then eqn else error ("Wrong head of defining equation,\nexpected constant "
+      then eqn else error ("Wrong head of code equation,\nexpected constant "
         ^ Code_Unit.string_of_const thy c ^ "\n" ^ Display.string_of_thm thm)
   in map (cert o recheck_eqn thy) eqns end;
 
@@ -554,11 +545,11 @@
         let
           val c = Code_Unit.const_eqn thm;
           val _ = if not default andalso (is_some o AxClass.class_of_param thy) c
-            then error ("Rejected polymorphic equation for overloaded constant:\n"
+            then error ("Rejected polymorphic code equation for overloaded constant:\n"
               ^ Display.string_of_thm thm)
             else ();
           val _ = if not default andalso (is_some o get_datatype_of_constr thy) c
-            then error ("Rejected equation for datatype constructor:\n"
+            then error ("Rejected code equation for datatype constructor:\n"
               ^ Display.string_of_thm thm)
             else ();
         in change_eqns false c (add_thm thy default (thm, linear)) thy end
@@ -583,7 +574,12 @@
 
 fun del_eqns c = change_eqns true c (K (false, Lazy.value []));
 
-fun get_case_scheme thy = Symtab.lookup ((fst o the_cases o the_exec) thy);
+fun get_case_scheme thy c = case Symtab.lookup ((fst o the_cases o the_exec) thy) c
+ of SOME (base_case_scheme as (_, case_pats)) =>
+      if forall (is_some o get_datatype_of_constr thy) case_pats
+      then SOME (1 + Int.max (1, length case_pats), base_case_scheme)
+      else NONE
+  | NONE => NONE;
 
 val is_undefined = Symtab.defined o snd o the_cases o the_exec;
 
@@ -727,18 +723,16 @@
 
 fun default_typscheme thy c =
   let
-    val typscheme = curry (Code_Unit.typscheme thy) c
-    val the_const_type = snd o dest_Const o TermSubst.zero_var_indexes
-      o curry Const "" o Sign.the_const_type thy;
+    fun the_const_typscheme c = (curry (Code_Unit.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
-   of SOME class => the_const_type c
-        |> Term.map_type_tvar (K (TVar ((Name.aT, 0), [class])))
-        |> typscheme
-    | NONE => (case get_constr_typ thy c
-       of SOME ty => typscheme ty
-        | NONE => (case get_eqns thy c
-           of (thm, _) :: _ => snd (Code_Unit.head_eqn thy (Drule.zero_var_indexes thm))
-            | [] => typscheme (the_const_type c))) end;
+   of SOME class => ([(Name.aT, [class])], snd (the_const_typscheme c))
+    | NONE => if is_some (get_datatype_of_constr thy c)
+        then strip_sorts (the_const_typscheme c)
+        else case get_eqns thy c
+         of (thm, _) :: _ => snd (Code_Unit.head_eqn thy (Drule.zero_var_indexes thm))
+          | [] => strip_sorts (the_const_typscheme c) end;
 
 end; (*local*)
 
--- a/src/Pure/Isar/code_unit.ML	Fri Feb 20 20:50:49 2009 +0100
+++ b/src/Pure/Isar/code_unit.ML	Fri Feb 20 20:51:06 2009 +0100
@@ -34,7 +34,7 @@
   val constrset_of_consts: theory -> (string * typ) list
     -> string * ((string * sort) list * (string * typ list) list)
 
-  (*defining equations*)
+  (*code equations*)
   val assert_eqn: theory -> thm -> thm
   val mk_eqn: theory -> thm -> thm * bool
   val assert_linear: (string -> bool) -> thm * bool -> thm * bool
@@ -76,10 +76,11 @@
 
 fun typscheme thy (c, ty) =
   let
-    fun dest (TVar ((v, 0), sort)) = (v, sort)
+    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, ty) end;
+    val vs = map dest (Sign.const_typargs thy (c, ty'));
+  in (vs, Type.strip_sorts ty') end;
 
 fun inst_thm thy tvars' thm =
   let
@@ -313,10 +314,10 @@
     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, cs''')) end;
+  in (tyco, (vs, rev cs''')) end;
 
 
-(* defining equations *)
+(* code equations *)
 
 fun assert_eqn thy thm =
   let
@@ -351,7 +352,7 @@
             ^ Display.string_of_thm thm)
       | check 0 (Var _) = ()
       | check _ (Var _) = bad_thm
-          ("Variable with application on left hand side of defining equation\n"
+          ("Variable with application on left hand side of code equation\n"
             ^ Display.string_of_thm thm)
       | check n (t1 $ t2) = (check (n+1) t1; check 0 t2)
       | check n (Const (_, ty)) = if n <> (length o fst o strip_type) ty
@@ -363,7 +364,7 @@
     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 defining equation\n"
+           ^ "\nof code equation\n"
            ^ Display.string_of_thm thm
            ^ "\nis incompatible with declared function type\n"
            ^ string_of_typ thy ty_decl)
@@ -388,7 +389,7 @@
 fun assert_linear is_cons (thm, false) = (thm, false)
   | assert_linear is_cons (thm, true) = if snd (add_linear (assert_pat is_cons thm)) then (thm, true)
       else bad_thm
-        ("Duplicate variables on left hand side of defining equation:\n"
+        ("Duplicate variables on left hand side of code equation:\n"
           ^ Display.string_of_thm thm);
 
 
--- a/src/Tools/code/code_funcgr.ML	Fri Feb 20 20:50:49 2009 +0100
+++ b/src/Tools/code/code_funcgr.ML	Fri Feb 20 20:51:06 2009 +0100
@@ -317,13 +317,13 @@
 in
 
 val _ =
-  OuterSyntax.improper_command "code_thms" "print system of defining equations for code" OuterKeyword.diag
+  OuterSyntax.improper_command "code_thms" "print system of code equations for code" OuterKeyword.diag
     (Scan.repeat P.term_group
       >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory
         o Toplevel.keep ((fn thy => code_thms_cmd thy cs) o Toplevel.theory_of)));
 
 val _ =
-  OuterSyntax.improper_command "code_deps" "visualize dependencies of defining equations for code" OuterKeyword.diag
+  OuterSyntax.improper_command "code_deps" "visualize dependencies of code equations for code" OuterKeyword.diag
     (Scan.repeat P.term_group
       >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory
         o Toplevel.keep ((fn thy => code_deps_cmd thy cs) o Toplevel.theory_of)));
--- a/src/Tools/code/code_target.ML	Fri Feb 20 20:50:49 2009 +0100
+++ b/src/Tools/code/code_target.ML	Fri Feb 20 20:51:06 2009 +0100
@@ -418,7 +418,7 @@
     val program4 = Graph.subgraph (member (op =) names_all) program3;
     val empty_funs = filter_out (member (op =) abortable)
       (Code_Thingol.empty_funs program3);
-    val _ = if null empty_funs then () else error ("No defining equations for "
+    val _ = if null empty_funs then () else error ("No code equations for "
       ^ commas (map (Sign.extern_const thy) empty_funs));
   in
     serializer module args (labelled_name thy program2) reserved includes
--- a/src/Tools/code/code_thingol.ML	Fri Feb 20 20:50:49 2009 +0100
+++ b/src/Tools/code/code_thingol.ML	Fri Feb 20 20:51:06 2009 +0100
@@ -453,7 +453,7 @@
   let
     val err_class = Sorts.class_error (Syntax.pp_global thy) e;
     val err_thm = case thm
-     of SOME thm => "\n(in defining equation " ^ Display.string_of_thm thm ^ ")" | NONE => "";
+     of SOME thm => "\n(in code equation " ^ Display.string_of_thm thm ^ ")" | NONE => "";
     val err_typ = "Type " ^ Syntax.string_of_typ_global thy ty ^ " not of sort "
       ^ Syntax.string_of_sort_global thy sort;
   in error ("Wellsortedness error" ^ err_thm ^ ":\n" ^ err_typ ^ "\n" ^ err_class) end;
@@ -582,9 +582,8 @@
     fun stmt_classparam class =
       ensure_class thy algbr funcgr class
       #>> (fn class => Classparam (c, class));
-    fun stmt_fun ((vs, raw_ty), raw_thms) =
+    fun stmt_fun ((vs, ty), raw_thms) =
       let
-        val ty = Logic.unvarifyT raw_ty; (*FIXME change convention here*)
         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;
@@ -643,11 +642,6 @@
     val ts_clause = nth_drop t_pos ts;
     fun mk_clause (co, num_co_args) t =
       let
-        val _ = if (is_some o Code.get_datatype_of_constr thy) co then ()
-          else error ("Non-constructor " ^ quote co
-            ^ " encountered in case pattern"
-            ^ (case thm of NONE => ""
-              | SOME thm => ", in equation\n" ^ Display.string_of_thm thm))
         val (vs, body) = Term.strip_abs_eta num_co_args t;
         val not_undefined = case body
          of (Const (c, _)) => not (Code.is_undefined thy c)
@@ -702,9 +696,7 @@
     translate_case thy algbr funcgr thm case_scheme ((c, ty), ts)
 and translate_app thy algbr funcgr thm (c_ty_ts as ((c, _), _)) =
   case Code.get_case_scheme thy c
-   of SOME (base_case_scheme as (_, case_pats)) =>
-        translate_app_case thy algbr funcgr thm
-          (1 + Int.max (1, length case_pats), base_case_scheme) c_ty_ts
+   of SOME case_scheme => translate_app_case thy algbr funcgr thm case_scheme c_ty_ts
     | NONE => translate_app_const thy algbr funcgr thm c_ty_ts;
 
 
--- a/src/Tools/code/code_wellsorted.ML	Fri Feb 20 20:50:49 2009 +0100
+++ b/src/Tools/code/code_wellsorted.ML	Fri Feb 20 20:51:06 2009 +0100
@@ -388,13 +388,13 @@
 in
 
 val _ =
-  OuterSyntax.improper_command "code_thms" "print system of defining equations for code" OuterKeyword.diag
+  OuterSyntax.improper_command "code_thms" "print system of code equations for code" OuterKeyword.diag
     (Scan.repeat P.term_group
       >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory
         o Toplevel.keep ((fn thy => code_thms_cmd thy cs) o Toplevel.theory_of)));
 
 val _ =
-  OuterSyntax.improper_command "code_deps" "visualize dependencies of defining equations for code" OuterKeyword.diag
+  OuterSyntax.improper_command "code_deps" "visualize dependencies of code equations for code" OuterKeyword.diag
     (Scan.repeat P.term_group
       >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory
         o Toplevel.keep ((fn thy => code_deps_cmd thy cs) o Toplevel.theory_of)));
--- a/src/Tools/nbe.ML	Fri Feb 20 20:50:49 2009 +0100
+++ b/src/Tools/nbe.ML	Fri Feb 20 20:51:06 2009 +0100
@@ -389,8 +389,8 @@
             val ts' = take_until is_dict ts;
             val c = const_of_idx idx;
             val (_, T) = Code.default_typscheme thy c;
-            val T' = map_type_tvar (fn ((v, i), S) => TypeInfer.param (typidx + i) (v, [])) T;
-            val typidx' = typidx + maxidx_of_typ T' + 1;
+            val T' = map_type_tfree (fn (v, _) => TypeInfer.param typidx (v, [])) T;
+            val typidx' = typidx + 1;
           in of_apps bounds (Term.Const (c, T'), ts') typidx' end
       | of_univ bounds (Free (name, ts)) typidx =
           of_apps bounds (Term.Free (name, dummyT), ts) typidx