renamed former datatype.ML to datatype_data.ML; datatype.ML provides uniform view on datatype.ML and datatype_rep_proofs.ML
--- a/src/HOL/Datatype.thy Fri Nov 27 08:41:08 2009 +0100
+++ b/src/HOL/Datatype.thy Fri Nov 27 08:41:10 2009 +0100
@@ -11,6 +11,7 @@
imports Product_Type Sum_Type Nat
uses
("Tools/Datatype/datatype_rep_proofs.ML")
+ ("Tools/Datatype/datatype.ML")
("Tools/inductive_realizer.ML")
("Tools/Datatype/datatype_realizer.ML")
begin
@@ -520,6 +521,7 @@
hide (open) const Push Node Atom Leaf Numb Lim Split Case
use "Tools/Datatype/datatype_rep_proofs.ML"
+use "Tools/Datatype/datatype.ML"
use "Tools/inductive_realizer.ML"
setup InductiveRealizer.setup
--- a/src/HOL/Inductive.thy Fri Nov 27 08:41:08 2009 +0100
+++ b/src/HOL/Inductive.thy Fri Nov 27 08:41:10 2009 +0100
@@ -14,7 +14,7 @@
"Tools/Datatype/datatype_prop.ML"
"Tools/Datatype/datatype_case.ML"
("Tools/Datatype/datatype_abs_proofs.ML")
- ("Tools/Datatype/datatype.ML")
+ ("Tools/Datatype/datatype_data.ML")
("Tools/old_primrec.ML")
("Tools/primrec.ML")
("Tools/Datatype/datatype_codegen.ML")
@@ -282,8 +282,8 @@
text {* Package setup. *}
use "Tools/Datatype/datatype_abs_proofs.ML"
-use "Tools/Datatype/datatype.ML"
-setup Datatype.setup
+use "Tools/Datatype/datatype_data.ML"
+setup Datatype_Data.setup
use "Tools/old_primrec.ML"
use "Tools/primrec.ML"
@@ -306,7 +306,7 @@
fun fun_tr ctxt [cs] =
let
val x = Free (Name.variant (Term.add_free_names cs []) "x", dummyT);
- val ft = DatatypeCase.case_tr true Datatype.info_of_constr
+ val ft = DatatypeCase.case_tr true Datatype_Data.info_of_constr
ctxt [x, cs]
in lambda x ft end
in [("_lam_pats_syntax", fun_tr)] end
--- a/src/HOL/IsaMakefile Fri Nov 27 08:41:08 2009 +0100
+++ b/src/HOL/IsaMakefile Fri Nov 27 08:41:10 2009 +0100
@@ -168,10 +168,11 @@
Tools/Datatype/datatype_aux.ML \
Tools/Datatype/datatype_case.ML \
Tools/Datatype/datatype_codegen.ML \
- Tools/Datatype/datatype.ML \
+ Tools/Datatype/datatype_data.ML \
Tools/Datatype/datatype_prop.ML \
Tools/Datatype/datatype_realizer.ML \
Tools/Datatype/datatype_rep_proofs.ML \
+ Tools/Datatype/datatype.ML \
Tools/dseq.ML \
Tools/Function/context_tree.ML \
Tools/Function/decompose.ML \
--- a/src/HOL/Tools/Datatype/datatype.ML Fri Nov 27 08:41:08 2009 +0100
+++ b/src/HOL/Tools/Datatype/datatype.ML Fri Nov 27 08:41:10 2009 +0100
@@ -1,471 +1,19 @@
(* Title: HOL/Tools/datatype.ML
Author: Stefan Berghofer, TU Muenchen
-Datatype package for Isabelle/HOL.
+Datatype package interface for Isabelle/HOL.
*)
signature DATATYPE =
sig
- include DATATYPE_COMMON
- val derive_datatype_props : config -> string list -> string list option
- -> descr list -> (string * sort) list -> thm -> thm list list -> thm list list
- -> theory -> string list * theory
- val rep_datatype : config -> (string list -> Proof.context -> Proof.context)
- -> string list option -> term list -> theory -> Proof.state
- val rep_datatype_cmd : string list option -> string list -> theory -> Proof.state
- val get_info : theory -> string -> info option
- val the_info : theory -> string -> info
- val the_descr : theory -> string list
- -> descr * (string * sort) list * string list
- * string * (string list * string list) * (typ list * typ list)
- val the_spec : theory -> string -> (string * sort) list * (string * typ list) list
- val all_distincts : theory -> typ list -> thm list list
- val get_constrs : theory -> string -> (string * typ) list option
- val get_all : theory -> info Symtab.table
- val info_of_constr : theory -> string * typ -> info option
- val info_of_case : theory -> string -> info option
- val interpretation : (config -> string list -> theory -> theory) -> theory -> theory
- val make_case : Proof.context -> DatatypeCase.config -> string list -> term ->
- (term * term) list -> term * (term * (int * bool)) list
- val strip_case : Proof.context -> bool -> term -> (term * (term * term) list) option
- val read_typ: theory -> string -> (string * sort) list -> typ * (string * sort) list
- val cert_typ: theory -> typ -> (string * sort) list -> typ * (string * sort) list
- val mk_case_names_induct: descr -> attribute
- val setup: theory -> theory
-end;
-
-structure Datatype : DATATYPE =
-struct
-
-open DatatypeAux;
-
-(** theory data **)
-
-(* data management *)
-
-structure DatatypesData = Theory_Data
-(
- type T =
- {types: info Symtab.table,
- constrs: (string * info) list Symtab.table,
- cases: info Symtab.table};
-
- val empty =
- {types = Symtab.empty, constrs = Symtab.empty, cases = Symtab.empty};
- val extend = I;
- fun merge
- ({types = types1, constrs = constrs1, cases = cases1},
- {types = types2, constrs = constrs2, cases = cases2}) : T =
- {types = Symtab.merge (K true) (types1, types2),
- constrs = Symtab.join (K (AList.merge (op =) (K true))) (constrs1, constrs2),
- cases = Symtab.merge (K true) (cases1, cases2)};
-);
-
-val get_all = #types o DatatypesData.get;
-val get_info = Symtab.lookup o get_all;
-
-fun the_info thy name =
- (case get_info thy name of
- SOME info => info
- | NONE => error ("Unknown datatype " ^ quote name));
-
-fun info_of_constr thy (c, T) =
- let
- val tab = Symtab.lookup_list ((#constrs o DatatypesData.get) thy) c;
- val hint = case strip_type T of (_, Type (tyco, _)) => SOME tyco | _ => NONE;
- val default = if null tab then NONE
- else SOME (snd (Library.last_elem tab))
- (*conservative wrt. overloaded constructors*);
- in case hint
- of NONE => default
- | SOME tyco => case AList.lookup (op =) tab tyco
- of NONE => default (*permissive*)
- | SOME info => SOME info
- end;
-
-val info_of_case = Symtab.lookup o #cases o DatatypesData.get;
-
-fun register (dt_infos : (string * info) list) =
- DatatypesData.map (fn {types, constrs, cases} =>
- {types = types |> fold Symtab.update dt_infos,
- constrs = constrs |> fold (fn (constr, dtname_info) =>
- Symtab.map_default (constr, []) (cons dtname_info))
- (maps (fn (dtname, info as {descr, index, ...}) =>
- map (rpair (dtname, info) o fst)
- (#3 (the (AList.lookup op = descr index)))) dt_infos),
- cases = cases |> fold Symtab.update
- (map (fn (_, info as {case_name, ...}) => (case_name, info)) dt_infos)});
-
-
-(* complex queries *)
-
-fun the_spec thy dtco =
- let
- val info as { descr, index, sorts = raw_sorts, ... } = the_info thy dtco;
- val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr index;
- val sorts = map ((fn v => (v, (the o AList.lookup (op =) raw_sorts) v))
- o DatatypeAux.dest_DtTFree) dtys;
- val cos = map
- (fn (co, tys) => (co, map (DatatypeAux.typ_of_dtyp descr sorts) tys)) raw_cos;
- in (sorts, cos) end;
-
-fun the_descr thy (raw_tycos as raw_tyco :: _) =
- let
- val info = the_info thy raw_tyco;
- val descr = #descr info;
-
- val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr (#index info);
- val vs = map ((fn v => (v, (the o AList.lookup (op =) (#sorts info)) v))
- o dest_DtTFree) dtys;
-
- fun is_DtTFree (DtTFree _) = true
- | is_DtTFree _ = false
- val k = find_index (fn (_, (_, dTs, _)) => not (forall is_DtTFree dTs)) descr;
- val protoTs as (dataTs, _) = chop k descr
- |> (pairself o map) (fn (_, (tyco, dTs, _)) => (tyco, map (typ_of_dtyp descr vs) dTs));
-
- val tycos = map fst dataTs;
- val _ = if eq_set (op =) (tycos, raw_tycos) then ()
- else error ("Type constructors " ^ commas (map quote raw_tycos)
- ^ " do not belong exhaustively to one mutual recursive datatype");
-
- val (Ts, Us) = (pairself o map) Type protoTs;
-
- val names = map Long_Name.base_name (the_default tycos (#alt_names info));
- val (auxnames, _) = Name.make_context names
- |> fold_map (yield_singleton Name.variants o name_of_typ) Us;
- val prefix = space_implode "_" names;
-
- in (descr, vs, tycos, prefix, (names, auxnames), (Ts, Us)) end;
-
-fun all_distincts thy Ts =
- let
- fun add_tycos (Type (tyco, Ts)) = insert (op =) tyco #> fold add_tycos Ts
- | add_tycos _ = I;
- val tycos = fold add_tycos Ts [];
- in map_filter (Option.map #distinct o get_info thy) tycos end;
-
-fun get_constrs thy dtco =
- case try (the_spec thy) dtco
- of SOME (sorts, cos) =>
- let
- fun subst (v, sort) = TVar ((v, 0), sort);
- fun subst_ty (TFree v) = subst v
- | subst_ty ty = ty;
- val dty = Type (dtco, map subst sorts);
- fun mk_co (co, tys) = (co, map (Term.map_atyps subst_ty) tys ---> dty);
- in SOME (map mk_co cos) end
- | NONE => NONE;
-
-
-
-(** various auxiliary **)
-
-(* prepare datatype specifications *)
-
-fun read_typ thy str sorts =
- let
- val ctxt = ProofContext.init thy
- |> fold (Variable.declare_typ o TFree) sorts;
- val T = Syntax.read_typ ctxt str;
- in (T, Term.add_tfreesT T sorts) end;
-
-fun cert_typ sign raw_T sorts =
- let
- val T = Type.no_tvars (Sign.certify_typ sign raw_T)
- handle TYPE (msg, _, _) => error msg;
- val sorts' = Term.add_tfreesT T sorts;
- val _ =
- case duplicates (op =) (map fst sorts') of
- [] => ()
- | dups => error ("Inconsistent sort constraints for " ^ commas dups)
- in (T, sorts') end;
-
-
-(* case names *)
-
-local
-
-fun dt_recs (DtTFree _) = []
- | dt_recs (DtType (_, dts)) = maps dt_recs dts
- | dt_recs (DtRec i) = [i];
-
-fun dt_cases (descr: descr) (_, args, constrs) =
- let
- fun the_bname i = Long_Name.base_name (#1 (the (AList.lookup (op =) descr i)));
- val bnames = map the_bname (distinct (op =) (maps dt_recs args));
- in map (fn (c, _) => space_implode "_" (Long_Name.base_name c :: bnames)) constrs end;
-
-fun induct_cases descr =
- DatatypeProp.indexify_names (maps (dt_cases descr) (map #2 descr));
-
-fun exhaust_cases descr i = dt_cases descr (the (AList.lookup (op =) descr i));
-
-in
-
-fun mk_case_names_induct descr = Rule_Cases.case_names (induct_cases descr);
-
-fun mk_case_names_exhausts descr new =
- map (Rule_Cases.case_names o exhaust_cases descr o #1)
- (filter (fn ((_, (name, _, _))) => member (op =) new name) descr);
-
+ include DATATYPE_DATA
+ include DATATYPE_REP_PROOFS
end;
-
-(* translation rules for case *)
-
-fun make_case ctxt = DatatypeCase.make_case
- (info_of_constr (ProofContext.theory_of ctxt)) ctxt;
-
-fun strip_case ctxt = DatatypeCase.strip_case
- (info_of_case (ProofContext.theory_of ctxt));
-
-fun add_case_tr' case_names thy =
- Sign.add_advanced_trfuns ([], [],
- map (fn case_name =>
- let val case_name' = Sign.const_syntax_name thy case_name
- in (case_name', DatatypeCase.case_tr' info_of_case case_name')
- end) case_names, []) thy;
-
-val trfun_setup =
- Sign.add_advanced_trfuns ([],
- [("_case_syntax", DatatypeCase.case_tr true info_of_constr)],
- [], []);
-
-
-
-(** document antiquotation **)
-
-val _ = ThyOutput.antiquotation "datatype" Args.tyname
- (fn {source = src, context = ctxt, ...} => fn dtco =>
- let
- val thy = ProofContext.theory_of ctxt;
- val (vs, cos) = the_spec thy dtco;
- val ty = Type (dtco, map TFree vs);
- fun pretty_typ_bracket (ty as Type (_, _ :: _)) =
- Pretty.enclose "(" ")" [Syntax.pretty_typ ctxt ty]
- | pretty_typ_bracket ty =
- Syntax.pretty_typ ctxt ty;
- fun pretty_constr (co, tys) =
- (Pretty.block o Pretty.breaks)
- (Syntax.pretty_term ctxt (Const (co, tys ---> ty)) ::
- map pretty_typ_bracket tys);
- val pretty_datatype =
- Pretty.block
- (Pretty.command "datatype" :: Pretty.brk 1 ::
- Syntax.pretty_typ ctxt ty ::
- Pretty.str " =" :: Pretty.brk 1 ::
- flat (separate [Pretty.brk 1, Pretty.str "| "]
- (map (single o pretty_constr) cos)));
- in ThyOutput.output (ThyOutput.maybe_pretty_source (K pretty_datatype) src [()]) end);
-
-
-
-(** abstract theory extensions relative to a datatype characterisation **)
-
-structure Datatype_Interpretation = Interpretation
- (type T = config * string list val eq: T * T -> bool = eq_snd op =);
-fun interpretation f = Datatype_Interpretation.interpretation (uncurry f);
-
-fun make_dt_info alt_names descr sorts induct inducts rec_names rec_rewrites
- (index, (((((((((((_, (tname, _, _))), inject), distinct),
- exhaust), nchotomy), case_name), case_rewrites), case_cong), weak_case_cong),
- (split, split_asm))) =
- (tname,
- {index = index,
- alt_names = alt_names,
- descr = descr,
- sorts = sorts,
- inject = inject,
- distinct = distinct,
- induct = induct,
- inducts = inducts,
- exhaust = exhaust,
- nchotomy = nchotomy,
- rec_names = rec_names,
- rec_rewrites = rec_rewrites,
- case_name = case_name,
- case_rewrites = case_rewrites,
- case_cong = case_cong,
- weak_case_cong = weak_case_cong,
- split = split,
- split_asm = split_asm});
-
-fun derive_datatype_props config dt_names alt_names descr sorts
- induct inject distinct thy1 =
- let
- val thy2 = thy1 |> Theory.checkpoint;
- val flat_descr = flat descr;
- val new_type_names = map Long_Name.base_name (the_default dt_names alt_names);
- val _ = message config ("Deriving properties for datatype(s) " ^ commas_quote new_type_names);
-
- val (exhaust, thy3) = DatatypeAbsProofs.prove_casedist_thms config new_type_names
- descr sorts induct (mk_case_names_exhausts flat_descr dt_names) thy2;
- val (nchotomys, thy4) = DatatypeAbsProofs.prove_nchotomys config new_type_names
- descr sorts exhaust thy3;
- val ((rec_names, rec_rewrites), thy5) = DatatypeAbsProofs.prove_primrec_thms
- config new_type_names descr sorts (#inject o the o Symtab.lookup (get_all thy4))
- inject (distinct, all_distincts thy2 (get_rec_types flat_descr sorts))
- induct thy4;
- val ((case_rewrites, case_names), thy6) = DatatypeAbsProofs.prove_case_thms
- config new_type_names descr sorts rec_names rec_rewrites thy5;
- val (case_congs, thy7) = DatatypeAbsProofs.prove_case_congs new_type_names
- descr sorts nchotomys case_rewrites thy6;
- val (weak_case_congs, thy8) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
- descr sorts thy7;
- val (splits, thy9) = DatatypeAbsProofs.prove_split_thms
- config new_type_names descr sorts inject distinct exhaust case_rewrites thy8;
+structure Datatype =
+struct
- val inducts = Project_Rule.projections (ProofContext.init thy2) induct;
- val dt_infos = map_index
- (make_dt_info alt_names flat_descr sorts induct inducts rec_names rec_rewrites)
- (hd descr ~~ inject ~~ distinct ~~ exhaust ~~ nchotomys ~~
- case_names ~~ case_rewrites ~~ case_congs ~~ weak_case_congs ~~ splits);
- val dt_names = map fst dt_infos;
- val prfx = Binding.qualify true (space_implode "_" new_type_names);
- val simps = flat (inject @ distinct @ case_rewrites) @ rec_rewrites;
- val named_rules = flat (map_index (fn (index, tname) =>
- [((Binding.empty, [nth inducts index]), [Induct.induct_type tname]),
- ((Binding.empty, [nth exhaust index]), [Induct.cases_type tname])]) dt_names);
- val unnamed_rules = map (fn induct =>
- ((Binding.empty, [induct]), [Rule_Cases.inner_rule, Induct.induct_type ""]))
- (drop (length dt_names) inducts);
- in
- thy9
- |> PureThy.add_thmss ([((prfx (Binding.name "simps"), simps), []),
- ((prfx (Binding.name "inducts"), inducts), []),
- ((prfx (Binding.name "splits"), maps (fn (x, y) => [x, y]) splits), []),
- ((Binding.empty, flat case_rewrites @ flat distinct @ rec_rewrites),
- [Simplifier.simp_add]),
- ((Binding.empty, rec_rewrites), [Code.add_default_eqn_attribute]),
- ((Binding.empty, flat inject), [iff_add]),
- ((Binding.empty, map (fn th => th RS notE) (flat distinct)),
- [Classical.safe_elim NONE]),
- ((Binding.empty, weak_case_congs), [Simplifier.attrib (op addcongs)])]
- @ named_rules @ unnamed_rules)
- |> snd
- |> add_case_tr' case_names
- |> register dt_infos
- |> Datatype_Interpretation.data (config, dt_names)
- |> pair dt_names
- end;
-
-
-
-(** declare existing type as datatype **)
-
-fun prove_rep_datatype config dt_names alt_names descr sorts
- raw_inject half_distinct raw_induct thy1 =
- let
- val raw_distinct = (map o maps) (fn thm => [thm, thm RS not_sym]) half_distinct;
- val new_type_names = map Long_Name.base_name (the_default dt_names alt_names);
- val (((inject, distinct), [induct]), thy2) =
- thy1
- |> store_thmss "inject" new_type_names raw_inject
- ||>> store_thmss "distinct" new_type_names raw_distinct
- ||> Sign.add_path (space_implode "_" new_type_names)
- ||>> PureThy.add_thms [((Binding.name "induct", raw_induct), [mk_case_names_induct descr])]
- ||> Sign.restore_naming thy1;
- in
- thy2
- |> derive_datatype_props config dt_names alt_names [descr] sorts
- induct inject distinct
- end;
-
-fun gen_rep_datatype prep_term config after_qed alt_names raw_ts thy =
- let
- fun constr_of_term (Const (c, T)) = (c, T)
- | constr_of_term t =
- error ("Not a constant: " ^ Syntax.string_of_term_global thy t);
- fun no_constr (c, T) = error ("Bad constructor: "
- ^ Sign.extern_const thy c ^ "::"
- ^ Syntax.string_of_typ_global thy T);
- fun type_of_constr (cT as (_, T)) =
- let
- val frees = OldTerm.typ_tfrees T;
- val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) T
- handle TYPE _ => no_constr cT
- val _ = if has_duplicates (eq_fst (op =)) vs then no_constr cT else ();
- val _ = if length frees <> length vs then no_constr cT else ();
- in (tyco, (vs, cT)) end;
-
- val raw_cs = AList.group (op =) (map (type_of_constr o constr_of_term o prep_term thy) raw_ts);
- val _ = case map_filter (fn (tyco, _) =>
- if Symtab.defined (get_all thy) tyco then SOME tyco else NONE) raw_cs
- of [] => ()
- | tycos => error ("Type(s) " ^ commas (map quote tycos)
- ^ " already represented inductivly");
- val raw_vss = maps (map (map snd o fst) o snd) raw_cs;
- val ms = case distinct (op =) (map length raw_vss)
- of [n] => 0 upto n - 1
- | _ => error ("Different types in given constructors");
- fun inter_sort m = map (fn xs => nth xs m) raw_vss
- |> Library.foldr1 (Sorts.inter_sort (Sign.classes_of thy))
- val sorts = map inter_sort ms;
- val vs = Name.names Name.context Name.aT sorts;
-
- fun norm_constr (raw_vs, (c, T)) = (c, map_atyps
- (TFree o (the o AList.lookup (op =) (map fst raw_vs ~~ vs)) o fst o dest_TFree) T);
-
- val cs = map (apsnd (map norm_constr)) raw_cs;
- val dtyps_of_typ = map (dtyp_of_typ (map (rpair (map fst vs) o fst) cs))
- o fst o strip_type;
- val dt_names = map fst cs;
-
- fun mk_spec (i, (tyco, constr)) = (i, (tyco,
- map (DtTFree o fst) vs,
- (map o apsnd) dtyps_of_typ constr))
- val descr = map_index mk_spec cs;
- val injs = DatatypeProp.make_injs [descr] vs;
- val half_distincts = map snd (DatatypeProp.make_distincts [descr] vs);
- val ind = DatatypeProp.make_ind [descr] vs;
- val rules = (map o map o map) Logic.close_form [[[ind]], injs, half_distincts];
-
- fun after_qed' raw_thms =
- let
- val [[[raw_induct]], raw_inject, half_distinct] =
- unflat rules (map Drule.zero_var_indexes_list raw_thms);
- (*FIXME somehow dubious*)
- in
- ProofContext.theory_result
- (prove_rep_datatype config dt_names alt_names descr vs
- raw_inject half_distinct raw_induct)
- #-> after_qed
- end;
- in
- thy
- |> ProofContext.init
- |> Proof.theorem_i NONE after_qed' ((map o map) (rpair []) (flat rules))
- end;
-
-val rep_datatype = gen_rep_datatype Sign.cert_term;
-val rep_datatype_cmd = gen_rep_datatype Syntax.read_term_global default_config (K I);
-
-
-
-(** package setup **)
-
-(* setup theory *)
-
-val setup =
- trfun_setup #>
- Datatype_Interpretation.init;
-
-
-(* outer syntax *)
-
-local
-
-structure P = OuterParse and K = OuterKeyword
-
-in
-
-val _ =
- OuterSyntax.command "rep_datatype" "represent existing types inductively" K.thy_goal
- (Scan.option (P.$$$ "(" |-- Scan.repeat1 P.name --| P.$$$ ")") -- Scan.repeat1 P.term
- >> (fn (alt_names, ts) => Toplevel.print
- o Toplevel.theory_to_proof (rep_datatype_cmd alt_names ts)));
+open Datatype_Data;
+open DatatypeRepProofs;
end;
-
-end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Datatype/datatype_data.ML Fri Nov 27 08:41:10 2009 +0100
@@ -0,0 +1,471 @@
+(* Title: HOL/Tools/datatype.ML
+ Author: Stefan Berghofer, TU Muenchen
+
+Datatype package for Isabelle/HOL.
+*)
+
+signature DATATYPE_DATA =
+sig
+ include DATATYPE_COMMON
+ val derive_datatype_props : config -> string list -> string list option
+ -> descr list -> (string * sort) list -> thm -> thm list list -> thm list list
+ -> theory -> string list * theory
+ val rep_datatype : config -> (string list -> Proof.context -> Proof.context)
+ -> string list option -> term list -> theory -> Proof.state
+ val rep_datatype_cmd : string list option -> string list -> theory -> Proof.state
+ val get_info : theory -> string -> info option
+ val the_info : theory -> string -> info
+ val the_descr : theory -> string list
+ -> descr * (string * sort) list * string list
+ * string * (string list * string list) * (typ list * typ list)
+ val the_spec : theory -> string -> (string * sort) list * (string * typ list) list
+ val all_distincts : theory -> typ list -> thm list list
+ val get_constrs : theory -> string -> (string * typ) list option
+ val get_all : theory -> info Symtab.table
+ val info_of_constr : theory -> string * typ -> info option
+ val info_of_case : theory -> string -> info option
+ val interpretation : (config -> string list -> theory -> theory) -> theory -> theory
+ val make_case : Proof.context -> DatatypeCase.config -> string list -> term ->
+ (term * term) list -> term * (term * (int * bool)) list
+ val strip_case : Proof.context -> bool -> term -> (term * (term * term) list) option
+ val read_typ: theory -> string -> (string * sort) list -> typ * (string * sort) list
+ val cert_typ: theory -> typ -> (string * sort) list -> typ * (string * sort) list
+ val mk_case_names_induct: descr -> attribute
+ val setup: theory -> theory
+end;
+
+structure Datatype_Data: DATATYPE_DATA =
+struct
+
+open DatatypeAux;
+
+(** theory data **)
+
+(* data management *)
+
+structure DatatypesData = Theory_Data
+(
+ type T =
+ {types: info Symtab.table,
+ constrs: (string * info) list Symtab.table,
+ cases: info Symtab.table};
+
+ val empty =
+ {types = Symtab.empty, constrs = Symtab.empty, cases = Symtab.empty};
+ val extend = I;
+ fun merge
+ ({types = types1, constrs = constrs1, cases = cases1},
+ {types = types2, constrs = constrs2, cases = cases2}) : T =
+ {types = Symtab.merge (K true) (types1, types2),
+ constrs = Symtab.join (K (AList.merge (op =) (K true))) (constrs1, constrs2),
+ cases = Symtab.merge (K true) (cases1, cases2)};
+);
+
+val get_all = #types o DatatypesData.get;
+val get_info = Symtab.lookup o get_all;
+
+fun the_info thy name =
+ (case get_info thy name of
+ SOME info => info
+ | NONE => error ("Unknown datatype " ^ quote name));
+
+fun info_of_constr thy (c, T) =
+ let
+ val tab = Symtab.lookup_list ((#constrs o DatatypesData.get) thy) c;
+ val hint = case strip_type T of (_, Type (tyco, _)) => SOME tyco | _ => NONE;
+ val default = if null tab then NONE
+ else SOME (snd (Library.last_elem tab))
+ (*conservative wrt. overloaded constructors*);
+ in case hint
+ of NONE => default
+ | SOME tyco => case AList.lookup (op =) tab tyco
+ of NONE => default (*permissive*)
+ | SOME info => SOME info
+ end;
+
+val info_of_case = Symtab.lookup o #cases o DatatypesData.get;
+
+fun register (dt_infos : (string * info) list) =
+ DatatypesData.map (fn {types, constrs, cases} =>
+ {types = types |> fold Symtab.update dt_infos,
+ constrs = constrs |> fold (fn (constr, dtname_info) =>
+ Symtab.map_default (constr, []) (cons dtname_info))
+ (maps (fn (dtname, info as {descr, index, ...}) =>
+ map (rpair (dtname, info) o fst)
+ (#3 (the (AList.lookup op = descr index)))) dt_infos),
+ cases = cases |> fold Symtab.update
+ (map (fn (_, info as {case_name, ...}) => (case_name, info)) dt_infos)});
+
+
+(* complex queries *)
+
+fun the_spec thy dtco =
+ let
+ val info as { descr, index, sorts = raw_sorts, ... } = the_info thy dtco;
+ val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr index;
+ val sorts = map ((fn v => (v, (the o AList.lookup (op =) raw_sorts) v))
+ o DatatypeAux.dest_DtTFree) dtys;
+ val cos = map
+ (fn (co, tys) => (co, map (DatatypeAux.typ_of_dtyp descr sorts) tys)) raw_cos;
+ in (sorts, cos) end;
+
+fun the_descr thy (raw_tycos as raw_tyco :: _) =
+ let
+ val info = the_info thy raw_tyco;
+ val descr = #descr info;
+
+ val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr (#index info);
+ val vs = map ((fn v => (v, (the o AList.lookup (op =) (#sorts info)) v))
+ o dest_DtTFree) dtys;
+
+ fun is_DtTFree (DtTFree _) = true
+ | is_DtTFree _ = false
+ val k = find_index (fn (_, (_, dTs, _)) => not (forall is_DtTFree dTs)) descr;
+ val protoTs as (dataTs, _) = chop k descr
+ |> (pairself o map) (fn (_, (tyco, dTs, _)) => (tyco, map (typ_of_dtyp descr vs) dTs));
+
+ val tycos = map fst dataTs;
+ val _ = if eq_set (op =) (tycos, raw_tycos) then ()
+ else error ("Type constructors " ^ commas (map quote raw_tycos)
+ ^ " do not belong exhaustively to one mutual recursive datatype");
+
+ val (Ts, Us) = (pairself o map) Type protoTs;
+
+ val names = map Long_Name.base_name (the_default tycos (#alt_names info));
+ val (auxnames, _) = Name.make_context names
+ |> fold_map (yield_singleton Name.variants o name_of_typ) Us;
+ val prefix = space_implode "_" names;
+
+ in (descr, vs, tycos, prefix, (names, auxnames), (Ts, Us)) end;
+
+fun all_distincts thy Ts =
+ let
+ fun add_tycos (Type (tyco, Ts)) = insert (op =) tyco #> fold add_tycos Ts
+ | add_tycos _ = I;
+ val tycos = fold add_tycos Ts [];
+ in map_filter (Option.map #distinct o get_info thy) tycos end;
+
+fun get_constrs thy dtco =
+ case try (the_spec thy) dtco
+ of SOME (sorts, cos) =>
+ let
+ fun subst (v, sort) = TVar ((v, 0), sort);
+ fun subst_ty (TFree v) = subst v
+ | subst_ty ty = ty;
+ val dty = Type (dtco, map subst sorts);
+ fun mk_co (co, tys) = (co, map (Term.map_atyps subst_ty) tys ---> dty);
+ in SOME (map mk_co cos) end
+ | NONE => NONE;
+
+
+
+(** various auxiliary **)
+
+(* prepare datatype specifications *)
+
+fun read_typ thy str sorts =
+ let
+ val ctxt = ProofContext.init thy
+ |> fold (Variable.declare_typ o TFree) sorts;
+ val T = Syntax.read_typ ctxt str;
+ in (T, Term.add_tfreesT T sorts) end;
+
+fun cert_typ sign raw_T sorts =
+ let
+ val T = Type.no_tvars (Sign.certify_typ sign raw_T)
+ handle TYPE (msg, _, _) => error msg;
+ val sorts' = Term.add_tfreesT T sorts;
+ val _ =
+ case duplicates (op =) (map fst sorts') of
+ [] => ()
+ | dups => error ("Inconsistent sort constraints for " ^ commas dups)
+ in (T, sorts') end;
+
+
+(* case names *)
+
+local
+
+fun dt_recs (DtTFree _) = []
+ | dt_recs (DtType (_, dts)) = maps dt_recs dts
+ | dt_recs (DtRec i) = [i];
+
+fun dt_cases (descr: descr) (_, args, constrs) =
+ let
+ fun the_bname i = Long_Name.base_name (#1 (the (AList.lookup (op =) descr i)));
+ val bnames = map the_bname (distinct (op =) (maps dt_recs args));
+ in map (fn (c, _) => space_implode "_" (Long_Name.base_name c :: bnames)) constrs end;
+
+fun induct_cases descr =
+ DatatypeProp.indexify_names (maps (dt_cases descr) (map #2 descr));
+
+fun exhaust_cases descr i = dt_cases descr (the (AList.lookup (op =) descr i));
+
+in
+
+fun mk_case_names_induct descr = Rule_Cases.case_names (induct_cases descr);
+
+fun mk_case_names_exhausts descr new =
+ map (Rule_Cases.case_names o exhaust_cases descr o #1)
+ (filter (fn ((_, (name, _, _))) => member (op =) new name) descr);
+
+end;
+
+
+(* translation rules for case *)
+
+fun make_case ctxt = DatatypeCase.make_case
+ (info_of_constr (ProofContext.theory_of ctxt)) ctxt;
+
+fun strip_case ctxt = DatatypeCase.strip_case
+ (info_of_case (ProofContext.theory_of ctxt));
+
+fun add_case_tr' case_names thy =
+ Sign.add_advanced_trfuns ([], [],
+ map (fn case_name =>
+ let val case_name' = Sign.const_syntax_name thy case_name
+ in (case_name', DatatypeCase.case_tr' info_of_case case_name')
+ end) case_names, []) thy;
+
+val trfun_setup =
+ Sign.add_advanced_trfuns ([],
+ [("_case_syntax", DatatypeCase.case_tr true info_of_constr)],
+ [], []);
+
+
+
+(** document antiquotation **)
+
+val _ = ThyOutput.antiquotation "datatype" Args.tyname
+ (fn {source = src, context = ctxt, ...} => fn dtco =>
+ let
+ val thy = ProofContext.theory_of ctxt;
+ val (vs, cos) = the_spec thy dtco;
+ val ty = Type (dtco, map TFree vs);
+ fun pretty_typ_bracket (ty as Type (_, _ :: _)) =
+ Pretty.enclose "(" ")" [Syntax.pretty_typ ctxt ty]
+ | pretty_typ_bracket ty =
+ Syntax.pretty_typ ctxt ty;
+ fun pretty_constr (co, tys) =
+ (Pretty.block o Pretty.breaks)
+ (Syntax.pretty_term ctxt (Const (co, tys ---> ty)) ::
+ map pretty_typ_bracket tys);
+ val pretty_datatype =
+ Pretty.block
+ (Pretty.command "datatype" :: Pretty.brk 1 ::
+ Syntax.pretty_typ ctxt ty ::
+ Pretty.str " =" :: Pretty.brk 1 ::
+ flat (separate [Pretty.brk 1, Pretty.str "| "]
+ (map (single o pretty_constr) cos)));
+ in ThyOutput.output (ThyOutput.maybe_pretty_source (K pretty_datatype) src [()]) end);
+
+
+
+(** abstract theory extensions relative to a datatype characterisation **)
+
+structure Datatype_Interpretation = Interpretation
+ (type T = config * string list val eq: T * T -> bool = eq_snd op =);
+fun interpretation f = Datatype_Interpretation.interpretation (uncurry f);
+
+fun make_dt_info alt_names descr sorts induct inducts rec_names rec_rewrites
+ (index, (((((((((((_, (tname, _, _))), inject), distinct),
+ exhaust), nchotomy), case_name), case_rewrites), case_cong), weak_case_cong),
+ (split, split_asm))) =
+ (tname,
+ {index = index,
+ alt_names = alt_names,
+ descr = descr,
+ sorts = sorts,
+ inject = inject,
+ distinct = distinct,
+ induct = induct,
+ inducts = inducts,
+ exhaust = exhaust,
+ nchotomy = nchotomy,
+ rec_names = rec_names,
+ rec_rewrites = rec_rewrites,
+ case_name = case_name,
+ case_rewrites = case_rewrites,
+ case_cong = case_cong,
+ weak_case_cong = weak_case_cong,
+ split = split,
+ split_asm = split_asm});
+
+fun derive_datatype_props config dt_names alt_names descr sorts
+ induct inject distinct thy1 =
+ let
+ val thy2 = thy1 |> Theory.checkpoint;
+ val flat_descr = flat descr;
+ val new_type_names = map Long_Name.base_name (the_default dt_names alt_names);
+ val _ = message config ("Deriving properties for datatype(s) " ^ commas_quote new_type_names);
+
+ val (exhaust, thy3) = DatatypeAbsProofs.prove_casedist_thms config new_type_names
+ descr sorts induct (mk_case_names_exhausts flat_descr dt_names) thy2;
+ val (nchotomys, thy4) = DatatypeAbsProofs.prove_nchotomys config new_type_names
+ descr sorts exhaust thy3;
+ val ((rec_names, rec_rewrites), thy5) = DatatypeAbsProofs.prove_primrec_thms
+ config new_type_names descr sorts (#inject o the o Symtab.lookup (get_all thy4))
+ inject (distinct, all_distincts thy2 (get_rec_types flat_descr sorts))
+ induct thy4;
+ val ((case_rewrites, case_names), thy6) = DatatypeAbsProofs.prove_case_thms
+ config new_type_names descr sorts rec_names rec_rewrites thy5;
+ val (case_congs, thy7) = DatatypeAbsProofs.prove_case_congs new_type_names
+ descr sorts nchotomys case_rewrites thy6;
+ val (weak_case_congs, thy8) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
+ descr sorts thy7;
+ val (splits, thy9) = DatatypeAbsProofs.prove_split_thms
+ config new_type_names descr sorts inject distinct exhaust case_rewrites thy8;
+
+ val inducts = Project_Rule.projections (ProofContext.init thy2) induct;
+ val dt_infos = map_index
+ (make_dt_info alt_names flat_descr sorts induct inducts rec_names rec_rewrites)
+ (hd descr ~~ inject ~~ distinct ~~ exhaust ~~ nchotomys ~~
+ case_names ~~ case_rewrites ~~ case_congs ~~ weak_case_congs ~~ splits);
+ val dt_names = map fst dt_infos;
+ val prfx = Binding.qualify true (space_implode "_" new_type_names);
+ val simps = flat (inject @ distinct @ case_rewrites) @ rec_rewrites;
+ val named_rules = flat (map_index (fn (index, tname) =>
+ [((Binding.empty, [nth inducts index]), [Induct.induct_type tname]),
+ ((Binding.empty, [nth exhaust index]), [Induct.cases_type tname])]) dt_names);
+ val unnamed_rules = map (fn induct =>
+ ((Binding.empty, [induct]), [Rule_Cases.inner_rule, Induct.induct_type ""]))
+ (drop (length dt_names) inducts);
+ in
+ thy9
+ |> PureThy.add_thmss ([((prfx (Binding.name "simps"), simps), []),
+ ((prfx (Binding.name "inducts"), inducts), []),
+ ((prfx (Binding.name "splits"), maps (fn (x, y) => [x, y]) splits), []),
+ ((Binding.empty, flat case_rewrites @ flat distinct @ rec_rewrites),
+ [Simplifier.simp_add]),
+ ((Binding.empty, rec_rewrites), [Code.add_default_eqn_attribute]),
+ ((Binding.empty, flat inject), [iff_add]),
+ ((Binding.empty, map (fn th => th RS notE) (flat distinct)),
+ [Classical.safe_elim NONE]),
+ ((Binding.empty, weak_case_congs), [Simplifier.attrib (op addcongs)])]
+ @ named_rules @ unnamed_rules)
+ |> snd
+ |> add_case_tr' case_names
+ |> register dt_infos
+ |> Datatype_Interpretation.data (config, dt_names)
+ |> pair dt_names
+ end;
+
+
+
+(** declare existing type as datatype **)
+
+fun prove_rep_datatype config dt_names alt_names descr sorts
+ raw_inject half_distinct raw_induct thy1 =
+ let
+ val raw_distinct = (map o maps) (fn thm => [thm, thm RS not_sym]) half_distinct;
+ val new_type_names = map Long_Name.base_name (the_default dt_names alt_names);
+ val (((inject, distinct), [induct]), thy2) =
+ thy1
+ |> store_thmss "inject" new_type_names raw_inject
+ ||>> store_thmss "distinct" new_type_names raw_distinct
+ ||> Sign.add_path (space_implode "_" new_type_names)
+ ||>> PureThy.add_thms [((Binding.name "induct", raw_induct), [mk_case_names_induct descr])]
+ ||> Sign.restore_naming thy1;
+ in
+ thy2
+ |> derive_datatype_props config dt_names alt_names [descr] sorts
+ induct inject distinct
+ end;
+
+fun gen_rep_datatype prep_term config after_qed alt_names raw_ts thy =
+ let
+ fun constr_of_term (Const (c, T)) = (c, T)
+ | constr_of_term t =
+ error ("Not a constant: " ^ Syntax.string_of_term_global thy t);
+ fun no_constr (c, T) = error ("Bad constructor: "
+ ^ Sign.extern_const thy c ^ "::"
+ ^ Syntax.string_of_typ_global thy T);
+ fun type_of_constr (cT as (_, T)) =
+ let
+ val frees = OldTerm.typ_tfrees T;
+ val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) T
+ handle TYPE _ => no_constr cT
+ val _ = if has_duplicates (eq_fst (op =)) vs then no_constr cT else ();
+ val _ = if length frees <> length vs then no_constr cT else ();
+ in (tyco, (vs, cT)) end;
+
+ val raw_cs = AList.group (op =) (map (type_of_constr o constr_of_term o prep_term thy) raw_ts);
+ val _ = case map_filter (fn (tyco, _) =>
+ if Symtab.defined (get_all thy) tyco then SOME tyco else NONE) raw_cs
+ of [] => ()
+ | tycos => error ("Type(s) " ^ commas (map quote tycos)
+ ^ " already represented inductivly");
+ val raw_vss = maps (map (map snd o fst) o snd) raw_cs;
+ val ms = case distinct (op =) (map length raw_vss)
+ of [n] => 0 upto n - 1
+ | _ => error ("Different types in given constructors");
+ fun inter_sort m = map (fn xs => nth xs m) raw_vss
+ |> Library.foldr1 (Sorts.inter_sort (Sign.classes_of thy))
+ val sorts = map inter_sort ms;
+ val vs = Name.names Name.context Name.aT sorts;
+
+ fun norm_constr (raw_vs, (c, T)) = (c, map_atyps
+ (TFree o (the o AList.lookup (op =) (map fst raw_vs ~~ vs)) o fst o dest_TFree) T);
+
+ val cs = map (apsnd (map norm_constr)) raw_cs;
+ val dtyps_of_typ = map (dtyp_of_typ (map (rpair (map fst vs) o fst) cs))
+ o fst o strip_type;
+ val dt_names = map fst cs;
+
+ fun mk_spec (i, (tyco, constr)) = (i, (tyco,
+ map (DtTFree o fst) vs,
+ (map o apsnd) dtyps_of_typ constr))
+ val descr = map_index mk_spec cs;
+ val injs = DatatypeProp.make_injs [descr] vs;
+ val half_distincts = map snd (DatatypeProp.make_distincts [descr] vs);
+ val ind = DatatypeProp.make_ind [descr] vs;
+ val rules = (map o map o map) Logic.close_form [[[ind]], injs, half_distincts];
+
+ fun after_qed' raw_thms =
+ let
+ val [[[raw_induct]], raw_inject, half_distinct] =
+ unflat rules (map Drule.zero_var_indexes_list raw_thms);
+ (*FIXME somehow dubious*)
+ in
+ ProofContext.theory_result
+ (prove_rep_datatype config dt_names alt_names descr vs
+ raw_inject half_distinct raw_induct)
+ #-> after_qed
+ end;
+ in
+ thy
+ |> ProofContext.init
+ |> Proof.theorem_i NONE after_qed' ((map o map) (rpair []) (flat rules))
+ end;
+
+val rep_datatype = gen_rep_datatype Sign.cert_term;
+val rep_datatype_cmd = gen_rep_datatype Syntax.read_term_global default_config (K I);
+
+
+
+(** package setup **)
+
+(* setup theory *)
+
+val setup =
+ trfun_setup #>
+ Datatype_Interpretation.init;
+
+
+(* outer syntax *)
+
+local
+
+structure P = OuterParse and K = OuterKeyword
+
+in
+
+val _ =
+ OuterSyntax.command "rep_datatype" "represent existing types inductively" K.thy_goal
+ (Scan.option (P.$$$ "(" |-- Scan.repeat1 P.name --| P.$$$ ")") -- Scan.repeat1 P.term
+ >> (fn (alt_names, ts) => Toplevel.print
+ o Toplevel.theory_to_proof (rep_datatype_cmd alt_names ts)));
+
+end;
+
+end;
--- a/src/HOL/Tools/Datatype/datatype_rep_proofs.ML Fri Nov 27 08:41:08 2009 +0100
+++ b/src/HOL/Tools/Datatype/datatype_rep_proofs.ML Fri Nov 27 08:41:10 2009 +0100
@@ -10,8 +10,7 @@
signature DATATYPE_REP_PROOFS =
sig
- include DATATYPE_COMMON
- val add_datatype : config -> string list -> (string list * binding * mixfix *
+ val add_datatype : DatatypeAux.config -> string list -> (string list * binding * mixfix *
(binding * typ list * mixfix) list) list -> theory -> string list * theory
val datatype_cmd : string list -> (string list * binding * mixfix *
(binding * string list * mixfix) list) list -> theory -> theory
@@ -50,6 +49,8 @@
val In0_not_In1 = @{thm In0_not_In1};
val In1_not_In0 = @{thm In1_not_In0};
val Lim_inject = @{thm Lim_inject};
+val Inl_inject = @{thm Inl_inject};
+val Inr_inject = @{thm Inr_inject};
val Suml_inject = @{thm Suml_inject};
val Sumr_inject = @{thm Sumr_inject};
@@ -697,7 +698,7 @@
val (dts', constr_syntax, sorts', i) =
fold2 prep_dt_spec dts new_type_names ([], [], [], 0);
val sorts = sorts' @ map (rpair (Sign.defaultS tmp_thy)) (subtract (op =) (map fst sorts') tyvars);
- val dt_info = Datatype.get_all thy;
+ val dt_info = Datatype_Data.get_all thy;
val (descr, _) = unfold_datatypes tmp_thy dts' sorts dt_info dts' i;
val _ = check_nonempty descr handle (exn as Datatype_Empty s) =>
if #strict config then error ("Nonemptiness check failed for datatype " ^ s)
@@ -708,14 +709,14 @@
in
thy
|> representation_proofs config dt_info new_type_names descr sorts
- types_syntax constr_syntax (Datatype.mk_case_names_induct (flat descr))
- |-> (fn (inject, distinct, induct) => Datatype.derive_datatype_props
+ types_syntax constr_syntax (Datatype_Data.mk_case_names_induct (flat descr))
+ |-> (fn (inject, distinct, induct) => Datatype_Data.derive_datatype_props
config dt_names (SOME new_type_names) descr sorts
induct inject distinct)
end;
-val add_datatype = gen_add_datatype Datatype.cert_typ;
-val datatype_cmd = snd ooo gen_add_datatype Datatype.read_typ default_config;
+val add_datatype = gen_add_datatype Datatype_Data.cert_typ;
+val datatype_cmd = snd ooo gen_add_datatype Datatype_Data.read_typ default_config;
local
@@ -744,11 +745,3 @@
end;
end;
-
-structure Datatype =
-struct
-
-open Datatype;
-open DatatypeRepProofs;
-
-end;
--- a/src/HOL/Tools/inductive_codegen.ML Fri Nov 27 08:41:08 2009 +0100
+++ b/src/HOL/Tools/inductive_codegen.ML Fri Nov 27 08:41:10 2009 +0100
@@ -104,7 +104,7 @@
let
val cnstrs = flat (maps
(map (fn (_, (_, _, cs)) => map (apsnd length) cs) o #descr o snd)
- (Symtab.dest (Datatype.get_all thy)));
+ (Symtab.dest (Datatype_Data.get_all thy)));
fun check t = (case strip_comb t of
(Var _, []) => true
| (Const (s, _), ts) => (case AList.lookup (op =) cnstrs s of
--- a/src/HOL/Tools/old_primrec.ML Fri Nov 27 08:41:08 2009 +0100
+++ b/src/HOL/Tools/old_primrec.ML Fri Nov 27 08:41:10 2009 +0100
@@ -246,7 +246,7 @@
fun gen_primrec_i note def alt_name eqns_atts thy =
let
val (eqns, atts) = split_list eqns_atts;
- val dt_info = Datatype.get_all thy;
+ val dt_info = Datatype_Data.get_all thy;
val rec_eqns = fold_rev (process_eqn thy o snd) eqns [] ;
val tnames = distinct (op =) (map (#1 o snd) rec_eqns);
val dts = find_dts dt_info tnames tnames;
--- a/src/HOL/Tools/primrec.ML Fri Nov 27 08:41:08 2009 +0100
+++ b/src/HOL/Tools/primrec.ML Fri Nov 27 08:41:10 2009 +0100
@@ -220,7 +220,7 @@
val eqns = fold_rev (process_eqn (fn v => Variable.is_fixed lthy v
orelse exists (fn ((w, _), _) => v = Binding.name_of w) fixes)) eqs [];
val tnames = distinct (op =) (map (#1 o snd) eqns);
- val dts = find_dts (Datatype.get_all (ProofContext.theory_of lthy)) tnames tnames;
+ val dts = find_dts (Datatype_Data.get_all (ProofContext.theory_of lthy)) tnames tnames;
val main_fns = map (fn (tname, {index, ...}) =>
(index, (fst o the o find_first (fn (_, x) => #1 x = tname)) eqns)) dts;
val {descr, rec_names, rec_rewrites, ...} =