--- a/src/HOL/HOLCF/Tools/fixrec.ML Fri Dec 16 12:00:59 2011 +0100
+++ b/src/HOL/HOLCF/Tools/fixrec.ML Fri Dec 16 12:01:10 2011 +0100
@@ -326,7 +326,7 @@
(*************************************************************************)
local
-(* code adapted from HOL/Tools/primrec.ML *)
+(* code adapted from HOL/Tools/Datatype/primrec.ML *)
fun gen_fixrec
prep_spec
--- a/src/HOL/Inductive.thy Fri Dec 16 12:00:59 2011 +0100
+++ b/src/HOL/Inductive.thy Fri Dec 16 12:01:10 2011 +0100
@@ -7,16 +7,16 @@
theory Inductive
imports Complete_Lattices
uses
+ "Tools/dseq.ML"
("Tools/inductive.ML")
- "Tools/dseq.ML"
- "Tools/Datatype/datatype_aux.ML"
- "Tools/Datatype/datatype_prop.ML"
+ ("Tools/Datatype/datatype_aux.ML")
+ ("Tools/Datatype/datatype_prop.ML")
("Tools/Datatype/datatype_abs_proofs.ML")
("Tools/Datatype/datatype_data.ML")
("Tools/Datatype/datatype_case.ML")
("Tools/Datatype/rep_datatype.ML")
- ("Tools/primrec.ML")
("Tools/Datatype/datatype_codegen.ML")
+ ("Tools/Datatype/primrec.ML")
begin
subsection {* Least and greatest fixed points *}
@@ -276,15 +276,14 @@
text {* Package setup. *}
+use "Tools/Datatype/datatype_aux.ML"
+use "Tools/Datatype/datatype_prop.ML"
use "Tools/Datatype/datatype_abs_proofs.ML"
use "Tools/Datatype/datatype_data.ML" setup Datatype_Data.setup
use "Tools/Datatype/datatype_case.ML" setup Datatype_Case.setup
use "Tools/Datatype/rep_datatype.ML"
-
-use "Tools/Datatype/datatype_codegen.ML"
-setup Datatype_Codegen.setup
-
-use "Tools/primrec.ML"
+use "Tools/Datatype/datatype_codegen.ML" setup Datatype_Codegen.setup
+use "Tools/Datatype/primrec.ML"
text{* Lambda-abstractions with pattern matching: *}
--- a/src/HOL/IsaMakefile Fri Dec 16 12:00:59 2011 +0100
+++ b/src/HOL/IsaMakefile Fri Dec 16 12:01:10 2011 +0100
@@ -218,6 +218,7 @@
Tools/Datatype/datatype_data.ML \
Tools/Datatype/datatype_prop.ML \
Tools/Datatype/datatype_realizer.ML \
+ Tools/Datatype/primrec.ML \
Tools/Datatype/rep_datatype.ML \
Tools/Function/context_tree.ML \
Tools/Function/fun.ML \
@@ -256,7 +257,6 @@
Tools/lin_arith.ML \
Tools/monomorph.ML \
Tools/nat_arith.ML \
- Tools/primrec.ML \
Tools/prop_logic.ML \
Tools/refute.ML \
Tools/rewrite_hol_proof.ML \
--- a/src/HOL/Nominal/nominal_datatype.ML Fri Dec 16 12:00:59 2011 +0100
+++ b/src/HOL/Nominal/nominal_datatype.ML Fri Dec 16 12:01:10 2011 +0100
@@ -42,8 +42,8 @@
(* theory data *)
type descr =
- (int * (string * Datatype_Aux.dtyp list *
- (string * (Datatype_Aux.dtyp list * Datatype_Aux.dtyp) list) list)) list;
+ (int * (string * Datatype.dtyp list *
+ (string * (Datatype.dtyp list * Datatype.dtyp) list) list)) list;
type nominal_datatype_info =
{index : int,
@@ -200,7 +200,7 @@
val (full_new_type_names',thy1) = Datatype.add_datatype config dts'' thy;
val {descr, induct, ...} = Datatype.the_info thy1 (hd full_new_type_names');
- fun nth_dtyp i = Datatype_Aux.typ_of_dtyp descr (Datatype_Aux.DtRec i);
+ fun nth_dtyp i = Datatype_Aux.typ_of_dtyp descr (Datatype.DtRec i);
val big_name = space_implode "_" new_type_names;
@@ -465,13 +465,13 @@
| (i, _) => SOME (Datatype_Aux.name_of_typ (nth_dtyp i))) descr)) ^ "_set";
val _ = warning ("big_rep_name: " ^ big_rep_name);
- fun strip_option (dtf as Datatype_Aux.DtType ("fun", [dt, Datatype_Aux.DtRec i])) =
+ fun strip_option (dtf as Datatype.DtType ("fun", [dt, Datatype.DtRec i])) =
(case AList.lookup op = descr i of
SOME ("Nominal.noption", _, [(_, [dt']), _]) =>
apfst (cons dt) (strip_option dt')
| _ => ([], dtf))
- | strip_option (Datatype_Aux.DtType ("fun",
- [dt, Datatype_Aux.DtType ("Nominal.noption", [dt'])])) =
+ | strip_option (Datatype.DtType ("fun",
+ [dt, Datatype.DtType ("Nominal.noption", [dt'])])) =
apfst (cons dt) (strip_option dt')
| strip_option dt = ([], dt);
@@ -497,7 +497,7 @@
end
in (j + 1, j' + length Ts,
case dt'' of
- Datatype_Aux.DtRec k => list_all (map (pair "x") Us,
+ Datatype.DtRec k => list_all (map (pair "x") Us,
HOLogic.mk_Trueprop (Free (nth rep_set_names k,
T --> HOLogic.boolT) $ free')) :: prems
| _ => prems,
@@ -676,8 +676,8 @@
(fn ((i, ("Nominal.noption", _, _)), p) => p
| ((i, _), (ty_idxs, j)) => (ty_idxs @ [(i, j)], j + 1)) ([], 0) descr;
- fun reindex (Datatype_Aux.DtType (s, dts)) = Datatype_Aux.DtType (s, map reindex dts)
- | reindex (Datatype_Aux.DtRec i) = Datatype_Aux.DtRec (the (AList.lookup op = ty_idxs i))
+ fun reindex (Datatype.DtType (s, dts)) = Datatype.DtType (s, map reindex dts)
+ | reindex (Datatype.DtRec i) = Datatype.DtRec (the (AList.lookup op = ty_idxs i))
| reindex dt = dt;
fun strip_suffix i s = implode (List.take (raw_explode s, size s - i)); (* FIXME Symbol.explode (?) *)
@@ -713,7 +713,7 @@
map (fn ((cname, cargs), idxs) => (cname, partition_cargs idxs cargs))
(constrs ~~ idxss)))) (descr'' ~~ ndescr);
- fun nth_dtyp' i = Datatype_Aux.typ_of_dtyp descr'' (Datatype_Aux.DtRec i);
+ fun nth_dtyp' i = Datatype_Aux.typ_of_dtyp descr'' (Datatype.DtRec i);
val rep_names = map (fn s =>
Sign.intern_const thy7 ("Rep_" ^ s)) new_type_names;
@@ -740,7 +740,7 @@
xs @ x :: l_args,
fold_rev mk_abs_fun xs
(case dt of
- Datatype_Aux.DtRec k => if k < length new_type_names then
+ Datatype.DtRec k => if k < length new_type_names then
Const (nth rep_names k, Datatype_Aux.typ_of_dtyp descr'' dt -->
Datatype_Aux.typ_of_dtyp descr dt) $ x
else error "nested recursion not (yet) supported"
@@ -1435,7 +1435,7 @@
val binders = maps fst frees';
val atomTs = distinct op = (maps (map snd o fst) frees');
val recs = map_filter
- (fn ((_, Datatype_Aux.DtRec i), p) => SOME (i, p) | _ => NONE)
+ (fn ((_, Datatype.DtRec i), p) => SOME (i, p) | _ => NONE)
(partition_cargs idxs cargs ~~ frees');
val frees'' = map (fn i => "y" ^ string_of_int i) (1 upto length recs) ~~
map (fn (i, _) => nth rec_result_Ts i) recs;
--- a/src/HOL/Tools/ATP/atp_util.ML Fri Dec 16 12:00:59 2011 +0100
+++ b/src/HOL/Tools/ATP/atp_util.ML Fri Dec 16 12:01:10 2011 +0100
@@ -23,9 +23,7 @@
val varify_type : Proof.context -> typ -> typ
val instantiate_type : theory -> typ -> typ -> typ -> typ
val varify_and_instantiate_type : Proof.context -> typ -> typ -> typ -> typ
- val typ_of_dtyp :
- Datatype_Aux.descr -> (Datatype_Aux.dtyp * typ) list -> Datatype_Aux.dtyp
- -> typ
+ val typ_of_dtyp : Datatype.descr -> (Datatype.dtyp * typ) list -> Datatype.dtyp -> typ
val is_type_surely_finite : Proof.context -> typ -> bool
val is_type_surely_infinite : Proof.context -> bool -> typ list -> typ -> bool
val s_not : term -> term
@@ -148,11 +146,11 @@
instantiate_type thy (varify_type ctxt T1) T1' (varify_type ctxt T2)
end
-fun typ_of_dtyp _ typ_assoc (Datatype_Aux.DtTFree a) =
- the (AList.lookup (op =) typ_assoc (Datatype_Aux.DtTFree a))
- | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtType (s, Us)) =
+fun typ_of_dtyp _ typ_assoc (Datatype.DtTFree a) =
+ the (AList.lookup (op =) typ_assoc (Datatype.DtTFree a))
+ | typ_of_dtyp descr typ_assoc (Datatype.DtType (s, Us)) =
Type (s, map (typ_of_dtyp descr typ_assoc) Us)
- | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtRec i) =
+ | typ_of_dtyp descr typ_assoc (Datatype.DtRec i) =
let val (s, ds, _) = the (AList.lookup (op =) descr i) in
Type (s, map (typ_of_dtyp descr typ_assoc) ds)
end
--- a/src/HOL/Tools/Datatype/datatype_abs_proofs.ML Fri Dec 16 12:00:59 2011 +0100
+++ b/src/HOL/Tools/Datatype/datatype_abs_proofs.ML Fri Dec 16 12:01:10 2011 +0100
@@ -8,7 +8,8 @@
signature DATATYPE_ABS_PROOFS =
sig
- include DATATYPE_COMMON
+ type config = Datatype_Aux.config
+ type descr = Datatype_Aux.descr
val prove_casedist_thms : config -> string list -> descr list -> thm ->
attribute list -> theory -> thm list * theory
val prove_primrec_thms : config -> string list -> descr list ->
@@ -29,10 +30,13 @@
structure Datatype_Abs_Proofs: DATATYPE_ABS_PROOFS =
struct
+type config = Datatype_Aux.config;
+type descr = Datatype_Aux.descr;
+
+
(************************ case distinction theorems ***************************)
-fun prove_casedist_thms (config : Datatype_Aux.config)
- new_type_names descr induct case_names_exhausts thy =
+fun prove_casedist_thms (config : config) new_type_names descr induct case_names_exhausts thy =
let
val _ = Datatype_Aux.message config "Proving case distinction theorems ...";
@@ -79,7 +83,7 @@
(*************************** primrec combinators ******************************)
-fun prove_primrec_thms (config : Datatype_Aux.config) new_type_names descr
+fun prove_primrec_thms (config : config) new_type_names descr
injects_of constr_inject (dist_rewrites, other_dist_rewrites) induct thy =
let
val _ = Datatype_Aux.message config "Constructing primrec combinators ...";
@@ -275,8 +279,7 @@
(***************************** case combinators *******************************)
-fun prove_case_thms (config : Datatype_Aux.config)
- new_type_names descr reccomb_names primrec_thms thy =
+fun prove_case_thms (config : config) new_type_names descr reccomb_names primrec_thms thy =
let
val _ = Datatype_Aux.message config "Proving characteristic theorems for case combinators ...";
@@ -350,7 +353,7 @@
(******************************* case splitting *******************************)
-fun prove_split_thms (config : Datatype_Aux.config)
+fun prove_split_thms (config : config)
new_type_names case_names descr constr_inject dist_rewrites casedist_thms case_thms thy =
let
val _ = Datatype_Aux.message config "Proving equations for case splitting ...";
@@ -399,7 +402,7 @@
(************************* additional theorems for TFL ************************)
-fun prove_nchotomys (config : Datatype_Aux.config) new_type_names descr casedist_thms thy =
+fun prove_nchotomys (config : config) new_type_names descr casedist_thms thy =
let
val _ = Datatype_Aux.message config "Proving additional theorems for TFL ...";
@@ -449,7 +452,4 @@
in thy |> Datatype_Aux.store_thms "case_cong" new_type_names case_congs end;
-
-open Datatype_Aux;
-
end;
--- a/src/HOL/Tools/Datatype/datatype_case.ML Fri Dec 16 12:00:59 2011 +0100
+++ b/src/HOL/Tools/Datatype/datatype_case.ML Fri Dec 16 12:01:10 2011 +0100
@@ -130,7 +130,7 @@
names = names,
constraints = cnstrts,
group = in_group'} :: part cs not_in_group
- end
+ end;
in part constructors rows end;
fun v_to_prfx (prfx, Free v :: pats) = (v :: prfx, pats)
@@ -143,7 +143,6 @@
let
val get_info = Datatype_Data.info_of_constr_permissive (Proof_Context.theory_of ctxt);
- val name = singleton (Name.variant_list used) "a";
fun expand constructors used ty ((_, []), _) = raise CASE_ERROR ("mk_case: expand_var_row", ~1)
| expand constructors used ty (row as ((prfx, p :: ps), (rhs, tag))) =
if is_Free p then
@@ -153,7 +152,10 @@
let val capp = list_comb (fresh_constr ty_match ty_inst ty used' c)
in ((prfx, capp :: ps), (subst_free [(p, capp)] rhs, tag)) end
in map expnd constructors end
- else [row]
+ else [row];
+
+ val name = singleton (Name.variant_list used) "a";
+
fun mk _ [] = raise CASE_ERROR ("no rows", ~1)
| mk [] (((_, []), (tm, tag)) :: _) = ([tag], tm) (* Done *)
| mk path (rows as ((row as ((_, [Free _]), _)) :: _ :: _)) = mk path [row]
@@ -277,19 +279,22 @@
val (u', used'') = prep_pat u used';
in (t' $ u', used'') end
| prep_pat t used = case_error ("Bad pattern: " ^ Syntax.string_of_term ctxt t);
+
fun dest_case1 (t as Const (@{syntax_const "_case1"}, _) $ l $ r) =
let val (l', cnstrts) = strip_constraints l
in ((fst (prep_pat l' (Term.add_free_names t [])), r), cnstrts) end
| dest_case1 t = case_error "dest_case1";
+
fun dest_case2 (Const (@{syntax_const "_case2"}, _) $ t $ u) = t :: dest_case2 u
| dest_case2 t = [t];
+
val (cases, cnstrts) = split_list (map dest_case1 (dest_case2 u));
- val case_tm =
- make_case_untyped ctxt
- (if err then Error else Warning) []
- (fold (fn tT => fn t => Syntax.const @{syntax_const "_constrain"} $ t $ tT)
- (flat cnstrts) t) cases;
- in case_tm end
+ in
+ make_case_untyped ctxt
+ (if err then Error else Warning) []
+ (fold (fn tT => fn t => Syntax.const @{syntax_const "_constrain"} $ t $ tT)
+ (flat cnstrts) t) cases
+ end
| case_tr _ _ _ = case_error "case_tr";
val trfun_setup =
--- a/src/HOL/Tools/Datatype/datatype_prop.ML Fri Dec 16 12:00:59 2011 +0100
+++ b/src/HOL/Tools/Datatype/datatype_prop.ML Fri Dec 16 12:01:10 2011 +0100
@@ -6,7 +6,7 @@
signature DATATYPE_PROP =
sig
- include DATATYPE_COMMON
+ type descr = Datatype_Aux.descr
val indexify_names: string list -> string list
val make_tnames: typ list -> string list
val make_injs : descr list -> term list list
@@ -26,6 +26,9 @@
structure Datatype_Prop : DATATYPE_PROP =
struct
+type descr = Datatype_Aux.descr;
+
+
fun indexify_names names =
let
fun index (x :: xs) tab =
@@ -429,7 +432,4 @@
(hd descr ~~ newTs)
end;
-
-open Datatype_Aux;
-
end;
--- a/src/HOL/Tools/Datatype/datatype_realizer.ML Fri Dec 16 12:00:59 2011 +0100
+++ b/src/HOL/Tools/Datatype/datatype_realizer.ML Fri Dec 16 12:01:10 2011 +0100
@@ -25,7 +25,7 @@
fun tname_of (Type (s, _)) = s
| tname_of _ = "";
-fun make_ind ({descr, rec_names, rec_rewrites, induct, ...} : Datatype_Aux.info) is thy =
+fun make_ind ({descr, rec_names, rec_rewrites, induct, ...} : Datatype.info) is thy =
let
val recTs = Datatype_Aux.get_rec_types descr;
val pnames =
@@ -157,7 +157,7 @@
in Extraction.add_realizers_i [(ind_name, (vs, r', prf))] thy' end;
-fun make_casedists ({index, descr, case_name, case_rewrites, exhaust, ...} : Datatype_Aux.info) thy =
+fun make_casedists ({index, descr, case_name, case_rewrites, exhaust, ...} : Datatype.info) thy =
let
val cert = cterm_of thy;
val rT = TFree ("'P", HOLogic.typeS);
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Datatype/primrec.ML Fri Dec 16 12:01:10 2011 +0100
@@ -0,0 +1,318 @@
+(* Title: HOL/Tools/Datatype/primrec.ML
+ Author: Norbert Voelker, FernUni Hagen
+ Author: Stefan Berghofer, TU Muenchen
+ Author: Florian Haftmann, TU Muenchen
+
+Primitive recursive functions on datatypes.
+*)
+
+signature PRIMREC =
+sig
+ val add_primrec: (binding * typ option * mixfix) list ->
+ (Attrib.binding * term) list -> local_theory -> (term list * thm list) * local_theory
+ val add_primrec_cmd: (binding * string option * mixfix) list ->
+ (Attrib.binding * string) list -> local_theory -> (term list * thm list) * local_theory
+ val add_primrec_global: (binding * typ option * mixfix) list ->
+ (Attrib.binding * term) list -> theory -> (term list * thm list) * theory
+ val add_primrec_overloaded: (string * (string * typ) * bool) list ->
+ (binding * typ option * mixfix) list ->
+ (Attrib.binding * term) list -> theory -> (term list * thm list) * theory
+ val add_primrec_simple: ((binding * typ) * mixfix) list -> term list ->
+ local_theory -> (string * (term list * thm list)) * local_theory
+end;
+
+structure Primrec : PRIMREC =
+struct
+
+exception PrimrecError of string * term option;
+
+fun primrec_error msg = raise PrimrecError (msg, NONE);
+fun primrec_error_eqn msg eqn = raise PrimrecError (msg, SOME eqn);
+
+
+(* preprocessing of equations *)
+
+fun process_eqn is_fixed spec rec_fns =
+ let
+ val (vs, Ts) = split_list (strip_qnt_vars "all" spec);
+ val body = strip_qnt_body "all" spec;
+ val (vs', _) = fold_map Name.variant vs (Name.make_context (fold_aterms
+ (fn Free (v, _) => insert (op =) v | _ => I) body []));
+ val eqn = curry subst_bounds (map2 (curry Free) vs' Ts |> rev) body;
+ val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop eqn)
+ handle TERM _ => primrec_error "not a proper equation";
+ val (recfun, args) = strip_comb lhs;
+ val fname =
+ (case recfun of
+ Free (v, _) =>
+ if is_fixed v then v
+ else primrec_error "illegal head of function equation"
+ | _ => primrec_error "illegal head of function equation");
+
+ val (ls', rest) = take_prefix is_Free args;
+ val (middle, rs') = take_suffix is_Free rest;
+ val rpos = length ls';
+
+ val (constr, cargs') =
+ if null middle then primrec_error "constructor missing"
+ else strip_comb (hd middle);
+ val (cname, T) = dest_Const constr
+ handle TERM _ => primrec_error "ill-formed constructor";
+ val (tname, _) = dest_Type (body_type T) handle TYPE _ =>
+ primrec_error "cannot determine datatype associated with function"
+
+ val (ls, cargs, rs) =
+ (map dest_Free ls', map dest_Free cargs', map dest_Free rs')
+ handle TERM _ => primrec_error "illegal argument in pattern";
+ val lfrees = ls @ rs @ cargs;
+
+ fun check_vars _ [] = ()
+ | check_vars s vars = primrec_error (s ^ commas_quote (map fst vars)) eqn;
+ in
+ if length middle > 1 then
+ primrec_error "more than one non-variable in pattern"
+ else
+ (check_vars "repeated variable names in pattern: " (duplicates (op =) lfrees);
+ check_vars "extra variables on rhs: "
+ (Term.add_frees rhs [] |> subtract (op =) lfrees
+ |> filter_out (is_fixed o fst));
+ (case AList.lookup (op =) rec_fns fname of
+ NONE =>
+ (fname, (tname, rpos, [(cname, (ls, cargs, rs, rhs, eqn))])) :: rec_fns
+ | SOME (_, rpos', eqns) =>
+ if AList.defined (op =) eqns cname then
+ primrec_error "constructor already occurred as pattern"
+ else if rpos <> rpos' then
+ primrec_error "position of recursive argument inconsistent"
+ else
+ AList.update (op =)
+ (fname, (tname, rpos, (cname, (ls, cargs, rs, rhs, eqn)) :: eqns))
+ rec_fns))
+ end handle PrimrecError (msg, NONE) => primrec_error_eqn msg spec;
+
+fun process_fun descr eqns (i, fname) (fnames, fnss) =
+ let
+ val (_, (tname, _, constrs)) = nth descr i;
+
+ (* substitute "fname ls x rs" by "y ls rs" for (x, (_, y)) in subs *)
+
+ fun subst [] t fs = (t, fs)
+ | subst subs (Abs (a, T, t)) fs =
+ fs
+ |> subst subs t
+ |-> (fn t' => pair (Abs (a, T, t')))
+ | subst subs (t as (_ $ _)) fs =
+ let
+ val (f, ts) = strip_comb t;
+ in
+ if is_Free f
+ andalso member (fn ((v, _), (w, _)) => v = w) eqns (dest_Free f) then
+ let
+ val (fname', _) = dest_Free f;
+ val (_, rpos, _) = the (AList.lookup (op =) eqns fname');
+ val (ls, rs) = chop rpos ts
+ val (x', rs') =
+ (case rs of
+ x' :: rs => (x', rs)
+ | [] => primrec_error ("not enough arguments in recursive application\n" ^
+ "of function " ^ quote fname' ^ " on rhs"));
+ val (x, xs) = strip_comb x';
+ in
+ (case AList.lookup (op =) subs x of
+ NONE =>
+ fs
+ |> fold_map (subst subs) ts
+ |-> (fn ts' => pair (list_comb (f, ts')))
+ | SOME (i', y) =>
+ fs
+ |> fold_map (subst subs) (xs @ ls @ rs')
+ ||> process_fun descr eqns (i', fname')
+ |-> (fn ts' => pair (list_comb (y, ts'))))
+ end
+ else
+ fs
+ |> fold_map (subst subs) (f :: ts)
+ |-> (fn f' :: ts' => pair (list_comb (f', ts')))
+ end
+ | subst _ t fs = (t, fs);
+
+ (* translate rec equations into function arguments suitable for rec comb *)
+
+ fun trans eqns (cname, cargs) (fnames', fnss', fns) =
+ (case AList.lookup (op =) eqns cname of
+ NONE => (warning ("No equation for constructor " ^ quote cname ^
+ "\nin definition of function " ^ quote fname);
+ (fnames', fnss', (Const (@{const_name undefined}, dummyT)) :: fns))
+ | SOME (ls, cargs', rs, rhs, eq) =>
+ let
+ val recs = filter (Datatype_Aux.is_rec_type o snd) (cargs' ~~ cargs);
+ val rargs = map fst recs;
+ val subs = map (rpair dummyT o fst)
+ (rev (Term.rename_wrt_term rhs rargs));
+ val (rhs', (fnames'', fnss'')) = subst (map2 (fn (x, y) => fn z =>
+ (Free x, (Datatype_Aux.body_index y, Free z))) recs subs) rhs (fnames', fnss')
+ handle PrimrecError (s, NONE) => primrec_error_eqn s eq
+ in
+ (fnames'', fnss'', fold_rev absfree (cargs' @ subs @ ls @ rs) rhs' :: fns)
+ end)
+
+ in
+ (case AList.lookup (op =) fnames i of
+ NONE =>
+ if exists (fn (_, v) => fname = v) fnames then
+ primrec_error ("inconsistent functions for datatype " ^ quote tname)
+ else
+ let
+ val (_, _, eqns) = the (AList.lookup (op =) eqns fname);
+ val (fnames', fnss', fns) = fold_rev (trans eqns) constrs
+ ((i, fname) :: fnames, fnss, [])
+ in
+ (fnames', (i, (fname, #1 (snd (hd eqns)), fns)) :: fnss')
+ end
+ | SOME fname' =>
+ if fname = fname' then (fnames, fnss)
+ else primrec_error ("inconsistent functions for datatype " ^ quote tname))
+ end;
+
+
+(* prepare functions needed for definitions *)
+
+fun get_fns fns ((i : int, (tname, _, constrs)), rec_name) (fs, defs) =
+ (case AList.lookup (op =) fns i of
+ NONE =>
+ let
+ val dummy_fns = map (fn (_, cargs) => Const (@{const_name undefined},
+ replicate (length cargs + length (filter Datatype_Aux.is_rec_type cargs))
+ dummyT ---> HOLogic.unitT)) constrs;
+ val _ = warning ("No function definition for datatype " ^ quote tname)
+ in
+ (dummy_fns @ fs, defs)
+ end
+ | SOME (fname, ls, fs') => (fs' @ fs, (fname, ls, rec_name, tname) :: defs));
+
+
+(* make definition *)
+
+fun make_def ctxt fixes fs (fname, ls, rec_name, tname) =
+ let
+ val SOME (var, varT) = get_first (fn ((b, T), mx) =>
+ if Binding.name_of b = fname then SOME ((b, mx), T) else NONE) fixes;
+ val def_name = Thm.def_name (Long_Name.base_name fname);
+ val raw_rhs = fold_rev (fn T => fn t => Abs ("", T, t)) (map snd ls @ [dummyT])
+ (list_comb (Const (rec_name, dummyT), fs @ map Bound (0 :: (length ls downto 1))))
+ val rhs = singleton (Syntax.check_terms ctxt) (Type.constraint varT raw_rhs);
+ in (var, ((Binding.conceal (Binding.name def_name), []), rhs)) end;
+
+
+(* find datatypes which contain all datatypes in tnames' *)
+
+fun find_dts _ _ [] = []
+ | find_dts dt_info tnames' (tname :: tnames) =
+ (case Symtab.lookup dt_info tname of
+ NONE => primrec_error (quote tname ^ " is not a datatype")
+ | SOME (dt : Datatype_Aux.info) =>
+ if subset (op =) (tnames', map (#1 o snd) (#descr dt)) then
+ (tname, dt) :: (find_dts dt_info tnames' tnames)
+ else find_dts dt_info tnames' tnames);
+
+
+(* distill primitive definition(s) from primrec specification *)
+
+fun distill ctxt fixes eqs =
+ let
+ val eqns = fold_rev (process_eqn (fn v => Variable.is_fixed ctxt 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_Data.get_all (Proof_Context.theory_of ctxt)) 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, ...} =
+ if null dts then primrec_error
+ ("datatypes " ^ commas_quote tnames ^ "\nare not mutually recursive")
+ else snd (hd dts);
+ val (fnames, fnss) = fold_rev (process_fun descr eqns) main_fns ([], []);
+ val (fs, raw_defs) = fold_rev (get_fns fnss) (descr ~~ rec_names) ([], []);
+ val defs = map (make_def ctxt fixes fs) raw_defs;
+ val names = map snd fnames;
+ val names_eqns = map fst eqns;
+ val _ =
+ if eq_set (op =) (names, names_eqns) then ()
+ else primrec_error ("functions " ^ commas_quote names_eqns ^
+ "\nare not mutually recursive");
+ val rec_rewrites' = map mk_meta_eq rec_rewrites;
+ val prefix = space_implode "_" (map (Long_Name.base_name o #1) raw_defs);
+ fun prove ctxt defs =
+ let
+ val frees = fold (Variable.add_free_names ctxt) eqs [];
+ val rewrites = rec_rewrites' @ map (snd o snd) defs;
+ fun tac _ = EVERY [rewrite_goals_tac rewrites, rtac refl 1];
+ in map (fn eq => Goal.prove ctxt frees [] eq tac) eqs end;
+ in ((prefix, (fs, defs)), prove) end
+ handle PrimrecError (msg, some_eqn) =>
+ error ("Primrec definition error:\n" ^ msg ^
+ (case some_eqn of
+ SOME eqn => "\nin\n" ^ quote (Syntax.string_of_term ctxt eqn)
+ | NONE => ""));
+
+
+(* primrec definition *)
+
+fun add_primrec_simple fixes ts lthy =
+ let
+ val ((prefix, (_, defs)), prove) = distill lthy fixes ts;
+ in
+ lthy
+ |> fold_map Local_Theory.define defs
+ |-> (fn defs => `(fn lthy => (prefix, (map fst defs, prove lthy defs))))
+ end;
+
+local
+
+fun gen_primrec prep_spec raw_fixes raw_spec lthy =
+ let
+ val (fixes, spec) = fst (prep_spec raw_fixes raw_spec lthy);
+ fun attr_bindings prefix = map (fn ((b, attrs), _) =>
+ (Binding.qualify false prefix b, Code.add_default_eqn_attrib :: attrs)) spec;
+ fun simp_attr_binding prefix =
+ (Binding.qualify true prefix (Binding.name "simps"), @{attributes [simp, nitpick_simp]});
+ in
+ lthy
+ |> add_primrec_simple fixes (map snd spec)
+ |-> (fn (prefix, (ts, simps)) =>
+ Spec_Rules.add Spec_Rules.Equational (ts, simps)
+ #> fold_map Local_Theory.note (attr_bindings prefix ~~ map single simps)
+ #-> (fn simps' => Local_Theory.note (simp_attr_binding prefix, maps snd simps')
+ #>> (fn (_, simps'') => (ts, simps''))))
+ end;
+
+in
+
+val add_primrec = gen_primrec Specification.check_spec;
+val add_primrec_cmd = gen_primrec Specification.read_spec;
+
+end;
+
+fun add_primrec_global fixes specs thy =
+ let
+ val lthy = Named_Target.theory_init thy;
+ val ((ts, simps), lthy') = add_primrec fixes specs lthy;
+ val simps' = Proof_Context.export lthy' lthy simps;
+ in ((ts, simps'), Local_Theory.exit_global lthy') end;
+
+fun add_primrec_overloaded ops fixes specs thy =
+ let
+ val lthy = Overloading.overloading ops thy;
+ val ((ts, simps), lthy') = add_primrec fixes specs lthy;
+ val simps' = Proof_Context.export lthy' lthy simps;
+ in ((ts, simps'), Local_Theory.exit_global lthy') end;
+
+
+(* outer syntax *)
+
+val _ =
+ Outer_Syntax.local_theory "primrec" "define primitive recursive functions on datatypes"
+ Keyword.thy_decl
+ (Parse.fixes -- Parse_Spec.where_alt_specs
+ >> (fn (fixes, specs) => add_primrec_cmd fixes specs #> snd));
+
+end;
--- a/src/HOL/Tools/Function/size.ML Fri Dec 16 12:00:59 2011 +0100
+++ b/src/HOL/Tools/Function/size.ML Fri Dec 16 12:01:10 2011 +0100
@@ -40,7 +40,7 @@
NONE => Abs ("x", T, HOLogic.zero)
| SOME t => t);
-fun is_poly thy (Datatype_Aux.DtType (name, dts)) =
+fun is_poly thy (Datatype.DtType (name, dts)) =
(case Datatype.get_info thy name of
NONE => false
| SOME _ => exists (is_poly thy) dts)
--- a/src/HOL/Tools/Nitpick/nitpick_util.ML Fri Dec 16 12:00:59 2011 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_util.ML Fri Dec 16 12:01:10 2011 +0100
@@ -61,9 +61,7 @@
val int_T : typ
val simple_string_of_typ : typ -> string
val is_real_constr : theory -> string * typ -> bool
- val typ_of_dtyp :
- Datatype_Aux.descr -> (Datatype_Aux.dtyp * typ) list -> Datatype_Aux.dtyp
- -> typ
+ val typ_of_dtyp : Datatype.descr -> (Datatype.dtyp * typ) list -> Datatype.dtyp -> typ
val varify_type : Proof.context -> typ -> typ
val instantiate_type : theory -> typ -> typ -> typ -> typ
val varify_and_instantiate_type : Proof.context -> typ -> typ -> typ -> typ
--- a/src/HOL/Tools/primrec.ML Fri Dec 16 12:00:59 2011 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,318 +0,0 @@
-(* Title: HOL/Tools/primrec.ML
- Author: Norbert Voelker, FernUni Hagen
- Author: Stefan Berghofer, TU Muenchen
- Author: Florian Haftmann, TU Muenchen
-
-Primitive recursive functions on datatypes.
-*)
-
-signature PRIMREC =
-sig
- val add_primrec: (binding * typ option * mixfix) list ->
- (Attrib.binding * term) list -> local_theory -> (term list * thm list) * local_theory
- val add_primrec_cmd: (binding * string option * mixfix) list ->
- (Attrib.binding * string) list -> local_theory -> (term list * thm list) * local_theory
- val add_primrec_global: (binding * typ option * mixfix) list ->
- (Attrib.binding * term) list -> theory -> (term list * thm list) * theory
- val add_primrec_overloaded: (string * (string * typ) * bool) list ->
- (binding * typ option * mixfix) list ->
- (Attrib.binding * term) list -> theory -> (term list * thm list) * theory
- val add_primrec_simple: ((binding * typ) * mixfix) list -> term list ->
- local_theory -> (string * (term list * thm list)) * local_theory
-end;
-
-structure Primrec : PRIMREC =
-struct
-
-exception PrimrecError of string * term option;
-
-fun primrec_error msg = raise PrimrecError (msg, NONE);
-fun primrec_error_eqn msg eqn = raise PrimrecError (msg, SOME eqn);
-
-
-(* preprocessing of equations *)
-
-fun process_eqn is_fixed spec rec_fns =
- let
- val (vs, Ts) = split_list (strip_qnt_vars "all" spec);
- val body = strip_qnt_body "all" spec;
- val (vs', _) = fold_map Name.variant vs (Name.make_context (fold_aterms
- (fn Free (v, _) => insert (op =) v | _ => I) body []));
- val eqn = curry subst_bounds (map2 (curry Free) vs' Ts |> rev) body;
- val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop eqn)
- handle TERM _ => primrec_error "not a proper equation";
- val (recfun, args) = strip_comb lhs;
- val fname =
- (case recfun of
- Free (v, _) =>
- if is_fixed v then v
- else primrec_error "illegal head of function equation"
- | _ => primrec_error "illegal head of function equation");
-
- val (ls', rest) = take_prefix is_Free args;
- val (middle, rs') = take_suffix is_Free rest;
- val rpos = length ls';
-
- val (constr, cargs') =
- if null middle then primrec_error "constructor missing"
- else strip_comb (hd middle);
- val (cname, T) = dest_Const constr
- handle TERM _ => primrec_error "ill-formed constructor";
- val (tname, _) = dest_Type (body_type T) handle TYPE _ =>
- primrec_error "cannot determine datatype associated with function"
-
- val (ls, cargs, rs) =
- (map dest_Free ls', map dest_Free cargs', map dest_Free rs')
- handle TERM _ => primrec_error "illegal argument in pattern";
- val lfrees = ls @ rs @ cargs;
-
- fun check_vars _ [] = ()
- | check_vars s vars = primrec_error (s ^ commas_quote (map fst vars)) eqn;
- in
- if length middle > 1 then
- primrec_error "more than one non-variable in pattern"
- else
- (check_vars "repeated variable names in pattern: " (duplicates (op =) lfrees);
- check_vars "extra variables on rhs: "
- (Term.add_frees rhs [] |> subtract (op =) lfrees
- |> filter_out (is_fixed o fst));
- (case AList.lookup (op =) rec_fns fname of
- NONE =>
- (fname, (tname, rpos, [(cname, (ls, cargs, rs, rhs, eqn))])) :: rec_fns
- | SOME (_, rpos', eqns) =>
- if AList.defined (op =) eqns cname then
- primrec_error "constructor already occurred as pattern"
- else if rpos <> rpos' then
- primrec_error "position of recursive argument inconsistent"
- else
- AList.update (op =)
- (fname, (tname, rpos, (cname, (ls, cargs, rs, rhs, eqn)) :: eqns))
- rec_fns))
- end handle PrimrecError (msg, NONE) => primrec_error_eqn msg spec;
-
-fun process_fun descr eqns (i, fname) (fnames, fnss) =
- let
- val (_, (tname, _, constrs)) = nth descr i;
-
- (* substitute "fname ls x rs" by "y ls rs" for (x, (_, y)) in subs *)
-
- fun subst [] t fs = (t, fs)
- | subst subs (Abs (a, T, t)) fs =
- fs
- |> subst subs t
- |-> (fn t' => pair (Abs (a, T, t')))
- | subst subs (t as (_ $ _)) fs =
- let
- val (f, ts) = strip_comb t;
- in
- if is_Free f
- andalso member (fn ((v, _), (w, _)) => v = w) eqns (dest_Free f) then
- let
- val (fname', _) = dest_Free f;
- val (_, rpos, _) = the (AList.lookup (op =) eqns fname');
- val (ls, rs) = chop rpos ts
- val (x', rs') =
- (case rs of
- x' :: rs => (x', rs)
- | [] => primrec_error ("not enough arguments in recursive application\n" ^
- "of function " ^ quote fname' ^ " on rhs"));
- val (x, xs) = strip_comb x';
- in
- (case AList.lookup (op =) subs x of
- NONE =>
- fs
- |> fold_map (subst subs) ts
- |-> (fn ts' => pair (list_comb (f, ts')))
- | SOME (i', y) =>
- fs
- |> fold_map (subst subs) (xs @ ls @ rs')
- ||> process_fun descr eqns (i', fname')
- |-> (fn ts' => pair (list_comb (y, ts'))))
- end
- else
- fs
- |> fold_map (subst subs) (f :: ts)
- |-> (fn f' :: ts' => pair (list_comb (f', ts')))
- end
- | subst _ t fs = (t, fs);
-
- (* translate rec equations into function arguments suitable for rec comb *)
-
- fun trans eqns (cname, cargs) (fnames', fnss', fns) =
- (case AList.lookup (op =) eqns cname of
- NONE => (warning ("No equation for constructor " ^ quote cname ^
- "\nin definition of function " ^ quote fname);
- (fnames', fnss', (Const (@{const_name undefined}, dummyT)) :: fns))
- | SOME (ls, cargs', rs, rhs, eq) =>
- let
- val recs = filter (Datatype_Aux.is_rec_type o snd) (cargs' ~~ cargs);
- val rargs = map fst recs;
- val subs = map (rpair dummyT o fst)
- (rev (Term.rename_wrt_term rhs rargs));
- val (rhs', (fnames'', fnss'')) = subst (map2 (fn (x, y) => fn z =>
- (Free x, (Datatype_Aux.body_index y, Free z))) recs subs) rhs (fnames', fnss')
- handle PrimrecError (s, NONE) => primrec_error_eqn s eq
- in
- (fnames'', fnss'', fold_rev absfree (cargs' @ subs @ ls @ rs) rhs' :: fns)
- end)
-
- in
- (case AList.lookup (op =) fnames i of
- NONE =>
- if exists (fn (_, v) => fname = v) fnames then
- primrec_error ("inconsistent functions for datatype " ^ quote tname)
- else
- let
- val (_, _, eqns) = the (AList.lookup (op =) eqns fname);
- val (fnames', fnss', fns) = fold_rev (trans eqns) constrs
- ((i, fname) :: fnames, fnss, [])
- in
- (fnames', (i, (fname, #1 (snd (hd eqns)), fns)) :: fnss')
- end
- | SOME fname' =>
- if fname = fname' then (fnames, fnss)
- else primrec_error ("inconsistent functions for datatype " ^ quote tname))
- end;
-
-
-(* prepare functions needed for definitions *)
-
-fun get_fns fns ((i : int, (tname, _, constrs)), rec_name) (fs, defs) =
- (case AList.lookup (op =) fns i of
- NONE =>
- let
- val dummy_fns = map (fn (_, cargs) => Const (@{const_name undefined},
- replicate (length cargs + length (filter Datatype_Aux.is_rec_type cargs))
- dummyT ---> HOLogic.unitT)) constrs;
- val _ = warning ("No function definition for datatype " ^ quote tname)
- in
- (dummy_fns @ fs, defs)
- end
- | SOME (fname, ls, fs') => (fs' @ fs, (fname, ls, rec_name, tname) :: defs));
-
-
-(* make definition *)
-
-fun make_def ctxt fixes fs (fname, ls, rec_name, tname) =
- let
- val SOME (var, varT) = get_first (fn ((b, T), mx) =>
- if Binding.name_of b = fname then SOME ((b, mx), T) else NONE) fixes;
- val def_name = Thm.def_name (Long_Name.base_name fname);
- val raw_rhs = fold_rev (fn T => fn t => Abs ("", T, t)) (map snd ls @ [dummyT])
- (list_comb (Const (rec_name, dummyT), fs @ map Bound (0 :: (length ls downto 1))))
- val rhs = singleton (Syntax.check_terms ctxt) (Type.constraint varT raw_rhs);
- in (var, ((Binding.conceal (Binding.name def_name), []), rhs)) end;
-
-
-(* find datatypes which contain all datatypes in tnames' *)
-
-fun find_dts (dt_info : Datatype_Aux.info Symtab.table) _ [] = []
- | find_dts dt_info tnames' (tname :: tnames) =
- (case Symtab.lookup dt_info tname of
- NONE => primrec_error (quote tname ^ " is not a datatype")
- | SOME dt =>
- if subset (op =) (tnames', map (#1 o snd) (#descr dt)) then
- (tname, dt) :: (find_dts dt_info tnames' tnames)
- else find_dts dt_info tnames' tnames);
-
-
-(* distill primitive definition(s) from primrec specification *)
-
-fun distill lthy fixes eqs =
- let
- 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_Data.get_all (Proof_Context.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, ...} =
- if null dts then primrec_error
- ("datatypes " ^ commas_quote tnames ^ "\nare not mutually recursive")
- else snd (hd dts);
- val (fnames, fnss) = fold_rev (process_fun descr eqns) main_fns ([], []);
- val (fs, raw_defs) = fold_rev (get_fns fnss) (descr ~~ rec_names) ([], []);
- val defs = map (make_def lthy fixes fs) raw_defs;
- val names = map snd fnames;
- val names_eqns = map fst eqns;
- val _ =
- if eq_set (op =) (names, names_eqns) then ()
- else primrec_error ("functions " ^ commas_quote names_eqns ^
- "\nare not mutually recursive");
- val rec_rewrites' = map mk_meta_eq rec_rewrites;
- val prefix = space_implode "_" (map (Long_Name.base_name o #1) raw_defs);
- fun prove lthy defs =
- let
- val frees = fold (Variable.add_free_names lthy) eqs [];
- val rewrites = rec_rewrites' @ map (snd o snd) defs;
- fun tac _ = EVERY [rewrite_goals_tac rewrites, rtac refl 1];
- in map (fn eq => Goal.prove lthy frees [] eq tac) eqs end;
- in ((prefix, (fs, defs)), prove) end
- handle PrimrecError (msg, some_eqn) =>
- error ("Primrec definition error:\n" ^ msg ^
- (case some_eqn of
- SOME eqn => "\nin\n" ^ quote (Syntax.string_of_term lthy eqn)
- | NONE => ""));
-
-
-(* primrec definition *)
-
-fun add_primrec_simple fixes ts lthy =
- let
- val ((prefix, (fs, defs)), prove) = distill lthy fixes ts;
- in
- lthy
- |> fold_map Local_Theory.define defs
- |-> (fn defs => `(fn lthy => (prefix, (map fst defs, prove lthy defs))))
- end;
-
-local
-
-fun gen_primrec prep_spec raw_fixes raw_spec lthy =
- let
- val (fixes, spec) = fst (prep_spec raw_fixes raw_spec lthy);
- fun attr_bindings prefix = map (fn ((b, attrs), _) =>
- (Binding.qualify false prefix b, Code.add_default_eqn_attrib :: attrs)) spec;
- fun simp_attr_binding prefix =
- (Binding.qualify true prefix (Binding.name "simps"), @{attributes [simp, nitpick_simp]});
- in
- lthy
- |> add_primrec_simple fixes (map snd spec)
- |-> (fn (prefix, (ts, simps)) =>
- Spec_Rules.add Spec_Rules.Equational (ts, simps)
- #> fold_map Local_Theory.note (attr_bindings prefix ~~ map single simps)
- #-> (fn simps' => Local_Theory.note (simp_attr_binding prefix, maps snd simps')
- #>> (fn (_, simps'') => (ts, simps''))))
- end;
-
-in
-
-val add_primrec = gen_primrec Specification.check_spec;
-val add_primrec_cmd = gen_primrec Specification.read_spec;
-
-end;
-
-fun add_primrec_global fixes specs thy =
- let
- val lthy = Named_Target.theory_init thy;
- val ((ts, simps), lthy') = add_primrec fixes specs lthy;
- val simps' = Proof_Context.export lthy' lthy simps;
- in ((ts, simps'), Local_Theory.exit_global lthy') end;
-
-fun add_primrec_overloaded ops fixes specs thy =
- let
- val lthy = Overloading.overloading ops thy;
- val ((ts, simps), lthy') = add_primrec fixes specs lthy;
- val simps' = Proof_Context.export lthy' lthy simps;
- in ((ts, simps'), Local_Theory.exit_global lthy') end;
-
-
-(* outer syntax *)
-
-val _ =
- Outer_Syntax.local_theory "primrec" "define primitive recursive functions on datatypes"
- Keyword.thy_decl
- (Parse.fixes -- Parse_Spec.where_alt_specs
- >> (fn (fixes, specs) => add_primrec_cmd fixes specs #> snd));
-
-end;
--- a/src/HOL/Tools/refute.ML Fri Dec 16 12:00:59 2011 +0100
+++ b/src/HOL/Tools/refute.ML Fri Dec 16 12:01:10 2011 +0100
@@ -862,7 +862,7 @@
(* sanity check: every element in 'dtyps' must be a *)
(* 'DtTFree' *)
val _ = if Library.exists (fn d =>
- case d of Datatype_Aux.DtTFree _ => false | _ => true) typs then
+ case d of Datatype.DtTFree _ => false | _ => true) typs then
raise REFUTE ("ground_types", "datatype argument (for type "
^ Syntax.string_of_typ ctxt T ^ ") is not a variable")
else ()
@@ -874,11 +874,11 @@
val dT = typ_of_dtyp descr typ_assoc d
in
case d of
- Datatype_Aux.DtTFree _ =>
+ Datatype.DtTFree _ =>
collect_types dT acc
- | Datatype_Aux.DtType (_, ds) =>
+ | Datatype.DtType (_, ds) =>
collect_types dT (fold_rev collect_dtyp ds acc)
- | Datatype_Aux.DtRec i =>
+ | Datatype.DtRec i =>
if member (op =) acc dT then
acc (* prevent infinite recursion *)
else
@@ -901,7 +901,7 @@
in
(* argument types 'Ts' could be added here, but they are also *)
(* added by 'collect_dtyp' automatically *)
- collect_dtyp (Datatype_Aux.DtRec index) acc
+ collect_dtyp (Datatype.DtRec index) acc
end
| NONE =>
(* not an inductive datatype, e.g. defined via "typedef" or *)
@@ -1853,7 +1853,7 @@
(* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
val _ =
if Library.exists (fn d =>
- case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps
+ case d of Datatype.DtTFree _ => false | _ => true) dtyps
then
raise REFUTE ("IDT_interpreter",
"datatype argument (for type "
@@ -1975,7 +1975,7 @@
(* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
val _ =
if Library.exists (fn d =>
- case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps
+ case d of Datatype.DtTFree _ => false | _ => true) dtyps
then
raise REFUTE ("IDT_constructor_interpreter",
"datatype argument (for type "
@@ -2035,7 +2035,7 @@
val typ_assoc = dtyps ~~ Ts'
(* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
val _ = if Library.exists (fn d =>
- case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps
+ case d of Datatype.DtTFree _ => false | _ => true) dtyps
then
raise REFUTE ("IDT_constructor_interpreter",
"datatype argument (for type "
@@ -2250,7 +2250,7 @@
(* 'DtTFree' *)
val _ =
if Library.exists (fn d =>
- case d of Datatype_Aux.DtTFree _ => false
+ case d of Datatype.DtTFree _ => false
| _ => true) dtyps
then
raise REFUTE ("IDT_recursion_interpreter",
@@ -2271,10 +2271,10 @@
(case AList.lookup op= acc d of
NONE =>
(case d of
- Datatype_Aux.DtTFree _ =>
+ Datatype.DtTFree _ =>
(* add the association, proceed *)
rec_typ_assoc ((d, T)::acc) xs
- | Datatype_Aux.DtType (s, ds) =>
+ | Datatype.DtType (s, ds) =>
let
val (s', Ts) = dest_Type T
in
@@ -2284,7 +2284,7 @@
raise REFUTE ("IDT_recursion_interpreter",
"DtType/Type mismatch")
end
- | Datatype_Aux.DtRec i =>
+ | Datatype.DtRec i =>
let
val (_, ds, _) = the (AList.lookup (op =) descr i)
val (_, Ts) = dest_Type T
@@ -2300,7 +2300,7 @@
raise REFUTE ("IDT_recursion_interpreter",
"different type associations for the same dtyp"))
val typ_assoc = filter
- (fn (Datatype_Aux.DtTFree _, _) => true | (_, _) => false)
+ (fn (Datatype.DtTFree _, _) => true | (_, _) => false)
(rec_typ_assoc []
(#2 (the (AList.lookup (op =) descr idt_index)) ~~ (snd o dest_Type) IDT))
(* sanity check: typ_assoc must associate types to the *)
@@ -2317,7 +2317,7 @@
val mc_intrs = map (fn (idx, (_, _, cs)) =>
let
val c_return_typ = typ_of_dtyp descr typ_assoc
- (Datatype_Aux.DtRec idx)
+ (Datatype.DtRec idx)
in
(idx, map (fn (cname, cargs) =>
(#1 o interpret ctxt (typs, []) {maxvars=0,
@@ -2371,7 +2371,7 @@
val _ = map (fn (idx, intrs) =>
let
val T = typ_of_dtyp descr typ_assoc
- (Datatype_Aux.DtRec idx)
+ (Datatype.DtRec idx)
in
if length intrs <> size_of_type ctxt (typs, []) T then
raise REFUTE ("IDT_recursion_interpreter",
@@ -2465,17 +2465,17 @@
(* takes the dtyp and interpretation of an element, *)
(* and computes the interpretation for the *)
(* corresponding recursive argument *)
- fun rec_intr (Datatype_Aux.DtRec i) (Leaf xs) =
+ fun rec_intr (Datatype.DtRec i) (Leaf xs) =
(* recursive argument is "rec_i params elem" *)
compute_array_entry i (find_index (fn x => x = True) xs)
- | rec_intr (Datatype_Aux.DtRec _) (Node _) =
+ | rec_intr (Datatype.DtRec _) (Node _) =
raise REFUTE ("IDT_recursion_interpreter",
"interpretation for IDT is a node")
- | rec_intr (Datatype_Aux.DtType ("fun", [dt1, dt2])) (Node xs) =
+ | rec_intr (Datatype.DtType ("fun", [dt1, dt2])) (Node xs) =
(* recursive argument is something like *)
(* "\<lambda>x::dt1. rec_? params (elem x)" *)
Node (map (rec_intr dt2) xs)
- | rec_intr (Datatype_Aux.DtType ("fun", [_, _])) (Leaf _) =
+ | rec_intr (Datatype.DtType ("fun", [_, _])) (Leaf _) =
raise REFUTE ("IDT_recursion_interpreter",
"interpretation for function dtyp is a leaf")
| rec_intr _ _ =
@@ -3024,7 +3024,7 @@
(* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
val _ =
if Library.exists (fn d =>
- case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps
+ case d of Datatype.DtTFree _ => false | _ => true) dtyps
then
raise REFUTE ("IDT_printer", "datatype argument (for type " ^
Syntax.string_of_typ ctxt (Type (s, Ts)) ^ ") is not a variable")