src/Pure/Tools/codegen_package.ML
changeset 18702 7dc7dcd63224
parent 18517 788fa99aba33
child 18704 2c86ced392a8
--- a/src/Pure/Tools/codegen_package.ML	Tue Jan 17 10:26:50 2006 +0100
+++ b/src/Pure/Tools/codegen_package.ML	Tue Jan 17 16:36:57 2006 +0100
@@ -12,10 +12,15 @@
 signature CODEGEN_PACKAGE =
 sig
   type auxtab;
-  type appgen;
+  type eqextr = theory -> auxtab
+    -> (string * typ) -> (thm list * typ) option;
   type defgen;
+  type appgen = theory -> auxtab
+    -> (string * typ) * term list -> CodegenThingol.transact -> CodegenThingol.iexpr * CodegenThingol.transact;
+
   val add_appconst: string * ((int * int) * appgen) -> theory -> theory;
   val add_appconst_i: xstring * ((int * int) * appgen) -> theory -> theory;
+  val add_eqextr: string * eqextr -> theory -> theory;
   val add_defgen: string * defgen -> theory -> theory;
   val add_prim_class: xstring -> string list -> (string * string)
     -> theory -> theory;
@@ -25,49 +30,31 @@
     -> theory -> theory;
   val add_prim_i: string -> string list -> (string * Pretty.T)
     -> theory -> theory;
-  val add_syntax_tyco: xstring -> (string * (string * CodegenSerializer.fixity))
-    -> theory -> theory;
-  val add_syntax_tyco_i: string -> (string * (CodegenThingol.itype Codegen.mixfix list * CodegenSerializer.fixity))
-      -> theory -> theory;
-  val add_syntax_const: (xstring * string option) -> (string * (string * CodegenSerializer.fixity))
-      -> theory -> theory;
-  val add_syntax_const_i: string -> (string * (CodegenThingol.iexpr Codegen.mixfix list * CodegenSerializer.fixity))
-    -> theory -> theory;
-  val add_lookup_tyco: string * string -> theory -> theory;
-  val add_lookup_const: (string * typ) * CodegenThingol.iexpr -> theory -> theory;
   val add_alias: string * string -> theory -> theory;
   val set_is_datatype: (theory -> string -> bool) -> theory -> theory;
   val set_get_all_datatype_cons : (theory -> (string * string) list)
     -> theory -> theory;
+  val set_int_tyco: string -> theory -> theory;
 
   val exprgen_type: theory -> auxtab
     -> typ -> CodegenThingol.transact -> CodegenThingol.itype * CodegenThingol.transact;
   val exprgen_term: theory -> auxtab
     -> term -> CodegenThingol.transact -> CodegenThingol.iexpr * CodegenThingol.transact;
-  val ensure_def_tyco: theory -> auxtab
-    -> string -> CodegenThingol.transact -> string * CodegenThingol.transact;
-  val ensure_def_const: theory -> auxtab
-    -> string * typ -> CodegenThingol.transact -> string * CodegenThingol.transact;
+  val appgen_default: appgen;
 
   val appgen_let: (int -> term -> term list * term)
     -> appgen;
   val appgen_split: (int -> term -> term list * term)
     -> appgen;
-  val appgen_number_of: (term -> IntInf.int) -> (term -> term)
-    -> appgen;
-  val appgen_datatype_case: (string * int) list
+  val appgen_number_of: (term -> term) -> (term -> IntInf.int) -> string -> string
     -> appgen;
-  val add_cg_case_const: (theory -> string -> (string * int) list option) -> xstring
+  val add_case_const: (theory -> string -> (string * int) list option) -> xstring
     -> theory -> theory;
-  val add_cg_case_const_i: (theory -> string -> (string * int) list option) -> string
+  val add_case_const_i: (theory -> string -> (string * int) list option) -> string
     -> theory -> theory;
   val defgen_datatype: (theory -> string -> ((string * sort) list * string list) option)
     -> (theory -> string * string -> typ list option)
     -> defgen;
-  val defgen_datacons: (theory -> string * string -> typ list option)
-    -> defgen;
-  val defgen_recfun: (theory -> string * typ -> (term list * term) list * typ)
-    -> defgen;
 
   val print_codegen_generated: theory -> unit;
   val rename_inconsistent: theory -> theory;
@@ -105,16 +92,13 @@
 
 val is_number = is_some o Int.fromString;
 
-fun newline_correct s =
-  s
-  |> space_explode "\n"
-  |> map (implode o (fn [] => []
-                      | (" "::xs) => xs
-                      | xs => xs) o explode)
-  |> space_implode "\n";
+fun merge_opt _ (x1, NONE) = x1
+  | merge_opt _ (NONE, x2) = x2
+  | merge_opt eq (SOME x1, SOME x2) =
+      if eq (x1, x2) then SOME x1 else error ("incompatible options during merge");
 
 
-(* shallo name spaces *)
+(* shallow name spaces *)
 
 val nsp_class = "class";
 val nsp_tyco = "tyco";
@@ -123,11 +107,9 @@
 val nsp_dtcon = "dtcon";
 val nsp_mem = "mem";
 val nsp_inst = "inst";
-val nsp_eq_inst = "eq_inst";
-val nsp_eq_pred = "eq";
 
 
-(* code generator data types *)
+(* code generator basics *)
 
 structure InstNameMangler = NameManglerFun (
   type ctxt = theory;
@@ -171,9 +153,7 @@
   type ctxt = theory;
   type src = string * string;
   val ord = prod_ord string_ord string_ord;
-  fun mk thy (("0", "nat"), _) =
-         "Nat.Zero"
-    | mk thy ((co, dtco), i) =
+  fun mk thy ((co, dtco), i) =
         let
           fun basename 0 = NameSpace.base co
             | basename 1 = NameSpace.base dtco ^ "_" ^ NameSpace.base co
@@ -194,69 +174,55 @@
   fun re_mangle _ dst = error ("no such datatype constructor: " ^ quote dst);
 );
 
-type auxtab = ((typ * (term list * term)) Symtab.table * string Symtab.table)
+type auxtab = ((typ * thm) list Symtab.table * string Symtab.table)
   * (InstNameMangler.T * ((typ * typ list) Symtab.table * ConstNameMangler.T) * DatatypeconsNameMangler.T);
-
-type appgen = theory -> auxtab -> ((string * typ) * term list, iexpr) gen_exprgen;
+type eqextr = theory -> auxtab
+  -> (string * typ) -> (thm list * typ) option;
 type defgen = theory -> auxtab -> gen_defgen;
-
-
-(* serializer *)
+type appgen = theory -> auxtab -> (string * typ) * term list -> transact -> iexpr * transact;
 
-val serializer_ml =
-  let
-    val name_root = "Generated";
-    val nsp_conn = [
-      [nsp_class, nsp_tyco], [nsp_const, nsp_overl, nsp_dtcon, nsp_mem, nsp_inst, nsp_eq_inst, nsp_eq_pred]
-    ];
-  in CodegenSerializer.ml_from_thingol nsp_conn nsp_class name_root end;
-
-val serializer_hs =
-  let
-    val name_root = "Generated";
-    val nsp_conn = [
-      [nsp_class], [nsp_tyco], [nsp_const, nsp_overl, nsp_mem, nsp_eq_pred], [nsp_dtcon], [nsp_inst, nsp_eq_inst]
-    ];
-  in CodegenSerializer.haskell_from_thingol nsp_conn nsp_dtcon name_root end;
+val serializers = ref (
+  Symtab.empty
+  |> Symtab.update (
+       #ml CodegenSerializer.serializers
+       |> apsnd (fn seri => seri
+            (nsp_dtcon, nsp_class, "")
+            [[nsp_class, nsp_tyco], [nsp_const, nsp_overl, nsp_dtcon, nsp_mem, nsp_inst]]
+          )
+     )
+  |> Symtab.update (
+       #haskell CodegenSerializer.serializers
+       |> apsnd (fn seri => seri
+            nsp_dtcon
+            [[nsp_class], [nsp_tyco], [nsp_const, nsp_overl, nsp_mem], [nsp_dtcon], [nsp_inst]]
+          )
+     )
+);
 
 
 (* theory data for code generator *)
 
 type gens = {
   appconst: ((int * int) * (appgen * stamp)) Symtab.table,
+  eqextrs: (string * (eqextr * stamp)) list,
   defgens: (string * (defgen * stamp)) list
 };
 
-fun map_gens f { appconst, defgens } =
+fun map_gens f { appconst, eqextrs, defgens } =
   let
-    val (appconst, defgens) =
-      f (appconst, defgens)
-  in { appconst = appconst, defgens = defgens } : gens end;
+    val (appconst, eqextrs, defgens) =
+      f (appconst, eqextrs, defgens)
+  in { appconst = appconst, eqextrs = eqextrs, defgens = defgens } : gens end;
 
 fun merge_gens
-  ({ appconst = appconst1 , defgens = defgens1 },
-   { appconst = appconst2 , defgens = defgens2 }) =
+  ({ appconst = appconst1 , eqextrs = eqextrs1, defgens = defgens1 },
+   { appconst = appconst2 , eqextrs = eqextrs2, defgens = defgens2 }) =
   { appconst = Symtab.merge
       (fn ((bounds1, (_, stamp1)), (bounds2, (_, stamp2))) => bounds1 = bounds2 andalso stamp1 = stamp2)
       (appconst1, appconst2),
-    defgens = AList.merge (op =) (eq_snd (op =)) (defgens1, defgens2) } : gens;
-
-type lookups = {
-  lookups_tyco: string Symtab.table,
-  lookups_const: (typ * iexpr) list Symtab.table
-}
-
-fun map_lookups f { lookups_tyco, lookups_const } =
-  let
-    val (lookups_tyco, lookups_const) =
-      f (lookups_tyco, lookups_const)
-  in { lookups_tyco = lookups_tyco, lookups_const = lookups_const } : lookups end;
-
-fun merge_lookups
-  ({ lookups_tyco = lookups_tyco1, lookups_const = lookups_const1 },
-   { lookups_tyco = lookups_tyco2, lookups_const = lookups_const2 }) =
-  { lookups_tyco = Symtab.merge (op =) (lookups_tyco1, lookups_tyco2),
-    lookups_const = Symtab.merge (op =) (lookups_const1, lookups_const2) } : lookups;
+    eqextrs = AList.merge (op =) (eq_snd (op =)) (eqextrs1, eqextrs2),
+    defgens = AList.merge (op =) (eq_snd (op =)) (defgens1, defgens2)
+  } : gens;
 
 type logic_data = {
   is_datatype: ((theory -> string -> bool) * stamp) option,
@@ -268,16 +234,15 @@
   let
     val ((is_datatype, get_all_datatype_cons), alias) =
       f ((is_datatype, get_all_datatype_cons), alias)
-  in { is_datatype = is_datatype, get_all_datatype_cons = get_all_datatype_cons, alias = alias } : logic_data end;
+  in { is_datatype = is_datatype, get_all_datatype_cons = get_all_datatype_cons,
+    alias = alias } : logic_data end;
 
 fun merge_logic_data
-  ({ is_datatype = is_datatype1, get_all_datatype_cons = get_all_datatype_cons1, alias = alias1 },
-   { is_datatype = is_datatype2, get_all_datatype_cons = get_all_datatype_cons2, alias = alias2 }) =
+  ({ is_datatype = is_datatype1, get_all_datatype_cons = get_all_datatype_cons1,
+       alias = alias1 },
+   { is_datatype = is_datatype2, get_all_datatype_cons = get_all_datatype_cons2,
+       alias = alias2 }) =
   let
-    fun merge_opt _ (x1, NONE) = x1
-      | merge_opt _ (NONE, x2) = x2
-      | merge_opt eq (SOME x1, SOME x2) =
-          if eq (x1, x2) then SOME x1 else error ("incompatible options during merge");
   in
     { is_datatype = merge_opt (eq_snd (op =)) (is_datatype1, is_datatype2),
       get_all_datatype_cons = merge_opt (eq_snd (op =)) (get_all_datatype_cons1, get_all_datatype_cons2),
@@ -285,28 +250,23 @@
                Symtab.merge (op =) (snd alias1, snd alias2)) } : logic_data
   end;
 
-type serialize_data = {
-  serializer: CodegenSerializer.serializer,
+type target_data = {
   syntax_tyco: (itype CodegenSerializer.pretty_syntax * stamp) Symtab.table,
   syntax_const: (iexpr CodegenSerializer.pretty_syntax * stamp) Symtab.table
 };
 
-fun map_serialize_data f { serializer, syntax_tyco, syntax_const } =
+fun map_target_data f { syntax_tyco, syntax_const } =
   let
     val (syntax_tyco, syntax_const) =
       f (syntax_tyco, syntax_const)
-  in { serializer = serializer,
-       syntax_tyco = syntax_tyco, syntax_const = syntax_const } : serialize_data
+  in { syntax_tyco = syntax_tyco, syntax_const = syntax_const } : target_data
   end;
 
-fun merge_serialize_data
-  ({ serializer = serializer,
-     syntax_tyco = syntax_tyco1, syntax_const = syntax_const1 },
-   {serializer = _,
-     syntax_tyco = syntax_tyco2, syntax_const = syntax_const2 }) =
-  { serializer = serializer,
-    syntax_tyco = Symtab.merge (eq_snd (op =)) (syntax_tyco1, syntax_tyco2),
-    syntax_const = Symtab.merge (eq_snd (op =)) (syntax_const1, syntax_const2) } : serialize_data;
+fun merge_target_data
+  ({ syntax_tyco = syntax_tyco1, syntax_const = syntax_const1 },
+   { syntax_tyco = syntax_tyco2, syntax_const = syntax_const2 }) =
+  { syntax_tyco = Symtab.merge (eq_snd (op =)) (syntax_tyco1, syntax_tyco2),
+    syntax_const = Symtab.merge (eq_snd (op =)) (syntax_const1, syntax_const2) } : target_data;
 
 structure CodegenData = TheoryDataFun
 (struct
@@ -314,50 +274,44 @@
   type T = {
     modl: module,
     gens: gens,
-    lookups: lookups,
     logic_data: logic_data,
-    serialize_data: serialize_data Symtab.table
+    target_data: target_data Symtab.table
   };
   val empty = {
     modl = empty_module,
-    gens = { appconst = Symtab.empty, defgens = [] } : gens,
-    lookups = { lookups_tyco = Symtab.empty, lookups_const = Symtab.empty } : lookups,
+    gens = { appconst = Symtab.empty, eqextrs = [], defgens = [] } : gens,
     logic_data = { is_datatype = NONE, get_all_datatype_cons = NONE,
       alias = (Symtab.empty, Symtab.empty) } : logic_data,
-    serialize_data =
+    target_data =
       Symtab.empty
-      |> Symtab.update ("ml",
-          { serializer = serializer_ml : CodegenSerializer.serializer,
-            syntax_tyco = Symtab.empty, syntax_const = Symtab.empty })
-      |> Symtab.update ("haskell",
-          { serializer = serializer_hs : CodegenSerializer.serializer,
-            syntax_tyco = Symtab.empty, syntax_const = Symtab.empty })
+      |> Symtab.fold (fn (target, _) =>
+           Symtab.update (target, { syntax_tyco = Symtab.empty, syntax_const = Symtab.empty })
+         ) (! serializers)
   } : T;
   val copy = I;
   val extend = I;
   fun merge _ (
-    { modl = modl1, gens = gens1, lookups = lookups1,
-      serialize_data = serialize_data1, logic_data = logic_data1 },
-    { modl = modl2, gens = gens2, lookups = lookups2,
-      serialize_data = serialize_data2, logic_data = logic_data2 }
+    { modl = modl1, gens = gens1,
+      target_data = target_data1, logic_data = logic_data1 },
+    { modl = modl2, gens = gens2,
+      target_data = target_data2, logic_data = logic_data2 }
   ) = {
     modl = merge_module (modl1, modl2),
     gens = merge_gens (gens1, gens2),
-    lookups = merge_lookups (lookups1, lookups2),
     logic_data = merge_logic_data (logic_data1, logic_data2),
-    serialize_data = Symtab.join (K (merge_serialize_data #> SOME))
-      (serialize_data1, serialize_data2)
+    target_data = Symtab.join (K (merge_target_data #> SOME))
+      (target_data1, target_data2)
   };
   fun print thy _ = writeln "sorry, this stuff is too complicated...";
 end);
 
 fun map_codegen_data f thy =
   case CodegenData.get thy
-   of { modl, gens, lookups, serialize_data, logic_data } =>
-      let val (modl, gens, lookups, serialize_data, logic_data) =
-        f (modl, gens, lookups, serialize_data, logic_data)
-      in CodegenData.put { modl = modl, gens = gens, lookups = lookups,
-           serialize_data = serialize_data, logic_data = logic_data } thy end;
+   of { modl, gens, target_data, logic_data } =>
+      let val (modl, gens, target_data, logic_data) =
+        f (modl, gens, target_data, logic_data)
+      in CodegenData.put { modl = modl, gens = gens,
+           target_data = target_data, logic_data = logic_data } thy end;
 
 fun print_codegen_generated thy =
   let
@@ -370,13 +324,13 @@
   let
     val c = prep_const thy raw_c;
   in map_codegen_data
-    (fn (modl, gens, lookups, serialize_data, logic_data) =>
+    (fn (modl, gens, target_data, logic_data) =>
        (modl,
         gens |> map_gens
-          (fn (appconst, defgens) =>
+          (fn (appconst, eqextrs, defgens) =>
             (appconst
              |> Symtab.update (c, (bounds, (ag, stamp ()))),
-             defgens)), lookups, serialize_data, logic_data)) thy
+             eqextrs, defgens)), target_data, logic_data)) thy
   end;
 
 val add_appconst = gen_add_appconst Sign.intern_const;
@@ -384,62 +338,36 @@
 
 fun add_defgen (name, dg) =
   map_codegen_data
-    (fn (modl, gens, lookups, serialize_data, logic_data) =>
+    (fn (modl, gens, target_data, logic_data) =>
        (modl,
         gens |> map_gens
-          (fn (appconst, defgens) =>
-            (appconst, defgens
+          (fn (appconst, eqextrs, defgens) =>
+            (appconst, eqextrs, defgens
              |> Output.update_warn (op =) ("overwriting existing definition definition generator " ^ name) (name, (dg, stamp ())))),
-             lookups, serialize_data, logic_data));
+             target_data, logic_data));
 
 fun get_defgens thy tabs =
   (map (apsnd (fn (dg, _) => dg thy tabs)) o #defgens o #gens o CodegenData.get) thy;
 
-fun add_lookup_tyco (src, dst) =
-  map_codegen_data
-    (fn (modl, gens, lookups, serialize_data, logic_data) =>
-       (modl, gens,
-        lookups |> map_lookups
-          (fn (lookups_tyco, lookups_const) =>
-            (lookups_tyco |> Symtab.update_new (src, dst),
-             lookups_const)),
-        serialize_data, logic_data));
-
-val lookup_tyco = Symtab.lookup o #lookups_tyco o #lookups o CodegenData.get;
-
-fun add_lookup_const ((src, ty), dst) =
+fun add_eqextr (name, eqx) =
   map_codegen_data
-    (fn (modl, gens, lookups, serialize_data, logic_data) =>
-       (modl, gens,
-        lookups |> map_lookups
-          (fn (lookups_tyco, lookups_const) =>
-            (lookups_tyco,
-             lookups_const |> Symtab.update_multi (src, (ty, dst)))),
-        serialize_data, logic_data));
+    (fn (modl, gens, target_data, logic_data) =>
+       (modl,
+        gens |> map_gens
+          (fn (appconst, eqextrs, defgens) =>
+            (appconst, eqextrs
+             |> Output.update_warn (op =) ("overwriting existing equation extractor " ^ name) (name, (eqx, stamp ())),
+             defgens)),
+             target_data, logic_data));
 
-fun lookup_const thy (f, ty) =
-  (Symtab.lookup_multi o #lookups_const o #lookups o CodegenData.get) thy f
-  |> (fn tab => AList.lookup (Sign.typ_instance thy) tab ty);
-
-fun set_is_datatype f =
-  map_codegen_data
-    (fn (modl, gens, lookups, serialize_data, logic_data) =>
-       (modl, gens, lookups, serialize_data,
-        logic_data
-        |> map_logic_data ((apfst o apfst) (K (SOME (f, stamp ()))))));
+fun get_eqextrs thy tabs =
+  (map (fn (_, (eqx, _)) => eqx thy tabs) o #eqextrs o #gens o CodegenData.get) thy;
 
 fun is_datatype thy =
   case (#is_datatype o #logic_data o CodegenData.get) thy
    of NONE => K false
     | SOME (f, _) => f thy;
 
-fun set_get_all_datatype_cons f =
-  map_codegen_data
-    (fn (modl, gens, lookups, serialize_data, logic_data) =>
-       (modl, gens, lookups, serialize_data,
-        logic_data
-        |> map_logic_data ((apfst o apsnd) (K (SOME (f, stamp ()))))));
-
 fun get_all_datatype_cons thy =
   case (#get_all_datatype_cons o #logic_data o CodegenData.get) thy
    of NONE => []
@@ -447,8 +375,8 @@
 
 fun add_alias (src, dst) =
   map_codegen_data
-    (fn (modl, gens, lookups, serialize_data, logic_data) =>
-       (modl, gens, lookups, serialize_data,
+    (fn (modl, gens, target_data, logic_data) =>
+       (modl, gens, target_data,
         logic_data |> map_logic_data
           (apsnd (fn (tab, tab_rev) =>
             (tab |> Symtab.update (src, dst),
@@ -457,16 +385,6 @@
 
 (* name handling *)
 
-val nsp_class = "class";
-val nsp_tyco = "tyco";
-val nsp_const = "const";
-val nsp_overl = "overl";
-val nsp_dtcon = "dtcon";
-val nsp_mem = "mem";
-val nsp_inst = "inst";
-val nsp_eq_inst = "eq_inst";
-val nsp_eq_pred = "eq";
-
 val alias_get = perhaps o Symtab.lookup o fst o #alias o #logic_data o CodegenData.get;
 val alias_rev = perhaps o Symtab.lookup o snd o #alias o #logic_data o CodegenData.get;
 
@@ -477,6 +395,7 @@
   |> apsnd (single #> cons shallow)
   |> (op @)
   |> NameSpace.pack;
+
 fun dest_nsp nsp idf =
   let
     val idf' = NameSpace.unpack idf;
@@ -489,17 +408,43 @@
   end;
 
 fun idf_of_name thy shallow name =
-  if is_number name
-  then name
-  else
-    name
-    |> alias_get thy
-    |> add_nsp shallow;
+  name
+  |> alias_get thy
+  |> add_nsp shallow;
+
 fun name_of_idf thy shallow idf =
   idf
   |> dest_nsp shallow
   |> Option.map (alias_rev thy);
 
+fun set_is_datatype f =
+  map_codegen_data
+    (fn (modl, gens, target_data, logic_data) =>
+       (modl, gens, target_data,
+        logic_data
+        |> map_logic_data (apfst (fn (is_datatype, get_all_datatype_cons)
+             => (SOME (f, stamp ()), get_all_datatype_cons)))));
+
+fun set_get_all_datatype_cons f =
+  map_codegen_data
+    (fn (modl, gens, target_data, logic_data) =>
+       (modl, gens, target_data,
+        logic_data
+        |> map_logic_data ((apfst (fn (is_datatype, get_all_datatype_cons)
+             => (is_datatype, SOME (f, stamp ())))))));
+
+fun set_int_tyco tyco thy =
+  (serializers := (
+    ! serializers
+    |> Symtab.update (
+         #ml CodegenSerializer.serializers
+         |> apsnd (fn seri => seri
+              (nsp_dtcon, nsp_class, idf_of_name thy nsp_tyco tyco)
+              [[nsp_class, nsp_tyco], [nsp_const, nsp_overl, nsp_dtcon, nsp_mem, nsp_inst]]
+            )
+       )
+    ); thy);
+
 
 (* code generator instantiation *)
 
@@ -515,7 +460,7 @@
 
 fun ensure_def_inst thy (tabs as (_, (insttab, _, _))) (cls, tyco) trns =
   let
-    val thyname = (the o AList.lookup (op =) (ClassPackage.the_tycos thy cls)) tyco;
+    val thyname = (the o AList.lookup (op =) (ClassPackage.the_instances thy cls)) tyco;
     val inst = idf_of_name thy nsp_inst (InstNameMangler.get thy insttab (thyname, (cls, tyco)));
   in
     trns
@@ -527,15 +472,11 @@
 fun ensure_def_tyco thy tabs tyco trns =
   let
     val tyco' = idf_of_name thy nsp_tyco tyco;
-  in case lookup_tyco thy tyco
-   of NONE =>
-        trns
-        |> debug 4 (fn _ => "generating type constructor " ^ quote tyco)
-        |> gen_ensure_def (get_defgens thy tabs) ("generating type constructor " ^ quote tyco) tyco'
-        |> pair tyco'
-    | SOME tyco =>
-        trns
-        |> pair tyco
+  in
+    trns
+    |> debug 4 (fn _ => "generating type constructor " ^ quote tyco)
+    |> gen_ensure_def (get_defgens thy tabs) ("generating type constructor " ^ quote tyco) tyco'
+    |> pair tyco'
   end;
 
 fun idf_of_const thy (tabs as ((_, clsmemtab), (_, (overltab1, overltab2), dtcontab))) (c, ty) =
@@ -553,10 +494,10 @@
        of Type (tyco, _) =>
             try (DatatypeconsNameMangler.get thy dtcontab) (c, tyco)
         | _ => NONE;
-  in case get_overloaded (c, ty)
+  in case get_datatypecons (c, ty)
+   of SOME c' => idf_of_name thy nsp_dtcon c'
+    | NONE => case get_overloaded (c, ty)
    of SOME idf => idf
-    | NONE => case get_datatypecons (c, ty)
-   of SOME c' => idf_of_name thy nsp_dtcon c'
     | NONE => case Symtab.lookup clsmemtab c
    of SOME _ => idf_of_name thy nsp_mem c
     | NONE => idf_of_name thy nsp_const c
@@ -564,7 +505,7 @@
 
 fun recconst_of_idf thy (_, (_, (_, overltab2), _)) idf =
   case name_of_idf thy nsp_const idf
-   of SOME c => SOME (c, Sign.the_const_constraint thy c)
+   of SOME c => SOME (c, Sign.the_const_type thy c)
     | NONE => (
         case dest_nsp nsp_overl idf
          of SOME _ =>
@@ -579,18 +520,14 @@
 fun ensure_def_const thy (tabs as ((_, clsmemtab), (_, overltab, dtcontab))) (c, ty) trns =
   let
     val c' = idf_of_const thy tabs (c, ty);
-  in case lookup_const thy (c, ty)
-   of NONE =>
-        trns
-        |> debug 4 (fn _ => "generating constant " ^ quote c)
-        |> gen_ensure_def (get_defgens thy tabs) ("generating constant " ^ quote c) c'
-        |> pair c'
-    | SOME (IConst (c, ty)) =>
-        trns
-        |> pair c
+  in
+    trns
+    |> debug 4 (fn _ => "generating constant " ^ quote c)
+    |> gen_ensure_def (get_defgens thy tabs) ("generating constant " ^ quote c) c'
+    |> pair c'
   end;
 
-fun ensure_def_eq thy tabs (dtco, (eqpred, arity)) trns =
+(* fun ensure_def_eq thy tabs (dtco, (eqpred, arity)) trns =
   let
     val name_dtco = (the ooo name_of_idf) thy nsp_tyco dtco;
     val idf_eqinst = idf_of_name thy nsp_eq_inst name_dtco;
@@ -598,30 +535,30 @@
     val inst_sortlookup = map (fn (v, _) => [ClassPackage.Lookup ([], (v, 0))]) arity;
     fun mk_eq_pred _ trns =
       trns
-      |> succeed (eqpred, [])
+      |> succeed (eqpred)
     fun mk_eq_inst _ trns =
       trns
       |> gen_ensure_def [("eqpred", mk_eq_pred)] ("generating equality predicate for " ^ quote dtco) idf_eqpred
-      |> succeed (Classinst ((class_eq, (dtco, arity)), ([], [(fun_eq, (idf_eqpred, inst_sortlookup))])), []);
+      |> succeed (Classinst ((class_eq, (dtco, arity)), ([], [(fun_eq, (idf_eqpred, inst_sortlookup))])));
   in
     trns
     |> gen_ensure_def [("eqinst", mk_eq_inst)] ("generating equality instance for " ^ quote dtco) idf_eqinst
-  end;
+  end; *)
 
 
 (* expression generators *)
 
-fun exprgen_sort thy tabs sort trns =
+fun exprgen_tyvar_sort thy tabs (v, sort) trns =
   trns
   |> fold_map (ensure_def_class thy tabs) (ClassPackage.syntactic_sort_of thy sort)
-  |-> (fn sort => pair sort);
+  |-> (fn sort => pair (unprefix "'" v, sort));
 
 fun exprgen_type thy tabs (TVar _) trns =
       error "TVar encountered during code generation"
-  | exprgen_type thy tabs (TFree (v, sort)) trns =
+  | exprgen_type thy tabs (TFree v_s) trns =
       trns
-      |> exprgen_sort thy tabs sort
-      |-> (fn sort => pair (IVarT (v |> unprefix "'", sort)))
+      |> exprgen_tyvar_sort thy tabs v_s
+      |-> (fn v_s => pair (IVarT v_s))
   | exprgen_type thy tabs (Type ("fun", [t1, t2])) trns =
       trns
       |> exprgen_type thy tabs t1
@@ -644,19 +581,19 @@
       |> fold_map (ensure_def_class thy tabs) clss
       |-> (fn clss => pair (ClassPackage.Lookup (clss, (v |> unprefix "'", i))));
 
-fun mk_itapp e [] = e
-  | mk_itapp e lookup = IInst (e, lookup);
-
-fun appgen thy tabs ((f, ty), ts) trns =
+fun appgen_default thy tabs ((c, ty), ts) trns =
+  trns
+  |> ensure_def_const thy tabs (c, ty)
+  ||>> (fold_map o fold_map) (mk_lookup thy tabs)
+    (ClassPackage.extract_sortlookup thy (c, ty))
+  ||>> exprgen_type thy tabs ty
+  ||>> fold_map (exprgen_term thy tabs) ts
+  |-> (fn (((c, ls), ty), es) =>
+         pair (Library.foldl IInst ((IConst (c, ty)), ls) `$$ es))
+and appgen thy tabs ((f, ty), ts) trns =
   case Symtab.lookup ((#appconst o #gens o CodegenData.get) thy) f
    of SOME ((imin, imax), (ag, _)) =>
-        let
-          fun invoke ts trns =
-            trns
-            |> gen_invoke [("const application", ag thy tabs)] ("generating application " ^ f ^ "::" ^ (Sign.string_of_typ thy) ty
-              ^ " " ^ enclose "(" ")" (commas (map (Sign.string_of_term thy) ts)))
-              ((f, ty), ts)
-        in if length ts < imin then
+        if length ts < imin then
           let
             val d = imin - length ts;
             val vs = Term.invent_names (add_term_names (Const (f, ty), [])) "x" d;
@@ -665,29 +602,22 @@
             trns
             |> debug 10 (fn _ => "eta-expanding")
             |> fold_map (exprgen_type thy tabs) tys
-            ||>> invoke (ts @ map2 (curry Free) vs tys)
+            ||>> ag thy tabs ((f, ty), ts @ map2 (curry Free) vs tys)
             |-> (fn (tys, e) => pair ((vs ~~ tys) `|--> e))
           end
         else if length ts > imax then
           trns
           |> debug 10 (fn _ => "splitting arguments (" ^ string_of_int imax ^ ", " ^ string_of_int (length ts) ^ ")")
-          |> invoke  (Library.take (imax, ts))
+          |> ag thy tabs ((f, ty), Library.take (imax, ts))
           ||>> fold_map (exprgen_term thy tabs) (Library.drop (imax, ts))
           |-> (fn es => pair (mk_apps es))
         else
           trns
           |> debug 10 (fn _ => "keeping arguments")
-          |> invoke ts
-        end
+          |> ag thy tabs ((f, ty), ts)
     | NONE =>
         trns
-        |> ensure_def_const thy tabs (f, ty)
-        ||>> (fold_map o fold_map) (mk_lookup thy tabs)
-          (ClassPackage.extract_sortlookup thy (Sign.the_const_constraint thy f, ty))
-        ||>> exprgen_type thy tabs ty
-        ||>> fold_map (exprgen_term thy tabs) ts
-        |-> (fn (((f, lookup), ty), es) =>
-               pair (mk_itapp (IConst (f, ty)) lookup `$$ es))
+        |> appgen_default thy tabs ((f, ty), ts)
 and exprgen_term thy tabs (Const (f, ty)) trns =
       trns
       |> appgen thy tabs ((f, ty), [])
@@ -723,103 +653,106 @@
 
 (* application generators *)
 
-fun appgen_neg thy tabs (("neg", Type ("fun", [ty, _])), ts) trns =
-  trns
-  |> exprgen_term thy tabs (Const ("op >", ty --> ty --> Type ("bool", [])) $ Const ("0", ty))
-  |-> succeed;
-
-fun appgen_eq thy tabs (("op =", Type ("fun", [ty, _])), [t1, t2]) trns =
+(* fun appgen_eq thy tabs (("op =", Type ("fun", [ty, _])), [t1, t2]) trns =
   trns
   |> invoke_eq (exprgen_type thy tabs) (ensure_def_eq thy tabs) ty
   |-> (fn false => error ("could not derive equality for " ^ Sign.string_of_typ thy ty)
         | true => fn trns => trns
   |> exprgen_term thy tabs t1
   ||>> exprgen_term thy tabs t2
-  |-> (fn (e1, e2) => succeed (Fun_eq `$ e1 `$ e2)));
+  |-> (fn (e1, e2) => pair (Fun_eq `$ e1 `$ e2))); *)
+
+
+(* function extractors *)
+
+fun mk_fun thy tabs (c, ty) trns =
+  case get_first (fn eqx => eqx (c, ty)) (get_eqextrs thy tabs)
+   of SOME (eq_thms, ty) =>
+        let
+          val sortctxt = ClassPackage.extract_sortctxt thy ty;
+          fun dest_eqthm eq_thm =
+            eq_thm
+            |> prop_of
+            |> Logic.dest_equals
+            |> apfst (strip_comb #> snd);
+          fun mk_eq (args, rhs) trns =
+            trns
+            |> fold_map (exprgen_term thy tabs o devarify_term) args
+            ||>> (exprgen_term thy tabs o devarify_term) rhs
+            |-> (fn (args, rhs) => pair (map ipat_of_iexpr args, rhs))
+        in
+          trns
+          |> fold_map (mk_eq o dest_eqthm) eq_thms
+          ||>> exprgen_type thy tabs (devarify_type ty)
+          ||>> fold_map (exprgen_tyvar_sort thy tabs) sortctxt
+          |-> (fn ((eqs, ty), sortctxt) => (pair o SOME) (eqs, (sortctxt, ty)))
+        end
+    | NONE => (NONE, trns);
+
+fun eqextr_defs thy ((deftab, _), _) (c, ty) =
+  let
+    fun eq_typ (ty1, ty2) = 
+      Sign.typ_instance thy (ty1, ty2)
+      andalso Sign.typ_instance thy (ty2, ty1)
+  in
+    Option.mapPartial (get_first (fn (ty', thm) => if eq_typ (ty, ty')
+      then SOME ([thm], ty')
+      else NONE
+    )) (Symtab.lookup deftab c)
+  end;
 
 
 (* definition generators *)
 
-fun mk_fun thy tabs eqs ty trns =
-  let
-    val sortctxt = ClassPackage.extract_sortctxt thy ty;
-    fun mk_sortvar (v, sort) trns =
-      trns
-      |> exprgen_sort thy tabs sort
-      |-> (fn sort => pair (unprefix "'" v, sort))
-    fun mk_eq (args, rhs) trns =
-      trns
-      |> fold_map (exprgen_term thy tabs o devarify_term) args
-      ||>> (exprgen_term thy tabs o devarify_term) rhs
-      |-> (fn (args, rhs) => pair (map ipat_of_iexpr args, rhs))
-  in
-    trns
-    |> fold_map mk_eq eqs
-    ||>> exprgen_type thy tabs (devarify_type ty)
-    ||>> fold_map mk_sortvar sortctxt
-    |-> (fn ((eqs, ty), sortctxt) => pair (Fun (eqs, (sortctxt, ty))))
-  end;
-
-fun defgen_tyco_fallback thy tabs tyco trns =
-  if Symtab.fold (fn (_, { syntax_tyco, ... }) => fn b => b orelse Symtab.defined syntax_tyco tyco)
-    ((#serialize_data o CodegenData.get) thy) false
-  then
-    trns
-    |> debug 5 (fn _ => "trying defgen tyco fallback for " ^ quote tyco)
-    |> succeed (Nop, [])
-  else
-    trns
-    |> fail ("no code generation fallback for " ^ quote tyco)
-
-fun defgen_const_fallback thy tabs c trns =
-  if Symtab.fold (fn (_, { syntax_const, ... }) => fn b => b orelse Symtab.defined syntax_const c)
-    ((#serialize_data o CodegenData.get) thy) false
-  then
-    trns
-    |> debug 5 (fn _ => "trying defgen const fallback for " ^ quote c)
-    |> succeed (Nop, [])
-  else
-    trns
-    |> fail ("no code generation fallback for " ^ quote c)
-
-fun defgen_defs thy (tabs as ((deftab, _), _)) c trns =
-  case Symtab.lookup deftab c
-   of SOME (ty, (args, rhs)) =>
-        trns
-        |> debug 5 (fn _ => "trying defgen def for " ^ quote c)
-        |> mk_fun thy tabs [(args, rhs)] (devarify_type ty)
-        |-> (fn def => succeed (def, []))
-      | _ => trns |> fail ("no definition found for " ^ quote c);
-
-fun defgen_clsdecl thy (tabs as (_, (insttab, _, _))) (cls : string) trns =
+fun defgen_clsdecl thy (tabs as (_, (insttab, _, _))) cls trns =
   case name_of_idf thy nsp_class cls
    of SOME cls =>
         let
-          val memnames = ClassPackage.the_consts thy (cls : string);
-          val memtypes = map (devarify_type o ClassPackage.get_const_sign thy "'a") memnames;
-          val memctxt = map (ClassPackage.extract_sortctxt thy) memtypes;
-          val memidfs = map (idf_of_name thy nsp_mem) memnames;
-          fun mk_instname (tyco, thyname) = idf_of_name thy nsp_inst (InstNameMangler.get thy insttab (thyname, (cls, tyco)))
-          val instnames = map mk_instname (ClassPackage.the_tycos thy cls);
+          val cs = (snd o ClassPackage.the_consts_sign thy) cls;
+          val sortctxts = map (ClassPackage.extract_sortctxt thy o snd) cs;
+          val idfs = map (idf_of_name thy nsp_mem o fst) cs;
         in
           trns
           |> debug 5 (fn _ => "trying defgen class declaration for " ^ quote cls)
-          |> fold_map (ensure_def_class thy tabs) (ClassPackage.get_superclasses thy cls)
-          ||>> fold_map (exprgen_type thy tabs) memtypes
-          |-> (fn (supcls, memtypes) => succeed (Class (supcls, "a", memidfs ~~ (memctxt ~~ memtypes), []),
-                 memidfs @ instnames))
+          |> fold_map (ensure_def_class thy tabs) (ClassPackage.the_superclasses thy cls)
+          ||>> fold_map (exprgen_type thy tabs o snd) cs
+          ||>> (fold_map o fold_map) (exprgen_tyvar_sort thy tabs) sortctxts
+          |-> (fn ((supcls, memtypes), sortctxts) => succeed
+            (Class ((supcls, ("a", idfs ~~ (sortctxts ~~ memtypes))), [])))
         end
     | _ =>
         trns
         |> fail ("no class definition found for " ^ quote cls);
 
+fun defgen_funs thy tabs c trns =
+  case recconst_of_idf thy tabs c
+   of SOME (c, ty) =>
+        trns
+        |> mk_fun thy tabs (c, ty)
+        |-> (fn (SOME funn) => succeed (Fun funn)
+              | NONE => fail ("no defining equations found for " ^ quote c))
+    | NONE =>
+        trns
+        |> fail ("not a constant: " ^ quote c);
+
+fun defgen_datatypecons thy (tabs as (_, (_, _, dtcontab))) co trns =
+  case Option.map (DatatypeconsNameMangler.rev thy dtcontab) (name_of_idf thy nsp_dtcon co)
+   of SOME (co, dtco) =>
+        trns
+        |> debug 5 (fn _ => "trying defgen datatype constructor for " ^ quote co)
+        |> ensure_def_tyco thy tabs dtco
+        |-> (fn dtco => succeed Undef)
+    | _ =>
+        trns
+        |> fail ("not a datatype constructor: " ^ quote co);
+
 fun defgen_clsmem thy tabs m trns =
   case name_of_idf thy nsp_mem m
    of SOME m =>
         trns
         |> debug 5 (fn _ => "trying defgen class member for " ^ quote m)
         |> ensure_def_class thy tabs ((the o ClassPackage.lookup_const_class thy) m)
-        |-> (fn cls => succeed (Classmember cls, []))
+        |-> (fn cls => succeed Undef)
     | _ =>
         trns |> fail ("no class member found for " ^ quote m)
 
@@ -827,70 +760,49 @@
   case Option.map (InstNameMangler.rev thy insttab) (name_of_idf thy nsp_inst inst)
    of SOME (_, (cls, tyco)) =>
         let
-          val arity = ClassPackage.get_arities thy [cls] tyco;
-          val ms = map (fn m => (m, Sign.the_const_constraint thy m)) (ClassPackage.the_consts thy cls);
-          val instmem_idfs = ClassPackage.get_inst_consts_sign thy (tyco, cls);
-          val supclss = ClassPackage.get_superclasses thy cls;
-          fun add_vars arity clsmems (trns as (_, modl)) =
-            case get_def modl (idf_of_name thy nsp_class cls)
-             of Class (_, _, members, _) => ((Term.invent_names
-              (tvars_of_itypes ((map (snd o snd)) members)) "a" (length arity) ~~ arity, clsmems), trns)
-          val ad_hoc_arity = map (fn (v, sort) => map_index (fn (i, _) => (ClassPackage.Lookup ([], (v, i)))) sort);
-          (*! THIS IS ACTUALLY VERY AD-HOC... !*)
+          val (arity, memdefs) = ClassPackage.the_inst_sign thy (cls, tyco);
+          fun gen_membr (m, ty) trns =
+            trns
+            |> mk_fun thy tabs (m, ty)
+            |-> (fn SOME funn => pair (m, funn)
+                  | NONE => error ("could not derive definition for member " ^ quote m));
         in
-          (trns
+          trns
           |> debug 5 (fn _ => "trying defgen class instance for (" ^ quote cls ^ ", " ^ quote tyco ^ ")")
-          |> (fold_map o fold_map) (ensure_def_class thy tabs) arity
-          ||>> fold_map (ensure_def_const thy tabs) ms
-          |-> (fn (arity, ms) => add_vars arity ms)
-          ||>> ensure_def_class thy tabs cls
+          |> ensure_def_class thy tabs cls
           ||>> ensure_def_tyco thy tabs tyco
-          ||>> fold_map (fn supcls => ensure_def_inst thy tabs (supcls, tyco)) supclss
-          ||>> fold_map (fn supcls => (fold_map o fold_map) (mk_lookup thy tabs)
-                 (ClassPackage.extract_sortlookup thy
-                   (Type (tyco, map_index (fn (i, _) => TVar (("'a", i), [])) (ClassPackage.get_arities thy [supcls] tyco)),
-                    Type (tyco, map_index (fn (i, sort) => TFree (string_of_int i, sort)) arity)))) supclss
-          ||>> fold_map (ensure_def_const thy tabs) instmem_idfs)
-          |-> (fn ((((((arity, ms), cls), tyco), supinsts), supinstlookup), instmem_idfs) 
-                : ((((((string * string list) list * string list) * string) * string)
-                * string list) * ClassPackage.sortlookup list list list) * string list
-                =>
-                 succeed (Classinst ((cls, (tyco, arity)), (supclss ~~ (supinsts ~~ supinstlookup), ms ~~ map (rpair (ad_hoc_arity arity)) instmem_idfs)), [])) 
+          ||>> fold_map (exprgen_tyvar_sort thy tabs) arity
+          ||>> fold_map gen_membr memdefs
+          |-> (fn (((cls, tyco), arity), memdefs) =>
+                 succeed (Classinst ((cls, (tyco, arity)), memdefs)))
         end
     | _ =>
         trns |> fail ("no class instance found for " ^ quote inst);
 
 
-(*    trns
-    |> ensure_def_const thy tabs (f, ty)
-
-    ||>> exprgen_type thy tabs ty
-    ||>> fold_map (exprgen_term thy tabs) ts
-    |-> (fn (((f, lookup), ty), es) =>
-           succeed (mk_itapp (IConst (f, ty)) lookup `$$ es))*)
-
-
 (* parametrized generators, for instantiation in HOL *)
 
-fun appgen_let strip_abs thy tabs (c, [t2, t3]) trns =
+fun appgen_let strip_abs thy tabs ((c, ty), [t2, t3]) trns =
   let
-    fun dest_let (l as Const ("Let", _) $ t $ u) =
-          (case strip_abs 1 u
-           of ([p], u') => apfst (cons (p, t)) (dest_let u')
-            | _ => ([], l))
+    fun dest_let (l as Const (c', _) $ t $ u) =
+          if c = c' then
+            case strip_abs 1 u
+             of ([p], u') => apfst (cons (p, t)) (dest_let u')
+              | _ => ([], l)
+          else ([], t)
       | dest_let t = ([], t);
     fun mk_let (l, r) trns =
       trns
       |> exprgen_term thy tabs l
       ||>> exprgen_term thy tabs r
       |-> (fn (l, r) => pair (r, ipat_of_iexpr l));
-    val (lets, body) = dest_let (Const c $ t2 $ t3)
+    val (lets, body) = dest_let (Const (c, ty) $ t2 $ t3)
   in
     trns
     |> fold_map mk_let lets
     ||>> exprgen_term thy tabs body
     |-> (fn (lets, body) =>
-         succeed (Library.foldr (fn ((e, p), body) => ICase (e, [(p, body)])) (lets, body)))
+         pair (Library.foldr (fn ((e, p), body) => ICase (e, [(p, body)])) (lets, body)))
   end
 
 fun appgen_split strip_abs thy tabs (c, [t2]) trns =
@@ -900,20 +812,19 @@
     trns
     |> exprgen_term thy tabs p
     ||>> exprgen_term thy tabs body
-    |-> (fn (IVarE v, body) => succeed (IAbs (v, body)))
+    |-> (fn (IVarE v, body) => pair (IAbs (v, body)))
   end;
 
-fun appgen_number_of dest_binum mk_int_to_nat thy tabs (("Numeral.number_of",
-      Type ("fun", [_, Type ("IntDef.int", [])])), [bin]) trns =
-        trns
-        |> (succeed (IConst ((IntInf.toString o dest_binum) bin, Type_integer))
-            handle TERM _
-            => error ("not a number: " ^ Sign.string_of_term thy bin))
-  | appgen_number_of dest_binum mk_int_to_nat thy tabs (("Numeral.number_of",
-      Type ("fun", [_, Type ("nat", [])])), [bin]) trns =
-        trns
-        |> exprgen_term thy tabs (mk_int_to_nat bin)
-        |-> succeed;
+fun appgen_number_of mk_int_to_nat bin_to_int tyco_int tyco_nat thy tabs ((_,
+  Type (_, [_, ty as Type (tyco, [])])), [bin]) trns =
+    if tyco = tyco_int then
+      trns
+      |> exprgen_type thy tabs ty
+      |-> (fn ty => pair (CodegenThingol.IConst ((IntInf.toString o bin_to_int) bin, ty)))
+    else if tyco = tyco_nat then
+      trns
+      |> exprgen_term thy tabs (mk_int_to_nat bin)
+    else error ("invalid type constructor for numeral: " ^ quote tyco);
 
 fun appgen_datatype_case cos thy tabs ((_, ty), ts) trns =
   let
@@ -938,13 +849,13 @@
     trns
     |> exprgen_term thy tabs t
     ||>> fold_map cg_case_d ((cos ~~ tys) ~~ ts')
-    |-> (fn (t, ds) => succeed (ICase (t, ds)))
+    |-> (fn (t, ds) => pair (ICase (t, ds)))
   end;
 
-fun gen_add_cg_case_const prep_c get_case_const_data raw_c thy =
+fun gen_add_case_const prep_c get_case_const_data raw_c thy =
   let
     val c = prep_c thy raw_c;
-    val (tys, dty) = (split_last o fst o strip_type o Sign.the_const_constraint thy) c;
+    val (tys, dty) = (split_last o fst o strip_type o Sign.the_const_type thy) c;
     val cos = (the o get_case_const_data thy) c;
     val n_eta = length cos + 1;
   in
@@ -952,8 +863,8 @@
     |> add_appconst_i (c, ((n_eta, n_eta), appgen_datatype_case cos))
   end;
 
-val add_cg_case_const = gen_add_cg_case_const Sign.intern_const;
-val add_cg_case_const_i = gen_add_cg_case_const (K I);
+val add_case_const = gen_add_case_const Sign.intern_const;
+val add_case_const_i = gen_add_case_const (K I);
 
 fun defgen_datatype get_datatype get_datacons thy (tabs as (_, (_, _, dtcontab))) dtco trns =
   case name_of_idf thy nsp_tyco dtco
@@ -967,11 +878,10 @@
               in
                 trns
                 |> debug 5 (fn _ => "trying defgen datatype for " ^ quote dtco)
-                |> fold_map (exprgen_sort thy tabs) (map snd vars)
+                |> fold_map (exprgen_tyvar_sort thy tabs) vars
                 ||>> (fold_map o fold_map) (exprgen_type thy tabs o devarify_type) cotys
                 |-> (fn (sorts, tys) => succeed (Datatype
-                     (map2 (fn (v, _) => fn sort => (unprefix "'" v, sort)) vars sorts, coidfs ~~ tys, []),
-                     coidfs))
+                     ((sorts, coidfs ~~ tys), [])))
               end
           | NONE =>
               trns
@@ -980,38 +890,6 @@
         trns
         |> fail ("not a type constructor: " ^ quote dtco)
 
-fun defgen_datacons get_datacons thy (tabs as (_, (_, _, dtcontab))) co trns =
-  case Option.map (DatatypeconsNameMangler.rev thy dtcontab) (name_of_idf thy nsp_dtcon co)
-   of SOME (co, dtco) =>
-        trns
-        |> debug 5 (fn _ => "trying defgen datatype constructor for " ^ quote co)
-        |> ensure_def_tyco thy tabs dtco
-        ||>> fold_map (exprgen_type thy tabs) ((the o get_datacons thy) (co, dtco))
-        |-> (fn (tyco, _) => succeed (Datatypecons tyco, []))
-    | _ =>
-        trns
-        |> fail ("not a datatype constructor: " ^ quote co);
-
-fun defgen_recfun get_equations thy tabs c trns =
-  case recconst_of_idf thy tabs c
-   of SOME (c, ty) =>
-        let
-          val (eqs, ty) = get_equations thy (c, ty);
-        in
-          case eqs
-           of (_::_) =>
-                trns
-                |> debug 5 (fn _ => "trying defgen recfun for " ^ quote c)
-                |> mk_fun thy tabs eqs (devarify_type ty)
-                |-> (fn def => succeed (def, []))
-            | _ =>
-                trns
-                |> fail ("no recursive definition found for " ^ quote c)
-        end
-    | NONE =>
-        trns
-        |> fail ("not a constant: " ^ quote c);
-
 
 
 (** theory interface **)
@@ -1020,46 +898,43 @@
   let
     fun extract_defs thy =
       let
-        fun dest t =
+        fun dest tm =
           let
-            val (lhs, rhs) = Logic.dest_equals t;
-            val (c, args) = strip_comb lhs;
-            val (s, T) = dest_Const c
-          in if forall is_Var args then SOME (s, (T, (args, rhs))) else NONE
+            val (lhs, rhs) = Logic.dest_equals (prop_of tm);
+            val (t, args) = strip_comb lhs;
+            val (c, ty) = dest_Const t
+          in if forall is_Var args then SOME ((c, ty), tm) else NONE
           end handle TERM _ => NONE;
         fun prep_def def = (case Codegen.preprocess thy [def] of
-          [def'] => prop_of def' | _ => error "mk_auxtab: bad preprocessor");
-        fun add_def (name, t) defs = (case dest t of
-            NONE => defs
-          | SOME _ => (case (dest o prep_def oo Thm.get_axiom) thy name of
-              NONE => defs
-            | SOME (s, (T, (args, rhs))) => Symtab.update
-                (s, (T, (split_last (args @ [rhs]))) ::
-                if_none (Symtab.lookup defs s) []) defs))
+          [def'] => def' | _ => error "mk_auxtab: bad preprocessor");
+        fun add_def (name, _) =
+          case (dest o prep_def o Thm.get_axiom thy) name
+           of SOME ((c, ty), tm) =>
+                Symtab.default (c, []) #> Symtab.map_entry c (cons (ty, tm))
+            | NONE => I
       in
         Symtab.empty
-        |> fold (Symtab.fold add_def) (map
-             (snd o #axioms o Theory.rep_theory) (thy :: Theory.ancestors_of thy))
+        |> fold (Symtab.fold add_def o snd o #axioms o Theory.rep_theory)
+             (thy :: Theory.ancestors_of thy)
       end;
     fun mk_insttab thy =
       InstNameMangler.empty
       |> Symtab.fold_map
-          (fn (cls, (_, clsinsts)) => fold_map
+          (fn (cls, (clsmems, clsinsts)) => fold_map
             (fn (tyco, thyname) => InstNameMangler.declare thy (thyname, (cls, tyco))) clsinsts)
                  (ClassPackage.get_classtab thy)
       |-> (fn _ => I);
-    fun mk_overltabs thy defs =
+    fun mk_overltabs thy deftab =
       (Symtab.empty, ConstNameMangler.empty)
       |> Symtab.fold
           (fn (c, [_]) => I
-            | ("0", _) => I
             | (c, tytab) =>
                 (fn (overltab1, overltab2) => (
                   overltab1
                   |> Symtab.update_new (c, (Sign.the_const_constraint thy c, map fst tytab)),
                   overltab2
                   |> fold (fn (ty, _) => ConstNameMangler.declare thy (idf_of_name thy nsp_overl c, (Sign.the_const_constraint thy c, ty)) #> snd) tytab
-                ))) defs;
+                ))) deftab;
     fun mk_dtcontab thy =
       DatatypeconsNameMangler.empty
       |> fold_map
@@ -1070,43 +945,31 @@
               in AList.default (op =) (key, []) #> AList.map_entry (op =) key (cons (co, dtco)) end
             ) (get_all_datatype_cons thy) [])
       |-> (fn _ => I);
-    fun mk_deftab thy defs overltab =
-      Symtab.empty
-      |> Symtab.fold
-          (fn (c, [ty_cdef]) =>
-                Symtab.update_new (idf_of_name thy nsp_const c, ty_cdef)
-            | ("0", _) => I
-            | (c, cdefs) => fold (fn (ty, cdef) =>
-                let
-                  val c' = ConstNameMangler.get thy overltab
-                    (idf_of_name thy nsp_overl c, (Sign.the_const_constraint thy c, ty))
-                in Symtab.update_new (c', (ty, cdef)) end) cdefs) defs;
     fun mk_clsmemtab thy =
       Symtab.empty
       |> Symtab.fold
           (fn (class, (clsmems, _)) => fold
             (fn clsmem => Symtab.update (clsmem, class)) clsmems)
               (ClassPackage.get_classtab thy);
-    val defs = extract_defs thy;
+    val deftab = extract_defs thy;
     val insttab = mk_insttab thy;
-    val overltabs = mk_overltabs thy defs;
+    val overltabs = mk_overltabs thy deftab;
     val dtcontab = mk_dtcontab thy;
-    val deftab = mk_deftab thy defs (snd overltabs);
     val clsmemtab = mk_clsmemtab thy;
   in ((deftab, clsmemtab), (insttab, overltabs, dtcontab)) end;
 
-fun check_for_serializer serial_name serialize_data =
-  if Symtab.defined serialize_data serial_name
-  then serialize_data
-  else error ("unknown code serializer: " ^ quote serial_name);
+fun check_for_target thy target =
+  if (Symtab.defined o #target_data o CodegenData.get) thy target
+  then ()
+  else error ("unknown code target language: " ^ quote target);
 
 fun map_module f =
-  map_codegen_data (fn (modl, gens, lookups, serialize_data, logic_data) =>
-    (f modl, gens, lookups, serialize_data, logic_data));
+  map_codegen_data (fn (modl, gens, target_data, logic_data) =>
+    (f modl, gens, target_data, logic_data));
 
-fun expand_module defs gen thy =
+fun expand_module gen thy =
   (#modl o CodegenData.get) thy
-  |> start_transact (gen thy defs)
+  |> start_transact (gen thy (mk_tabs thy))
   |-> (fn x:'a => fn modl => (x, map_module (K modl) thy));
 
 fun rename_inconsistent thy =
@@ -1141,7 +1004,7 @@
       then
         (warning ("case constant " ^ quote case_c ^ " already present in application table, will not overwrite"); thy)
       else
-        add_cg_case_const_i get_case_const_data case_c thy;
+        add_case_const_i get_case_const_data case_c thy;
   in
     fold ensure (get_datatype_case_consts thy) thy
   end;
@@ -1152,17 +1015,28 @@
 
 (* primitive definitions *)
 
+fun read_typ thy =
+  Sign.read_typ (thy, K NONE);
+
 fun read_const thy (raw_c, raw_ty) =
   let
     val c = Sign.intern_const thy raw_c;
+    val _ = if Sign.declared_const thy c
+      then ()
+      else error ("no such constant: " ^ quote c);
     val ty = case raw_ty
      of NONE => Sign.the_const_constraint thy c
-      | SOME raw_ty => Sign.read_typ (thy, K NONE) raw_ty;
+      | SOME raw_ty => read_typ thy raw_ty;
   in (c, ty) end;
 
+fun read_quote reader gen raw thy =
+  expand_module
+    (fn thy => fn tabs => gen thy tabs (reader thy raw))
+    thy;
+
 fun gen_add_prim prep_name prep_primdef raw_name deps (target, raw_primdef) thy =
   let
-    val _ = if Symtab.defined ((#serialize_data o CodegenData.get) thy) target
+    val _ = if Symtab.defined ((#target_data o CodegenData.get) thy) target
       then () else error ("unknown target language: " ^ quote target);
     val tabs = mk_tabs thy;
     val name = prep_name thy tabs raw_name;
@@ -1175,133 +1049,128 @@
 val add_prim_i = gen_add_prim ((K o K) I) I;
 val add_prim_class = gen_add_prim
   (fn thy => K (idf_of_name thy nsp_class o Sign.intern_class thy))
-  (Pretty.str o newline_correct o Symbol.strip_blanks);
+  (Pretty.str o CodegenSerializer.parse_targetdef I);
 val add_prim_tyco = gen_add_prim
   (fn thy => K (idf_of_name thy nsp_tyco o Sign.intern_type thy))
-  (Pretty.str o newline_correct o Symbol.strip_blanks);
+  (Pretty.str o CodegenSerializer.parse_targetdef I);
 val add_prim_const = gen_add_prim
   (fn thy => fn tabs => idf_of_const thy tabs o read_const thy)
-  (Pretty.str o newline_correct o Symbol.strip_blanks);
+  (Pretty.str o CodegenSerializer.parse_targetdef I);
 
-val ensure_prim = (map_module o CodegenThingol.ensure_prim);
+val ensure_prim = (map_module oo CodegenThingol.ensure_prim);
 
 
 (* syntax *)
 
-fun gen_prep_mfx read_quote mk_quote tabs mfx thy =
-  let
-    val proto_mfx = Codegen.parse_mixfix (read_quote thy) mfx;
-    fun generate thy tabs = fold_map (mk_quote thy tabs)
-      (Codegen.quotes_of proto_mfx);
-  in
-    thy
-    |> expand_module tabs generate
-    |-> (fn tys => pair (Codegen.replace_quotes tys proto_mfx))
-  end;
-
-fun gen_add_syntax_tyco prep_tyco prep_mfx raw_tyco (serial_name, (raw_mfx, fixity)) thy =
+val parse_syntax_tyco =
   let
-    val tyco = prep_tyco thy raw_tyco;
-    val tabs = mk_tabs thy;
+    fun mk reader raw_tyco target thy =
+      let
+        val _ = check_for_target thy target;
+        fun check_tyco tyco =
+          if Sign.declared_tyname thy tyco
+          then tyco
+          else error ("no such type constructor: " ^ quote tyco);
+        fun prep_tyco thy tyco =
+          tyco
+          |> Sign.intern_type thy
+          |> check_tyco
+          |> idf_of_name thy nsp_tyco;
+        val tyco = prep_tyco thy raw_tyco;
+      in
+        thy
+        |> ensure_prim tyco target
+        |> reader
+        |-> (fn pretty => map_codegen_data
+          (fn (modl, gens, target_data, logic_data) =>
+             (modl, gens,
+              target_data |> Symtab.map_entry target
+                (map_target_data
+                  (fn (syntax_tyco, syntax_const) =>
+                   (syntax_tyco |> Symtab.update_new
+                      (tyco, (pretty, stamp ())),
+                    syntax_const))),
+              logic_data)))
+      end;
   in
-    thy
-    |> ensure_prim tyco
-    |> prep_mfx tabs raw_mfx
-    |-> (fn mfx => map_codegen_data
-      (fn (modl, gens, lookups, serialize_data, logic_data) =>
-         (modl, gens, lookups,
-          serialize_data |> check_for_serializer serial_name |> Symtab.map_entry serial_name
-            (map_serialize_data
-              (fn (syntax_tyco, syntax_const) =>
-               (syntax_tyco |> Symtab.update_new
-                  (tyco,
-                   (((Codegen.num_args_of mfx, fixity), Codegen.fillin_mixfix mfx), stamp ())),
-                syntax_const))),
-          logic_data)))
+    CodegenSerializer.parse_syntax (read_quote read_typ exprgen_type)
+    #-> (fn reader => pair (mk reader))
   end;
 
-val add_syntax_tyco_i = gen_add_syntax_tyco (K I) (K pair);
-val add_syntax_tyco = gen_add_syntax_tyco
-  (fn thy => idf_of_name thy nsp_tyco o Sign.intern_type thy)
-  (gen_prep_mfx (fn thy => typ_of o read_ctyp thy)
-    (fn thy => fn tabs => exprgen_type thy tabs o devarify_type));
-
-fun gen_add_syntax_const prep_const prep_mfx raw_c (serial_name, (raw_mfx, fixity)) thy =
+val parse_syntax_const =
   let
-    val tabs = mk_tabs thy;
-    val c = prep_const thy tabs raw_c;
+    fun mk reader raw_const target thy =
+      let
+        val _ = check_for_target thy target;
+        val tabs = mk_tabs thy;
+        val c = idf_of_const thy tabs (read_const thy raw_const);
+      in
+        thy
+        |> ensure_prim c target
+        |> reader
+        |-> (fn pretty => map_codegen_data
+          (fn (modl, gens, target_data, logic_data) =>
+             (modl, gens,
+              target_data |> Symtab.map_entry target
+                (map_target_data
+                  (fn (syntax_tyco, syntax_const) =>
+                   (syntax_tyco,
+                    syntax_const
+                    |> Symtab.update_new
+                       (c, (pretty, stamp ()))))),
+              logic_data)))
+      end;
   in
-    thy
-    |> ensure_prim c
-    |> prep_mfx tabs raw_mfx
-    |-> (fn mfx => map_codegen_data
-      (fn (modl, gens, lookups, serialize_data, logic_data) =>
-         (modl, gens, lookups,
-          serialize_data |> check_for_serializer serial_name |> Symtab.map_entry serial_name
-            (map_serialize_data
-              (fn (syntax_tyco, syntax_const) =>
-               (syntax_tyco,
-                syntax_const |> Symtab.update_new
-                  (c,
-                    (((Codegen.num_args_of mfx, fixity), Codegen.fillin_mixfix mfx), stamp ()))))),
-          logic_data)))
+    CodegenSerializer.parse_syntax (read_quote Sign.read_term exprgen_term)
+    #-> (fn reader => pair (mk reader))
   end;
 
-val add_syntax_const_i = gen_add_syntax_const ((K o K) I) (K pair);
-val add_syntax_const = gen_add_syntax_const
-  (fn thy => fn tabs => idf_of_const thy tabs o read_const thy)
-  (gen_prep_mfx (fn thy => term_of o read_cterm thy o rpair TypeInfer.logicT)
-    (fn thy => fn tabs => exprgen_term thy tabs o devarify_term));
-
 
 
 (** code generation **)
 
-fun get_serializer thy serial_name =
-  (#serializer o (fn data => (the oo Symtab.lookup) data serial_name)
-    o #serialize_data o CodegenData.get) thy;
-
-fun mk_const thy (f, s_ty) =
+fun generate_code raw_consts thy =
   let
-    val f' = Sign.intern_const thy f;
-    val ty = case s_ty
-     of NONE => Sign.the_const_constraint thy f'
-      | SOME s => Sign.read_typ (thy, K NONE) s;
-  in (f', ty) end;
-
-fun generate_code consts thy =
-  let
-    val tabs = mk_tabs thy;
-    val consts' = map (mk_const thy) consts;
-    fun generate thy tabs = fold_map (ensure_def_const thy tabs) consts'
+    val consts = map (read_const thy) raw_consts;
+    fun generate thy tabs = fold_map (ensure_def_const thy tabs) consts
   in
     thy
-    |> expand_module tabs generate
-    |-> (fn consts => pair consts)
+    |> expand_module generate
   end;
 
-fun serialize_code serial_name filename consts thy =
+fun serialize_code target filename raw_consts thy =
   let
-    val serialize_data =
-      thy
-      |> CodegenData.get
-      |> #serialize_data
-      |> check_for_serializer serial_name
-      |> (fn data => (the oo Symtab.lookup) data serial_name)
-    val serializer' = (get_serializer thy serial_name) serial_name
-      ((Option.map fst oo Symtab.lookup o #syntax_tyco) serialize_data)
-      ((Option.map fst oo Symtab.lookup o #syntax_const) serialize_data);
-    val compile_it = serial_name = "ml" andalso filename = "-";
+    fun get_serializer thy target =
+      let
+        val _ = check_for_target thy target;
+        val target_data =
+          thy
+          |> CodegenData.get
+          |> #target_data
+          |> (fn data => (the oo Symtab.lookup) data target);
+      in
+        (the o Symtab.lookup (! serializers)) target (
+          (Option.map fst oo Symtab.lookup o #syntax_tyco) target_data,
+          (Option.map fst oo Symtab.lookup o #syntax_const) target_data
+        )
+      end;
     fun use_code code =
-      if compile_it
+      if target = "ml" andalso filename = "-"
       then use_text Context.ml_output false code
       else File.write (Path.unpack filename) (code ^ "\n");
+    fun serialize thy cs =
+      let
+        val module = (#modl o CodegenData.get) thy;
+        val seri = get_serializer thy target "Generated";
+      in case cs
+       of [] => seri NONE module () |> fst |> Pretty.output |> use_code
+        | cs => seri (SOME cs) module () |> fst |> Pretty.output |> use_code
+      end;
   in
     thy
-    |> (if is_some consts then generate_code (the consts) else pair [])
-    |-> (fn [] => `(serializer' NONE o #modl o CodegenData.get)
-          | consts => `(serializer' (SOME consts) o #modl o CodegenData.get))
-    |-> (fn code => ((use_code o Pretty.output) code; I))
+    |> (if is_some raw_consts then generate_code (the raw_consts) else pair [])
+    |-> (fn cs => `(fn thy => serialize thy cs))
+    |-> (fn _ => I)
   end;
 
 
@@ -1347,68 +1216,71 @@
          P.$$$ constantsK
          |-- Scan.repeat1 (P.name -- Scan.option (P.$$$ "::" |-- P.typ))
        )
-    >> (fn ((serial_name, filename), consts) =>
-          Toplevel.theory (serialize_code serial_name filename consts))
+    >> (fn ((target, filename), raw_consts) =>
+          Toplevel.theory (serialize_code target filename raw_consts))
   );
 
 val aliasP =
   OuterSyntax.command aliasK "declare an alias for a theory identifier" K.thy_decl (
-    P.xname
-    -- P.xname
-      >> (fn (src, dst) => Toplevel.theory (add_alias (src, dst)))
+    Scan.repeat1 (P.name -- P.name)
+      >> (Toplevel.theory oo fold) add_alias
   );
 
 val primclassP =
   OuterSyntax.command primclassK "define target-lanugage specific class" K.thy_decl (
     P.xname
+    -- Scan.optional (P.$$$ dependingK |--
+         P.$$$ "(" |-- P.list1 P.name --| P.$$$ ")") []
     -- Scan.repeat1 (P.name -- P.text)
-    -- Scan.optional (P.$$$ dependingK |-- Scan.repeat1 P.name) []
-      >> (fn ((raw_class, primdefs), depends) =>
+      >> (fn ((raw_class, depends), primdefs) =>
             (Toplevel.theory oo fold) (add_prim_class raw_class depends) primdefs)
   );
 
 val primtycoP =
   OuterSyntax.command primtycoK "define target-lanugage specific type constructor" K.thy_decl (
     P.xname
+    -- Scan.optional (P.$$$ dependingK |--
+         P.$$$ "(" |-- P.list1 P.name --| P.$$$ ")") []
     -- Scan.repeat1 (P.name -- P.text)
-    -- Scan.optional (P.$$$ dependingK |-- Scan.repeat1 P.name) []
-      >> (fn ((raw_tyco, primdefs), depends) =>
+      >> (fn ((raw_tyco, depends), primdefs) =>
             (Toplevel.theory oo fold) (add_prim_tyco raw_tyco depends) primdefs)
   );
 
 val primconstP =
   OuterSyntax.command primconstK "define target-lanugage specific constant" K.thy_decl (
-    (P.xname -- Scan.option P.typ)
+    (P.xname -- Scan.option (P.$$$ "::" |-- P.typ))
+    -- Scan.optional (P.$$$ dependingK |--
+         P.$$$ "(" |-- P.list1 P.name --| P.$$$ ")") []
     -- Scan.repeat1 (P.name -- P.text)
-    -- Scan.optional (P.$$$ dependingK |-- Scan.repeat1 P.name) []
-      >> (fn ((raw_const, primdefs), depends) =>
+      >> (fn ((raw_const, depends), primdefs) =>
             (Toplevel.theory oo fold) (add_prim_const raw_const depends) primdefs)
   );
 
+val _ : OuterParse.token list -> (string -> string -> theory -> theory) * OuterParse.token list
+ = parse_syntax_tyco;
+
 val syntax_tycoP =
   OuterSyntax.command syntax_tycoK "define code syntax for type constructor" K.thy_decl (
-    P.xname
-    -- Scan.repeat1 (
-         P.name -- (P.$$$ "(" |-- P.string --| P.$$$ ")")
-         -- CodegenSerializer.parse_fixity
-       )
-    >> (fn (raw_tyco, stxs) =>
-          (Toplevel.theory oo fold)
-            (fn ((target, raw_mfx), fixity) =>
-              add_syntax_tyco raw_tyco (target, (raw_mfx, fixity))) stxs)
+    Scan.repeat1 (
+      P.xname
+      -- Scan.repeat1 (
+           P.name -- parse_syntax_tyco
+         )
+    )
+    >> (Toplevel.theory oo fold) (fn (raw_tyco, syns) =>
+          fold (fn (target, modifier) => modifier raw_tyco target) syns)
   );
 
 val syntax_constP =
   OuterSyntax.command syntax_constK "define code syntax for constant" K.thy_decl (
-    (P.xname -- Scan.option (P.$$$ "::" |-- P.typ))
-    -- Scan.repeat1 (
-         P.name -- (P.$$$ "(" |-- P.string --| P.$$$ ")")
-         -- CodegenSerializer.parse_fixity
-       )
-    >> (fn (raw_c, stxs) =>
-          (Toplevel.theory oo fold)
-            (fn ((target, raw_mfx), fixity) =>
-              add_syntax_const raw_c (target, (raw_mfx, fixity))) stxs)
+    Scan.repeat1 (
+      (P.xname -- Scan.option (P.$$$ "::" |-- P.typ))
+      -- Scan.repeat1 (
+           P.name -- parse_syntax_const
+         )
+    )
+    >> (Toplevel.theory oo fold) (fn (raw_c, syns) =>
+          fold (fn (target, modifier) => modifier raw_c target) syns)
   );
 
 val _ = OuterSyntax.add_parsers [classP, generateP, serializeP, aliasP,
@@ -1419,67 +1291,16 @@
 
 (** setup **)
 
-val _ =
-  let
-    val bool = Type ("bool", []);
-    val nat = Type ("nat", []);
-    val int = Type ("IntDef.int", []);
-    fun list t = Type ("List.list", [t]);
-    fun pair t1 t2 = Type ("*", [t1, t2]);
-    val A = TVar (("'a", 0), []);
-    val B = TVar (("'b", 0), []);
-  in Context.add_setup [
-    CodegenData.init,
-    add_appconst_i ("neg", ((0, 0), appgen_neg)),
-    add_appconst_i ("op =", ((2, 2), appgen_eq)),
-    add_defgen ("clsdecl", defgen_clsdecl),
-    add_defgen ("tyco_fallback", defgen_tyco_fallback),
-    add_defgen ("const_fallback", defgen_const_fallback),
-    add_defgen ("defs", defgen_defs),
-    add_defgen ("clsmem", defgen_clsmem),
-    add_defgen ("clsinst", defgen_clsinst),
-    add_alias ("op -->", "HOL.op_implies"),
-    add_alias ("op +", "HOL.op_add"),
-    add_alias ("op -", "HOL.op_minus"),
-    add_alias ("op *", "HOL.op_times"),
-    add_alias ("op <=", "Orderings.op_le"),
-    add_alias ("op <", "Orderings.op_lt"),
-    add_alias ("List.op @", "List.append"),
-    add_alias ("List.op mem", "List.member"),
-    add_alias ("Divides.op div", "Divides.div"),
-    add_alias ("Divides.op dvd", "Divides.dvd"),
-    add_alias ("Divides.op mod", "Divides.mod"),
-    add_lookup_tyco ("bool", type_bool),
-    add_lookup_tyco ("*", type_pair),
-    add_lookup_tyco ("IntDef.int", type_integer),
-    add_lookup_tyco ("List.list", type_list),
-    add_lookup_const (("True", bool), Cons_true),
-    add_lookup_const (("False", bool), Cons_false),
-    add_lookup_const (("Not", bool --> bool), Fun_not),
-    add_lookup_const (("op &", bool --> bool --> bool), Fun_and),
-    add_lookup_const (("op |", bool --> bool --> bool), Fun_or),
-    add_lookup_const (("HOL.If", bool --> A --> A --> A), Fun_if),
-    add_lookup_const (("Pair", A --> B --> pair A B), Cons_pair),
-    add_lookup_const (("fst", pair A B --> A), Fun_fst),
-    add_lookup_const (("snd", pair A B --> B), Fun_snd),
-    add_lookup_const (("List.list.Cons", A --> list A --> list A), Cons_cons),
-    add_lookup_const (("List.list.Nil", list A), Cons_nil),
-    add_lookup_const (("1", nat),
-      IApp (
-        IConst ("const.Suc", IFun (IType ("type.nat", []), IFun (IType ("type.nat", []), IType ("type.nat", [])))),
-        IConst ("const.Zero", IType ("type.nat", []))
-      )),
-    add_lookup_const (("0", int), Fun_0),
-    add_lookup_const (("1", int), Fun_1),
-    add_lookup_const (("op +", int --> int --> int), Fun_add),
-    add_lookup_const (("op *", int --> int --> int), Fun_mult),
-    add_lookup_const (("uminus", int --> int), Fun_minus),
-    add_lookup_const (("op <", int --> int --> bool), Fun_lt),
-    add_lookup_const (("op <=", int --> int --> bool), Fun_le),
-    add_lookup_const (("Wellfounded_Recursion.wfrec", ((A --> B) --> A --> B) --> A --> B), Fun_wfrec)
-  ] end;
-
-(* "op /" ??? *)
+val _ = Context.add_setup [
+  CodegenData.init,
+(*   add_appconst_i ("op =", ((2, 2), appgen_eq)),  *)
+  add_eqextr ("defs", eqextr_defs),
+  add_defgen ("clsdecl", defgen_clsdecl),
+  add_defgen ("funs", defgen_funs),
+  add_defgen ("clsmem", defgen_clsmem),
+  add_defgen ("datatypecons", defgen_datatypecons)(*,
+  add_defgen ("clsinst", defgen_clsinst)  *)
+];
 
 end; (* local *)