1 (* Title: HOL/Tools/datatype_package.ML
3 Author: Stefan Berghofer
4 Copyright 1998 TU Muenchen
6 Datatype package for Isabelle/HOL.
9 signature BASIC_DATATYPE_PACKAGE =
11 val induct_tac : string -> int -> tactic
12 val case_tac : string -> int -> tactic
13 val distinct_simproc : simproc
16 signature DATATYPE_PACKAGE =
18 include BASIC_DATATYPE_PACKAGE
19 val quiet_mode : bool ref
20 val add_datatype : bool -> string list -> (string list * bstring * mixfix *
21 (bstring * string list * mixfix) list) list -> theory -> theory *
22 {distinct : thm list list,
23 inject : thm list list,
24 exhaustion : thm list,
26 case_thms : thm list list,
27 split_thms : (thm * thm) list,
31 val add_datatype_i : bool -> string list -> (string list * bstring * mixfix *
32 (bstring * typ list * mixfix) list) list -> theory -> theory *
33 {distinct : thm list list,
34 inject : thm list list,
35 exhaustion : thm list,
37 case_thms : thm list list,
38 split_thms : (thm * thm) list,
42 val rep_datatype_i : string list option -> (thm * theory attribute list) list list ->
43 (thm * theory attribute list) list list -> (thm * theory attribute list) -> theory -> theory *
44 {distinct : thm list list,
45 inject : thm list list,
46 exhaustion : thm list,
48 case_thms : thm list list,
49 split_thms : (thm * thm) list,
53 val rep_datatype : string list option -> (xstring * Args.src list) list list ->
54 (xstring * Args.src list) list list -> xstring * Args.src list -> theory -> theory *
55 {distinct : thm list list,
56 inject : thm list list,
57 exhaustion : thm list,
59 case_thms : thm list list,
60 split_thms : (thm * thm) list,
64 val get_datatypes : theory -> DatatypeAux.datatype_info Symtab.table
65 val print_datatypes : theory -> unit
66 val datatype_info_sg : Sign.sg -> string -> DatatypeAux.datatype_info option
67 val datatype_info : theory -> string -> DatatypeAux.datatype_info option
68 val datatype_info_sg_err : Sign.sg -> string -> DatatypeAux.datatype_info
69 val datatype_info_err : theory -> string -> DatatypeAux.datatype_info
70 val constrs_of : theory -> string -> term list option
71 val constrs_of_sg : Sign.sg -> string -> term list option
72 val case_const_of : theory -> string -> term option
73 val setup: (theory -> theory) list
76 structure DatatypePackage : DATATYPE_PACKAGE =
81 val quiet_mode = quiet_mode;
84 (* data kind 'HOL/datatypes' *)
86 structure DatatypesArgs =
88 val name = "HOL/datatypes";
89 type T = datatype_info Symtab.table;
91 val empty = Symtab.empty;
94 val merge: T * T -> T = Symtab.merge (K true);
97 Pretty.writeln (Pretty.strs ("datatypes:" ::
98 map #1 (Sign.cond_extern_table sg Sign.typeK tab)));
101 structure DatatypesData = TheoryDataFun(DatatypesArgs);
102 val get_datatypes_sg = DatatypesData.get_sg;
103 val get_datatypes = DatatypesData.get;
104 val put_datatypes = DatatypesData.put;
105 val print_datatypes = DatatypesData.print;
108 (** theory information about datatypes **)
110 fun datatype_info_sg sg name = Symtab.lookup (get_datatypes_sg sg, name);
112 fun datatype_info_sg_err sg name = (case datatype_info_sg sg name of
114 | None => error ("Unknown datatype " ^ quote name));
116 val datatype_info = datatype_info_sg o Theory.sign_of;
118 fun datatype_info_err thy name = (case datatype_info thy name of
120 | None => error ("Unknown datatype " ^ quote name));
122 fun constrs_of_sg sg tname = (case datatype_info_sg sg tname of
123 Some {index, descr, ...} =>
124 let val (_, _, constrs) = the (assoc (descr, index))
125 in Some (map (fn (cname, _) => Const (cname, the (Sign.const_type sg cname))) constrs)
129 val constrs_of = constrs_of_sg o Theory.sign_of;
131 fun case_const_of thy tname = (case datatype_info thy tname of
132 Some {case_name, ...} => Some (Const (case_name, the (Sign.const_type
133 (Theory.sign_of thy) case_name)))
136 fun find_tname var Bi =
137 let val frees = map dest_Free (term_frees Bi)
138 val params = Logic.strip_params Bi;
139 in case assoc (frees @ params, var) of
140 None => error ("No such variable in subgoal: " ^ quote var)
141 | Some(Type (tn, _)) => tn
142 | _ => error ("Cannot determine type of " ^ quote var)
145 fun infer_tname state i aterm =
147 val sign = Thm.sign_of_thm state;
148 val (_, _, Bi, _) = Thm.dest_state (state, i)
149 val params = Logic.strip_params Bi; (*params of subgoal i*)
150 val params = rev (rename_wrt_term Bi params); (*as they are printed*)
151 val (types, sorts) = types_sorts state;
152 fun types' (a, ~1) = (case assoc (params, a) of None => types(a, ~1) | sm => sm)
153 | types' ixn = types ixn;
154 val (ct, _) = read_def_cterm (sign, types', sorts) [] false (aterm, TypeInfer.logicT);
155 in case #T (rep_cterm ct) of
157 | _ => error ("Cannot determine type of " ^ quote aterm)
160 (*Warn if the (induction) variable occurs Free among the premises, which
161 usually signals a mistake. But calls the tactic either way!*)
162 fun occs_in_prems tacf vars =
163 SUBGOAL (fn (Bi, i) =>
164 (if exists (fn Free (a, _) => a mem vars)
165 (foldr add_term_frees (#2 (strip_context Bi), []))
166 then warning "Induction variable occurs also among premises!"
171 (* generic induction tactic for datatypes *)
175 fun prep_var (Var (ixn, _), Some x) = Some (implode (tl (explode (Syntax.string_of_vname ixn))), x)
178 fun prep_inst (concl, xs) = (*exception LIST*)
179 let val vs = InductMethod.vars_of concl
180 in mapfilter prep_var (Library.drop (length vs - length xs, vs) ~~ xs) end;
184 fun gen_induct_tac (varss, opt_rule) i state =
186 val (_, _, Bi, _) = Thm.dest_state (state, i);
187 val {sign, ...} = Thm.rep_thm state;
188 val (rule, rule_name) =
190 Some r => (r, "Induction rule")
192 let val tn = find_tname (hd (mapfilter I (flat varss))) Bi
193 in (#induction (datatype_info_sg_err sign tn), "Induction rule for type " ^ tn) end);
195 val concls = InductMethod.concls_of rule;
196 val insts = flat (map prep_inst (concls ~~ varss)) handle LIST _ =>
197 error (rule_name ^ " has different numbers of variables");
198 in occs_in_prems (Tactic.res_inst_tac insts rule) (map #2 insts) i state end;
200 fun induct_tac s = gen_induct_tac (map (Library.single o Some) (Syntax.read_idents s), None);
205 (* generic case tactic for datatypes *)
207 fun case_inst_tac t rule i state =
209 val _ $ Var (ixn, _) $ _ = HOLogic.dest_Trueprop
210 (hd (Logic.strip_assums_hyp (hd (Thm.prems_of rule))));
211 val exh_vname = implode (tl (explode (Syntax.string_of_vname ixn)));
212 in Tactic.res_inst_tac [(exh_vname, t)] rule i state end;
214 fun gen_case_tac (t, Some rule) i state = case_inst_tac t rule i state
215 | gen_case_tac (t, None) i state =
216 let val tn = infer_tname state i t in
217 if tn = HOLogic.boolN then Tactic.res_inst_tac [("P", t)] case_split_thm i state
218 else case_inst_tac t (#exhaustion (datatype_info_sg_err (Thm.sign_of_thm state) tn)) i state
221 fun case_tac t = gen_case_tac (t, None);
225 (** Isar tactic emulations **)
229 val rule_spec = Scan.lift (Args.$$$ "rule" -- Args.$$$ ":");
230 val opt_rule = Scan.option (rule_spec |-- Attrib.local_thm);
232 val varss = Args.and_list (Scan.repeat (Scan.unless rule_spec (Scan.lift Args.name_dummy)));
234 fun tactic_syntax args tac =
235 #2 oo Method.syntax (Args.goal_spec HEADGOAL -- (args -- opt_rule) >>
236 (fn (quant, x) => Method.METHOD (fn facts => quant (Method.insert_tac facts THEN' tac x))));
240 val tactic_emulations =
241 [("induct_tac", tactic_syntax varss gen_induct_tac,
242 "induct_tac emulation (dynamic instantiation!)"),
243 ("case_tac", tactic_syntax (Scan.lift Args.name) gen_case_tac,
244 "case_tac emulation (dynamic instantiation!)")];
250 (** induct method setup **)
256 fun dt_recs (DtTFree _) = []
257 | dt_recs (DtType (_, dts)) = flat (map dt_recs dts)
258 | dt_recs (DtRec i) = [i];
260 fun dt_cases (descr: descr) (_, args, constrs) =
262 fun the_bname i = Sign.base_name (#1 (the (assoc (descr, i))));
263 val bnames = map the_bname (distinct (flat (map dt_recs args)));
264 in map (fn (c, _) => space_implode "_" (Sign.base_name c :: bnames)) constrs end;
267 fun induct_cases descr =
268 DatatypeProp.indexify_names (flat (map (dt_cases descr) (map #2 descr)));
270 fun exhaust_cases descr i = dt_cases descr (the (assoc (descr, i)));
274 fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr);
276 fun mk_case_names_exhausts descr new =
277 map (RuleCases.case_names o exhaust_cases descr o #1)
278 (filter (fn ((_, (name, _, _))) => name mem_string new) descr);
283 (* add_cases_induct *)
285 fun add_cases_induct infos =
287 fun proj _ 1 thm = thm
289 (if i + 1 < n then (fn th => th RS conjunct1) else I)
290 (Library.funpow i (fn th => th RS conjunct2) thm)
292 |> RuleCases.name (RuleCases.get thm);
294 fun add (ths, (name, {index, descr, induction, exhaustion, ...}: datatype_info)) =
295 (("", proj index (length descr) induction), [InductMethod.induct_type_global name]) ::
296 (("", exhaustion), [InductMethod.cases_type_global name]) :: ths;
297 in #1 o PureThy.add_thms (foldl add ([], infos)) end;
301 (**** simplification procedure for showing distinctness of constructors ****)
303 fun stripT (i, Type ("fun", [_, T])) = stripT (i + 1, T)
306 fun stripC (i, f $ x) = stripC (i + 1, f)
309 val distinctN = "constr_distinct";
311 exception ConstrDistinct of term;
313 fun distinct_proc sg _ (t as Const ("op =", _) $ t1 $ t2) =
314 (case (stripC (0, t1), stripC (0, t2)) of
315 ((i, Const (cname1, T1)), (j, Const (cname2, T2))) =>
316 (case (stripT (0, T1), stripT (0, T2)) of
317 ((i', Type (tname1, _)), (j', Type (tname2, _))) =>
318 if tname1 = tname2 andalso not (cname1 = cname2) andalso i = i' andalso j = j' then
319 (case (constrs_of_sg sg tname1) of
320 Some constrs => let val cnames = map (fst o dest_Const) constrs
321 in if cname1 mem cnames andalso cname2 mem cnames then
322 let val eq_t = Logic.mk_equals (t, Const ("False", HOLogic.boolT));
323 val eq_ct = cterm_of sg eq_t;
324 val Datatype_thy = theory "Datatype";
325 val [In0_inject, In1_inject, In0_not_In1, In1_not_In0] =
326 map (get_thm Datatype_thy)
327 ["In0_inject", "In1_inject", "In0_not_In1", "In1_not_In0"]
328 in (case (#distinct (datatype_info_sg_err sg tname1)) of
329 QuickAndDirty => Some (Thm.invoke_oracle
330 Datatype_thy distinctN (sg, ConstrDistinct eq_t))
331 | FewConstrs thms => Some (prove_goalw_cterm [] eq_ct (K
332 [rtac eq_reflection 1, rtac iffI 1, rtac notE 1,
333 atac 2, resolve_tac thms 1, etac FalseE 1]))
334 | ManyConstrs (thm, ss) => Some (prove_goalw_cterm [] eq_ct (K
335 [rtac eq_reflection 1, rtac iffI 1, dtac thm 1,
337 REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
338 eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1,
347 | distinct_proc sg _ _ = None;
349 val distinct_pats = [Thm.read_cterm (Theory.sign_of HOL.thy) ("s = t", HOLogic.termT)];
350 val distinct_simproc = mk_simproc distinctN distinct_pats distinct_proc;
352 val dist_ss = HOL_ss addsimprocs [distinct_simproc];
355 [Theory.add_oracle (distinctN, fn (_, ConstrDistinct t) => t),
356 fn thy => (simpset_ref_of thy := simpset_of thy addsimprocs [distinct_simproc]; thy)];
361 fun read_typ sign ((Ts, sorts), str) =
363 val T = Type.no_tvars (Sign.read_typ (sign, (curry assoc)
364 (map (apfst (rpair ~1)) sorts)) str) handle TYPE (msg, _, _) => error msg
365 in (Ts @ [T], add_typ_tfrees (T, sorts)) end;
367 fun cert_typ sign ((Ts, sorts), raw_T) =
369 val T = Type.no_tvars (Sign.certify_typ sign raw_T) handle
370 TYPE (msg, _, _) => error msg;
371 val sorts' = add_typ_tfrees (T, sorts)
373 case duplicates (map fst sorts') of
375 | dups => error ("Inconsistent sort constraints for " ^ commas dups))
379 (**** make datatype info ****)
381 fun make_dt_info descr induct reccomb_names rec_thms
382 ((((((((i, (_, (tname, _, _))), case_name), case_thms),
383 exhaustion_thm), distinct_thm), inject), nchotomy), case_cong) = (tname,
386 rec_names = reccomb_names,
387 rec_rewrites = rec_thms,
388 case_name = case_name,
389 case_rewrites = case_thms,
391 exhaustion = exhaustion_thm,
392 distinct = distinct_thm,
395 case_cong = case_cong});
398 (********************* axiomatic introduction of datatypes ********************)
400 fun add_and_get_axioms_atts label tnames attss ts thy =
401 foldr (fn (((tname, atts), t), (thy', axs)) =>
403 val (thy'', [ax]) = thy' |>
404 Theory.add_path tname |>
405 PureThy.add_axioms_i [((label, t), atts)];
406 in (Theory.parent_path thy'', ax::axs)
407 end) (tnames ~~ attss ~~ ts, (thy, []));
409 fun add_and_get_axioms label tnames =
410 add_and_get_axioms_atts label tnames (replicate (length tnames) []);
412 fun add_and_get_axiomss label tnames tss thy =
413 foldr (fn ((tname, ts), (thy', axss)) =>
415 val (thy'', [axs]) = thy' |>
416 Theory.add_path tname |>
417 PureThy.add_axiomss_i [((label, ts), [])];
418 in (Theory.parent_path thy'', axs::axss)
419 end) (tnames ~~ tss, (thy, []));
421 fun add_datatype_axm flat_names new_type_names descr sorts types_syntax constr_syntax dt_info
422 case_names_induct case_names_exhausts thy =
424 val descr' = flat descr;
425 val recTs = get_rec_types descr' sorts;
426 val used = foldr add_typ_tfree_names (recTs, []);
427 val newTs = take (length (hd descr), recTs);
429 val no_size = exists (fn (_, (_, _, constrs)) => exists (fn (_, cargs) => exists
430 (fn (DtType ("fun", [_, DtRec _])) => true | _ => false) cargs) constrs) descr';
432 (**** declare new types and constants ****)
434 val tyvars = map (fn (_, (_, Ts, _)) => map dest_DtTFree Ts) (hd descr);
436 val constr_decls = map (fn (((_, (_, _, constrs)), T), constr_syntax') =>
437 map (fn ((_, cargs), (cname, mx)) =>
438 (cname, map (typ_of_dtyp descr' sorts) cargs ---> T, mx))
439 (constrs ~~ constr_syntax')) ((hd descr) ~~ newTs ~~ constr_syntax);
441 val rec_result_Ts = map TFree (variantlist (replicate (length descr') "'t", used) ~~
442 replicate (length descr') HOLogic.termS);
444 val reccomb_fn_Ts = flat (map (fn (i, (_, _, constrs)) =>
445 map (fn (_, cargs) =>
447 val Ts = map (typ_of_dtyp descr' sorts) cargs;
448 val recs = filter (is_rec_type o fst) (cargs ~~ Ts);
450 fun mk_argT (DtRec k, _) = nth_elem (k, rec_result_Ts)
451 | mk_argT (DtType ("fun", [_, DtRec k]), Type ("fun", [T, _])) =
452 T --> nth_elem (k, rec_result_Ts);
454 val argTs = Ts @ map mk_argT recs
455 in argTs ---> nth_elem (i, rec_result_Ts)
456 end) constrs) descr');
458 val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
459 val reccomb_names = if length descr' = 1 then [big_reccomb_name] else
460 (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
461 (1 upto (length descr')));
463 val big_size_name = space_implode "_" new_type_names ^ "_size";
464 val size_names = if length (flat (tl descr)) = 1 then [big_size_name] else
465 map (fn i => big_size_name ^ "_" ^ string_of_int i)
466 (1 upto length (flat (tl descr)));
468 val freeT = TFree (variant used "'t", HOLogic.termS);
469 val case_fn_Ts = map (fn (i, (_, _, constrs)) =>
470 map (fn (_, cargs) =>
471 let val Ts = map (typ_of_dtyp descr' sorts) cargs
472 in Ts ---> freeT end) constrs) (hd descr);
474 val case_names = map (fn s => (s ^ "_case")) new_type_names;
480 curry (foldr (fn (((name, mx), tvs), thy') => thy' |>
481 TypedefPackage.add_typedecls [(name, tvs, mx)]))
482 (types_syntax ~~ tyvars) |>
483 add_path flat_names (space_implode "_" new_type_names) |>
485 (** primrec combinators **)
487 Theory.add_consts_i (map (fn ((name, T), T') =>
488 (name, reccomb_fn_Ts @ [T] ---> T', NoSyn))
489 (reccomb_names ~~ recTs ~~ rec_result_Ts)) |>
491 (** case combinators **)
493 Theory.add_consts_i (map (fn ((name, T), Ts) =>
494 (name, Ts @ [T] ---> freeT, NoSyn))
495 (case_names ~~ newTs ~~ case_fn_Ts)) |>
496 Theory.add_trrules_i (DatatypeProp.make_case_trrules new_type_names descr);
498 val reccomb_names' = map (Sign.intern_const (Theory.sign_of thy2')) reccomb_names;
499 val case_names' = map (Sign.intern_const (Theory.sign_of thy2')) case_names;
503 (** size functions **)
505 (if no_size then I else Theory.add_consts_i (map (fn (s, T) =>
506 (Sign.base_name s, T --> HOLogic.natT, NoSyn))
507 (size_names ~~ drop (length (hd descr), recTs)))) |>
511 parent_path flat_names |>
512 curry (foldr (fn (((((_, (_, _, constrs)), T), tname),
513 constr_syntax'), thy') => thy' |>
514 add_path flat_names tname |>
515 Theory.add_consts_i (map (fn ((_, cargs), (cname, mx)) =>
516 (cname, map (typ_of_dtyp descr' sorts) cargs ---> T, mx))
517 (constrs ~~ constr_syntax')) |>
518 parent_path flat_names))
519 (hd descr ~~ newTs ~~ new_type_names ~~ constr_syntax);
521 (**** introduction of axioms ****)
523 val rec_axs = DatatypeProp.make_primrecs new_type_names descr sorts thy2;
524 val size_axs = if no_size then [] else DatatypeProp.make_size new_type_names descr sorts thy2;
526 val (thy3, (([induct], [rec_thms]), inject)) =
528 Theory.add_path (space_implode "_" new_type_names) |>
529 PureThy.add_axioms_i [(("induct", DatatypeProp.make_ind descr sorts), [case_names_induct])] |>>>
530 PureThy.add_axiomss_i [(("recs", rec_axs), [])] |>>
531 (if no_size then I else #1 o PureThy.add_axiomss_i [(("size", size_axs), [])]) |>>
532 Theory.parent_path |>>>
533 add_and_get_axiomss "inject" new_type_names
534 (DatatypeProp.make_injs descr sorts);
535 val size_thms = if no_size then [] else get_thms thy3 "size";
536 val (thy4, distinct) = add_and_get_axiomss "distinct" new_type_names
537 (DatatypeProp.make_distincts new_type_names descr sorts thy3) thy3;
539 val exhaust_ts = DatatypeProp.make_casedists descr sorts;
540 val (thy5, exhaustion) = add_and_get_axioms_atts "exhaust" new_type_names
541 (map Library.single case_names_exhausts) exhaust_ts thy4;
542 val (thy6, case_thms) = add_and_get_axiomss "cases" new_type_names
543 (DatatypeProp.make_cases new_type_names descr sorts thy5) thy5;
544 val (split_ts, split_asm_ts) = ListPair.unzip
545 (DatatypeProp.make_splits new_type_names descr sorts thy6);
546 val (thy7, split) = add_and_get_axioms "split" new_type_names split_ts thy6;
547 val (thy8, split_asm) = add_and_get_axioms "split_asm" new_type_names
549 val (thy9, nchotomys) = add_and_get_axioms "nchotomy" new_type_names
550 (DatatypeProp.make_nchotomys descr sorts) thy8;
551 val (thy10, case_congs) = add_and_get_axioms "case_cong" new_type_names
552 (DatatypeProp.make_case_congs new_type_names descr sorts thy9) thy9;
553 val (thy11, weak_case_congs) = add_and_get_axioms "weak_case_cong" new_type_names
554 (DatatypeProp.make_weak_case_congs new_type_names descr sorts thy10) thy10;
556 val dt_infos = map (make_dt_info descr' induct reccomb_names' rec_thms)
557 ((0 upto length (hd descr) - 1) ~~ (hd descr) ~~ case_names' ~~ case_thms ~~
558 exhaustion ~~ replicate (length (hd descr)) QuickAndDirty ~~ inject ~~
559 nchotomys ~~ case_congs);
561 val simps = flat (distinct @ inject @ case_thms) @ size_thms @ rec_thms;
564 Theory.add_path (space_implode "_" new_type_names) |>
565 (#1 o PureThy.add_thmss [(("simps", simps), []),
566 (("", flat case_thms @ size_thms @ rec_thms), [Simplifier.simp_add_global]),
567 (("", flat (inject @ distinct)), [iff_add_global]),
568 (("", weak_case_congs), [cong_add_global])]) |>
569 put_datatypes (foldr Symtab.update (dt_infos, dt_info)) |>
570 add_cases_induct dt_infos |>
575 {distinct = distinct,
577 exhaustion = exhaustion,
579 case_thms = case_thms,
580 split_thms = split ~~ split_asm,
587 (******************* definitional introduction of datatypes *******************)
589 fun add_datatype_def flat_names new_type_names descr sorts types_syntax constr_syntax dt_info
590 case_names_induct case_names_exhausts thy =
592 val _ = message ("Proofs for datatype(s) " ^ commas_quote new_type_names);
594 val (thy2, inject, distinct, dist_rewrites, simproc_dists, induct) = thy |>
595 DatatypeRepProofs.representation_proofs flat_names dt_info new_type_names descr sorts
596 types_syntax constr_syntax case_names_induct;
598 val (thy3, casedist_thms) = DatatypeAbsProofs.prove_casedist_thms new_type_names descr
599 sorts induct case_names_exhausts thy2;
600 val (thy4, (reccomb_names, rec_thms)) = DatatypeAbsProofs.prove_primrec_thms
601 flat_names new_type_names descr sorts dt_info inject dist_rewrites dist_ss induct thy3;
602 val (thy6, (case_thms, case_names)) = DatatypeAbsProofs.prove_case_thms
603 flat_names new_type_names descr sorts reccomb_names rec_thms thy4;
604 val (thy7, split_thms) = DatatypeAbsProofs.prove_split_thms new_type_names
605 descr sorts inject dist_rewrites casedist_thms case_thms thy6;
606 val (thy8, nchotomys) = DatatypeAbsProofs.prove_nchotomys new_type_names
607 descr sorts casedist_thms thy7;
608 val (thy9, case_congs) = DatatypeAbsProofs.prove_case_congs new_type_names
609 descr sorts nchotomys case_thms thy8;
610 val (thy10, weak_case_congs) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
612 val (thy11, size_thms) = DatatypeAbsProofs.prove_size_thms flat_names new_type_names
613 descr sorts reccomb_names rec_thms thy10;
615 val dt_infos = map (make_dt_info (flat descr) induct reccomb_names rec_thms)
616 ((0 upto length (hd descr) - 1) ~~ (hd descr) ~~ case_names ~~ case_thms ~~
617 casedist_thms ~~ simproc_dists ~~ inject ~~ nchotomys ~~ case_congs);
619 val simps = flat (distinct @ inject @ case_thms) @ size_thms @ rec_thms;
622 Theory.add_path (space_implode "_" new_type_names) |>
623 (#1 o PureThy.add_thmss [(("simps", simps), []),
624 (("", flat case_thms @ size_thms @ rec_thms), [Simplifier.simp_add_global]),
625 (("", flat (inject @ distinct)), [iff_add_global]),
626 (("", weak_case_congs), [Simplifier.change_global_ss (op addcongs)])]) |>
627 put_datatypes (foldr Symtab.update (dt_infos, dt_info)) |>
628 add_cases_induct dt_infos |>
633 {distinct = distinct,
635 exhaustion = casedist_thms,
637 case_thms = case_thms,
638 split_thms = split_thms,
645 (*********************** declare existing type as datatype *********************)
647 fun gen_rep_datatype apply_theorems alt_names raw_distinct raw_inject raw_induction thy0 =
649 fun app_thmss srcs thy = foldl_map (fn (thy, x) => apply_theorems x thy) (thy, srcs);
650 fun app_thm src thy = apsnd Library.hd (apply_theorems [src] thy);
652 val (((thy1, induction), inject), distinct) = thy0
653 |> app_thmss raw_distinct
654 |> apfst (app_thmss raw_inject)
655 |> apfst (apfst (app_thm raw_induction));
656 val sign = Theory.sign_of thy1;
658 val induction' = freezeT induction;
660 fun err t = error ("Ill-formed predicate in induction rule: " ^
661 Sign.string_of_term sign t);
663 fun get_typ (t as _ $ Var (_, Type (tname, Ts))) =
664 ((tname, map dest_TFree Ts) handle TERM _ => err t)
667 val dtnames = map get_typ (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induction')));
668 val new_type_names = if_none alt_names (map fst dtnames);
670 fun get_constr t = (case Logic.strip_assums_concl t of
671 _ $ (_ $ t') => (case head_of t' of
672 Const (cname, cT) => (case strip_type cT of
673 (Ts, Type (tname, _)) => (tname, (cname, map (dtyp_of_typ dtnames) Ts))
678 fun make_dt_spec [] _ _ = []
679 | make_dt_spec ((tname, tvs)::dtnames') i constrs =
680 let val (constrs', constrs'') = take_prefix (equal tname o fst) constrs
681 in (i, (tname, map DtTFree tvs, map snd constrs'))::
682 (make_dt_spec dtnames' (i + 1) constrs'')
685 val descr = make_dt_spec dtnames 0 (map get_constr (prems_of induction'));
686 val sorts = add_term_tfrees (concl_of induction', []);
687 val dt_info = get_datatypes thy1;
689 val case_names_induct = mk_case_names_induct descr;
690 val case_names_exhausts = mk_case_names_exhausts descr (map #1 dtnames);
693 val _ = message ("Proofs for datatype(s) " ^ commas_quote new_type_names);
695 val (thy2, casedist_thms) = thy1 |>
696 DatatypeAbsProofs.prove_casedist_thms new_type_names [descr] sorts induction
698 val (thy3, (reccomb_names, rec_thms)) = DatatypeAbsProofs.prove_primrec_thms
699 false new_type_names [descr] sorts dt_info inject distinct dist_ss induction thy2;
700 val (thy4, (case_thms, case_names)) = DatatypeAbsProofs.prove_case_thms false
701 new_type_names [descr] sorts reccomb_names rec_thms thy3;
702 val (thy5, split_thms) = DatatypeAbsProofs.prove_split_thms
703 new_type_names [descr] sorts inject distinct casedist_thms case_thms thy4;
704 val (thy6, nchotomys) = DatatypeAbsProofs.prove_nchotomys new_type_names
705 [descr] sorts casedist_thms thy5;
706 val (thy7, case_congs) = DatatypeAbsProofs.prove_case_congs new_type_names
707 [descr] sorts nchotomys case_thms thy6;
708 val (thy8, weak_case_congs) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
710 val (thy9, size_thms) =
711 if exists (equal "Arith") (Sign.stamp_names_of (Theory.sign_of thy8)) then
712 DatatypeAbsProofs.prove_size_thms false new_type_names
713 [descr] sorts reccomb_names rec_thms thy8
716 val dt_infos = map (make_dt_info descr induction reccomb_names rec_thms)
717 ((0 upto length descr - 1) ~~ descr ~~ case_names ~~ case_thms ~~
718 casedist_thms ~~ map FewConstrs distinct ~~ inject ~~ nchotomys ~~ case_congs);
720 val simps = flat (distinct @ inject @ case_thms) @ size_thms @ rec_thms;
722 val (thy10, [induction']) = thy9 |>
723 (#1 o store_thmss "inject" new_type_names inject) |>
724 (#1 o store_thmss "distinct" new_type_names distinct) |>
725 Theory.add_path (space_implode "_" new_type_names) |>
726 PureThy.add_thms [(("induct", induction), [case_names_induct])] |>>
727 (#1 o PureThy.add_thmss [(("simps", simps), []),
728 (("", flat case_thms @ size_thms @ rec_thms), [Simplifier.simp_add_global]),
729 (("", flat (inject @ distinct)), [iff_add_global]),
730 (("", weak_case_congs), [Simplifier.change_global_ss (op addcongs)])]) |>>
731 put_datatypes (foldr Symtab.update (dt_infos, dt_info)) |>>
732 add_cases_induct dt_infos |>>
737 {distinct = distinct,
739 exhaustion = casedist_thms,
741 case_thms = case_thms,
742 split_thms = split_thms,
743 induction = induction',
748 val rep_datatype = gen_rep_datatype IsarThy.apply_theorems;
749 val rep_datatype_i = gen_rep_datatype IsarThy.apply_theorems_i;
752 (******************************** add datatype ********************************)
754 fun gen_add_datatype prep_typ flat_names new_type_names dts thy =
756 val _ = Theory.requires thy "Datatype" "datatype definitions";
758 (* this theory is used just for parsing *)
762 Theory.add_types (map (fn (tvs, tname, mx, _) =>
763 (tname, length tvs, mx)) dts);
765 val sign = Theory.sign_of tmp_thy;
767 val (tyvars, _, _, _)::_ = dts;
768 val (new_dts, types_syntax) = ListPair.unzip (map (fn (tvs, tname, mx, _) =>
769 let val full_tname = Sign.full_name sign (Syntax.type_name tname mx)
770 in (case duplicates tvs of
771 [] => if eq_set (tyvars, tvs) then ((full_tname, tvs), (tname, mx))
772 else error ("Mutually recursive datatypes must have same type parameters")
773 | dups => error ("Duplicate parameter(s) for datatype " ^ full_tname ^
774 " : " ^ commas dups))
777 val _ = (case duplicates (map fst new_dts) @ duplicates new_type_names of
778 [] => () | dups => error ("Duplicate datatypes: " ^ commas dups));
780 fun prep_dt_spec ((dts', constr_syntax, sorts, i), (tvs, tname, mx, constrs)) =
782 fun prep_constr ((constrs, constr_syntax', sorts'), (cname, cargs, mx')) =
784 val (cargs', sorts'') = foldl (prep_typ sign) (([], sorts'), cargs);
785 val _ = (case foldr add_typ_tfree_names (cargs', []) \\ tvs of
787 | vs => error ("Extra type variables on rhs: " ^ commas vs))
788 in (constrs @ [((if flat_names then Sign.full_name sign else
789 Sign.full_name_path sign tname) (Syntax.const_name cname mx'),
790 map (dtyp_of_typ new_dts) cargs')],
791 constr_syntax' @ [(cname, mx')], sorts'')
793 error ("The error above occured in constructor " ^ cname ^
794 " of datatype " ^ tname);
796 val (constrs', constr_syntax', sorts') =
797 foldl prep_constr (([], [], sorts), constrs)
800 case duplicates (map fst constrs') of
802 (dts' @ [(i, (Sign.full_name sign (Syntax.type_name tname mx),
803 map DtTFree tvs, constrs'))],
804 constr_syntax @ [constr_syntax'], sorts', i + 1)
805 | dups => error ("Duplicate constructors " ^ commas dups ^
806 " in datatype " ^ tname)
809 val (dts', constr_syntax, sorts', i) = foldl prep_dt_spec (([], [], [], 0), dts);
810 val sorts = sorts' @ (map (rpair (Sign.defaultS sign)) (tyvars \\ map fst sorts'));
811 val dt_info = get_datatypes thy;
812 val (descr, _) = unfold_datatypes sign dts' sorts dt_info dts' i;
813 val _ = check_nonempty descr;
815 val descr' = flat descr;
816 val case_names_induct = mk_case_names_induct descr';
817 val case_names_exhausts = mk_case_names_exhausts descr' (map #1 new_dts);
819 (if (!quick_and_dirty) then add_datatype_axm else add_datatype_def)
820 flat_names new_type_names descr sorts types_syntax constr_syntax dt_info
821 case_names_induct case_names_exhausts thy
824 val add_datatype_i = gen_add_datatype cert_typ;
825 val add_datatype = gen_add_datatype read_typ;
828 (** package setup **)
832 val setup = [DatatypesData.init, Method.add_methods tactic_emulations] @ simproc_setup;
837 local structure P = OuterParse and K = OuterSyntax.Keyword in
840 Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.name -- P.opt_infix --
841 (P.$$$ "=" |-- P.enum1 "|" (P.name -- Scan.repeat P.typ -- P.opt_mixfix --| P.marg_comment));
843 fun mk_datatype args =
845 val names = map (fn ((((None, _), t), _), _) => t | ((((Some t, _), _), _), _) => t) args;
846 val specs = map (fn ((((_, vs), t), mx), cons) => (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args;
847 in #1 o add_datatype false names specs end;
850 OuterSyntax.command "datatype" "define inductive datatypes" K.thy_decl
851 (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype));
854 val rep_datatype_decl =
855 Scan.option (Scan.repeat1 P.name) --
856 Scan.optional (P.$$$ "distinct" |-- P.!!! (P.and_list1 P.xthms1)) [] --
857 Scan.optional (P.$$$ "inject" |-- P.!!! (P.and_list1 P.xthms1)) [] --
858 (P.$$$ "induction" |-- P.!!! P.xthm);
860 fun mk_rep_datatype (((opt_ts, dss), iss), ind) = #1 o rep_datatype opt_ts dss iss ind;
863 OuterSyntax.command "rep_datatype" "represent existing types inductively" K.thy_decl
864 (rep_datatype_decl >> (Toplevel.theory o mk_rep_datatype));
867 val _ = OuterSyntax.add_keywords ["distinct", "inject", "induction"];
868 val _ = OuterSyntax.add_parsers [datatypeP, rep_datatypeP];
875 structure BasicDatatypePackage: BASIC_DATATYPE_PACKAGE = DatatypePackage;
876 open BasicDatatypePackage;