1 (* Title: HOL/Tools/datatype_rep_proofs.ML
3 Author: Stefan Berghofer, TU Muenchen
5 Definitional introduction of datatypes
6 Proof of characteristic theorems:
8 - injectivity of constructors
9 - distinctness of constructors
14 signature DATATYPE_REP_PROOFS =
16 val representation_proofs : bool -> DatatypeAux.datatype_info Symtab.table ->
17 string list -> DatatypeAux.descr list -> (string * sort) list ->
18 (string * mixfix) list -> (string * mixfix) list list -> attribute
19 -> theory -> (thm list list * thm list list * thm list list *
20 DatatypeAux.simproc_dist list * thm) * theory
23 structure DatatypeRepProofs : DATATYPE_REP_PROOFS =
28 val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma);
30 val collect_simp = rewrite_rule [mk_meta_eq mem_Collect_eq];
33 (** theory context references **)
35 val f_myinv_f = thm "f_myinv_f";
36 val myinv_f_f = thm "myinv_f_f";
39 fun exh_thm_of (dt_info : datatype_info Symtab.table) tname =
40 #exhaustion (the (Symtab.lookup dt_info tname));
42 (******************************************************************************)
44 fun representation_proofs flat_names (dt_info : datatype_info Symtab.table)
45 new_type_names descr sorts types_syntax constr_syntax case_names_induct thy =
47 val Datatype_thy = ThyInfo.the_theory "Datatype" thy;
48 val node_name = "Datatype.node";
49 val In0_name = "Datatype.In0";
50 val In1_name = "Datatype.In1";
51 val Scons_name = "Datatype.Scons";
52 val Leaf_name = "Datatype.Leaf";
53 val Numb_name = "Datatype.Numb";
54 val Lim_name = "Datatype.Lim";
55 val Suml_name = "Datatype.Suml";
56 val Sumr_name = "Datatype.Sumr";
58 val [In0_inject, In1_inject, Scons_inject, Leaf_inject,
59 In0_eq, In1_eq, In0_not_In1, In1_not_In0,
60 Lim_inject, Suml_inject, Sumr_inject] = map (get_thm Datatype_thy o Name)
61 ["In0_inject", "In1_inject", "Scons_inject", "Leaf_inject",
62 "In0_eq", "In1_eq", "In0_not_In1", "In1_not_In0",
63 "Lim_inject", "Suml_inject", "Sumr_inject"];
65 val descr' = List.concat descr;
67 val big_name = space_implode "_" new_type_names;
68 val thy1 = add_path flat_names big_name thy;
69 val big_rec_name = big_name ^ "_rep_set";
71 (if length descr' = 1 then [big_rec_name] else
72 (map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int)
73 (1 upto (length descr'))));
74 val rep_set_names = map (Sign.full_name (Theory.sign_of thy1)) rep_set_names';
76 val tyvars = map (fn (_, (_, Ts, _)) => map dest_DtTFree Ts) (hd descr);
77 val leafTs' = get_nonrec_types descr' sorts;
78 val branchTs = get_branching_types descr' sorts;
79 val branchT = if null branchTs then HOLogic.unitT
80 else fold_bal (fn (T, U) => Type ("+", [T, U])) branchTs;
81 val arities = get_arities descr' \ 0;
82 val unneeded_vars = hd tyvars \\ foldr add_typ_tfree_names [] (leafTs' @ branchTs);
83 val leafTs = leafTs' @ (map (fn n => TFree (n, (the o AList.lookup (op =) sorts) n)) unneeded_vars);
84 val recTs = get_rec_types descr' sorts;
85 val newTs = Library.take (length (hd descr), recTs);
86 val oldTs = Library.drop (length (hd descr), recTs);
87 val sumT = if null leafTs then HOLogic.unitT
88 else fold_bal (fn (T, U) => Type ("+", [T, U])) leafTs;
89 val Univ_elT = HOLogic.mk_setT (Type (node_name, [sumT, branchT]));
90 val UnivT = HOLogic.mk_setT Univ_elT;
91 val UnivT' = Univ_elT --> HOLogic.boolT;
92 val Collect = Const ("Collect", UnivT' --> UnivT);
94 val In0 = Const (In0_name, Univ_elT --> Univ_elT);
95 val In1 = Const (In1_name, Univ_elT --> Univ_elT);
96 val Leaf = Const (Leaf_name, sumT --> Univ_elT);
97 val Lim = Const (Lim_name, (branchT --> Univ_elT) --> Univ_elT);
99 (* make injections needed for embedding types in leaves *)
105 let val n2 = n div 2;
106 val Type (_, [T1, T2]) = T
109 Const ("Sum_Type.Inl", T1 --> T) $ (mk_inj' T1 n2 i)
111 Const ("Sum_Type.Inr", T2 --> T) $ (mk_inj' T2 (n - n2) (i - n2))
113 in mk_inj' sumT (length leafTs) (1 + find_index_eq T' leafTs)
116 (* make injections for constructors *)
118 fun mk_univ_inj ts = access_bal (fn t => In0 $ t, fn t => In1 $ t, if ts = [] then
119 Const ("arbitrary", Univ_elT)
121 foldr1 (HOLogic.mk_binop Scons_name) ts);
123 (* function spaces *)
125 fun mk_fun_inj T' x =
131 val Type (_, [T1, T2]) = T;
132 fun mkT U = (U --> Univ_elT) --> T --> Univ_elT
134 if i <= n2 then Const (Suml_name, mkT T1) $ mk_inj T1 n2 i
135 else Const (Sumr_name, mkT T2) $ mk_inj T2 (n - n2) (i - n2)
137 in mk_inj branchT (length branchTs) (1 + find_index_eq T' branchTs)
140 val mk_lim = foldr (fn (T, t) => Lim $ mk_fun_inj T (Abs ("x", T, t)));
142 (************** generate introduction rules for representing set **********)
144 val _ = message "Constructing representing sets ...";
146 (* make introduction rule for a single constructor *)
148 fun make_intr s n (i, (_, cargs)) =
150 fun mk_prem (dt, (j, prems, ts)) = (case strip_dtyp dt of
153 val Ts = map (typ_of_dtyp descr' sorts) dts;
155 app_bnds (mk_Free "x" (Ts ---> Univ_elT) j) (length Ts)
156 in (j + 1, list_all (map (pair "x") Ts,
158 (Free (List.nth (rep_set_names', k), UnivT') $ free_t)) :: prems,
159 mk_lim free_t Ts :: ts)
162 let val T = typ_of_dtyp descr' sorts dt
163 in (j + 1, prems, (Leaf $ mk_inj T (mk_Free "x" T j))::ts)
166 val (_, prems, ts) = foldr mk_prem (1, [], []) cargs;
167 val concl = HOLogic.mk_Trueprop
168 (Free (s, UnivT') $ mk_univ_inj ts n i)
169 in Logic.list_implies (prems, concl)
172 val intr_ts = List.concat (map (fn ((_, (_, _, constrs)), rep_set_name) =>
173 map (make_intr rep_set_name (length constrs))
174 ((1 upto (length constrs)) ~~ constrs)) (descr' ~~ rep_set_names'));
176 val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy2) =
177 setmp InductivePackage.quiet_mode (!quiet_mode)
178 (TheoryTarget.init NONE #>
179 InductivePackage.add_inductive_i false big_rec_name false true false
180 (map (fn s => (s, SOME UnivT', NoSyn)) rep_set_names') []
181 (map (fn x => (("", []), x)) intr_ts) [] #>
182 apsnd (ProofContext.theory_of o LocalTheory.exit)) thy1;
184 (********************************* typedef ********************************)
186 val (typedefs, thy3) = thy2 |>
187 parent_path flat_names |>
188 fold_map (fn ((((name, mx), tvs), c), name') =>
189 setmp TypedefPackage.quiet_mode true
190 (TypedefPackage.add_typedef_i false (SOME name') (name, tvs, mx)
191 (Collect $ Const (c, UnivT')) NONE
192 (rtac exI 1 THEN rtac CollectI 1 THEN
193 QUIET_BREADTH_FIRST (has_fewer_prems 1)
194 (resolve_tac rep_intrs 1))))
195 (types_syntax ~~ tyvars ~~
196 (Library.take (length newTs, rep_set_names)) ~~ new_type_names) ||>
197 add_path flat_names big_name;
199 (*********************** definition of constructors ***********************)
201 val big_rep_name = (space_implode "_" new_type_names) ^ "_Rep_";
202 val rep_names = map (curry op ^ "Rep_") new_type_names;
203 val rep_names' = map (fn i => big_rep_name ^ (string_of_int i))
204 (1 upto (length (List.concat (tl descr))));
205 val all_rep_names = map (Sign.intern_const (Theory.sign_of thy3)) rep_names @
206 map (Sign.full_name (Theory.sign_of thy3)) rep_names';
208 (* isomorphism declarations *)
210 val iso_decls = map (fn (T, s) => (s, T --> Univ_elT, NoSyn))
211 (oldTs ~~ rep_names');
213 (* constructor definitions *)
215 fun make_constr_def tname T n ((thy, defs, eqns, i), ((cname, cargs), (cname', mx))) =
217 fun constr_arg (dt, (j, l_args, r_args)) =
218 let val T = typ_of_dtyp descr' sorts dt;
219 val free_t = mk_Free "x" T j
220 in (case (strip_dtyp dt, strip_type T) of
221 ((_, DtRec m), (Us, U)) => (j + 1, free_t :: l_args, mk_lim
222 (Const (List.nth (all_rep_names, m), U --> Univ_elT) $
223 app_bnds free_t (length Us)) Us :: r_args)
224 | _ => (j + 1, free_t::l_args, (Leaf $ mk_inj T free_t)::r_args))
227 val (_, l_args, r_args) = foldr constr_arg (1, [], []) cargs;
228 val constrT = (map (typ_of_dtyp descr' sorts) cargs) ---> T;
229 val abs_name = Sign.intern_const (Theory.sign_of thy) ("Abs_" ^ tname);
230 val rep_name = Sign.intern_const (Theory.sign_of thy) ("Rep_" ^ tname);
231 val lhs = list_comb (Const (cname, constrT), l_args);
232 val rhs = mk_univ_inj r_args n i;
233 val def = equals T $ lhs $ (Const (abs_name, Univ_elT --> T) $ rhs);
234 val def_name = (Sign.base_name cname) ^ "_def";
235 val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
236 (Const (rep_name, T --> Univ_elT) $ lhs, rhs));
237 val ([def_thm], thy') =
239 |> Theory.add_consts_i [(cname', constrT, mx)]
240 |> (PureThy.add_defs_i false o map Thm.no_attributes) [(def_name, def)];
242 in (thy', defs @ [def_thm], eqns @ [eqn], i + 1) end;
244 (* constructor definitions for datatype *)
246 fun dt_constr_defs ((thy, defs, eqns, rep_congs, dist_lemmas),
247 ((((_, (_, _, constrs)), tname), T), constr_syntax)) =
249 val _ $ (_ $ (cong_f $ _) $ _) = concl_of arg_cong;
250 val sg = Theory.sign_of thy;
251 val rep_const = cterm_of sg
252 (Const (Sign.intern_const sg ("Rep_" ^ tname), T --> Univ_elT));
253 val cong' = standard (cterm_instantiate [(cterm_of sg cong_f, rep_const)] arg_cong);
254 val dist = standard (cterm_instantiate [(cterm_of sg distinct_f, rep_const)] distinct_lemma);
255 val (thy', defs', eqns', _) = Library.foldl ((make_constr_def tname T) (length constrs))
256 ((add_path flat_names tname thy, defs, [], 1), constrs ~~ constr_syntax)
258 (parent_path flat_names thy', defs', eqns @ [eqns'],
259 rep_congs @ [cong'], dist_lemmas @ [dist])
262 val (thy4, constr_defs, constr_rep_eqns, rep_congs, dist_lemmas) = Library.foldl dt_constr_defs
263 ((thy3 |> Theory.add_consts_i iso_decls |> parent_path flat_names, [], [], [], []),
264 hd descr ~~ new_type_names ~~ newTs ~~ constr_syntax);
266 (*********** isomorphisms for new types (introduced by typedef) ***********)
268 val _ = message "Proving isomorphism properties ...";
270 val newT_iso_axms = map (fn (_, td) =>
271 (collect_simp (#Abs_inverse td), #Rep_inverse td,
272 collect_simp (#Rep td))) typedefs;
274 val newT_iso_inj_thms = map (fn (_, td) =>
275 (collect_simp (#Abs_inject td) RS iffD1, #Rep_inject td RS iffD1)) typedefs;
277 (********* isomorphisms between existing types and "unfolded" types *******)
279 (*---------------------------------------------------------------------*)
280 (* isomorphisms are defined using primrec-combinators: *)
281 (* generate appropriate functions for instantiating primrec-combinator *)
283 (* e.g. dt_Rep_i = list_rec ... (%h t y. In1 (Scons (Leaf h) y)) *)
285 (* also generate characteristic equations for isomorphisms *)
287 (* e.g. dt_Rep_i (cons h t) = In1 (Scons (dt_Rep_j h) (dt_Rep_i t)) *)
288 (*---------------------------------------------------------------------*)
290 fun make_iso_def k ks n ((fs, eqns, i), (cname, cargs)) =
292 val argTs = map (typ_of_dtyp descr' sorts) cargs;
293 val T = List.nth (recTs, k);
294 val rep_name = List.nth (all_rep_names, k);
295 val rep_const = Const (rep_name, T --> Univ_elT);
296 val constr = Const (cname, argTs ---> T);
298 fun process_arg ks' ((i2, i2', ts, Ts), dt) =
300 val T' = typ_of_dtyp descr' sorts dt;
301 val (Us, U) = strip_type T'
302 in (case strip_dtyp dt of
303 (_, DtRec j) => if j mem ks' then
304 (i2 + 1, i2' + 1, ts @ [mk_lim (app_bnds
305 (mk_Free "y" (Us ---> Univ_elT) i2') (length Us)) Us],
306 Ts @ [Us ---> Univ_elT])
308 (i2 + 1, i2', ts @ [mk_lim
309 (Const (List.nth (all_rep_names, j), U --> Univ_elT) $
310 app_bnds (mk_Free "x" T' i2) (length Us)) Us], Ts)
311 | _ => (i2 + 1, i2', ts @ [Leaf $ mk_inj T' (mk_Free "x" T' i2)], Ts))
314 val (i2, i2', ts, Ts) = Library.foldl (process_arg ks) ((1, 1, [], []), cargs);
315 val xs = map (uncurry (mk_Free "x")) (argTs ~~ (1 upto (i2 - 1)));
316 val ys = map (uncurry (mk_Free "y")) (Ts ~~ (1 upto (i2' - 1)));
317 val f = list_abs_free (map dest_Free (xs @ ys), mk_univ_inj ts n i);
319 val (_, _, ts', _) = Library.foldl (process_arg []) ((1, 1, [], []), cargs);
320 val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
321 (rep_const $ list_comb (constr, xs), mk_univ_inj ts' n i))
323 in (fs @ [f], eqns @ [eqn], i + 1) end;
325 (* define isomorphisms for all mutually recursive datatypes in list ds *)
327 fun make_iso_defs (ds, (thy, char_thms)) =
330 val (_, (tname, _, _)) = hd ds;
331 val {rec_rewrites, rec_names, ...} = the (Symtab.lookup dt_info tname);
333 fun process_dt ((fs, eqns, isos), (k, (tname, _, constrs))) =
335 val (fs', eqns', _) = Library.foldl (make_iso_def k ks (length constrs))
336 ((fs, eqns, 1), constrs);
337 val iso = (List.nth (recTs, k), List.nth (all_rep_names, k))
338 in (fs', eqns', isos @ [iso]) end;
340 val (fs, eqns, isos) = Library.foldl process_dt (([], [], []), ds);
341 val fTs = map fastype_of fs;
342 val defs = map (fn (rec_name, (T, iso_name)) => ((Sign.base_name iso_name) ^ "_def",
343 equals (T --> Univ_elT) $ Const (iso_name, T --> Univ_elT) $
344 list_comb (Const (rec_name, fTs @ [T] ---> Univ_elT), fs))) (rec_names ~~ isos);
345 val (def_thms, thy') = (PureThy.add_defs_i false o map Thm.no_attributes) defs thy;
347 (* prove characteristic equations *)
349 val rewrites = def_thms @ (map mk_meta_eq rec_rewrites);
350 val char_thms' = map (fn eqn => Goal.prove_global thy' [] [] eqn
351 (fn _ => EVERY [rewrite_goals_tac rewrites, rtac refl 1])) eqns;
353 in (thy', char_thms' @ char_thms) end;
355 val (thy5, iso_char_thms) = foldr make_iso_defs
356 (add_path flat_names big_name thy4, []) (tl descr);
358 (* prove isomorphism properties *)
360 fun mk_funs_inv thm =
362 val {sign, prop, ...} = rep_thm thm;
363 val _ $ (_ $ ((S as Const (_, Type (_, [U, _]))) $ _ )) $
364 (_ $ (_ $ (r $ (a $ _)) $ _)) = Type.freeze prop;
365 val used = add_term_tfree_names (a, []);
369 val Ts = map (TFree o rpair HOLogic.typeS)
370 (Name.variant_list used (replicate i "'t"));
371 val f = Free ("f", Ts ---> U)
372 in Goal.prove_global sign [] [] (Logic.mk_implies
373 (HOLogic.mk_Trueprop (HOLogic.list_all
374 (map (pair "x") Ts, S $ app_bnds f i)),
375 HOLogic.mk_Trueprop (HOLogic.mk_eq (list_abs (map (pair "x") Ts,
376 r $ (a $ app_bnds f i)), f))))
377 (fn _ => EVERY [REPEAT (rtac ext 1), REPEAT (etac allE 1), rtac thm 1, atac 1])
379 in map (fn r => r RS subst) (thm :: map mk_thm arities) end;
381 (* prove inj dt_Rep_i and dt_Rep_i x : dt_rep_set_i *)
383 fun prove_iso_thms (ds, (inj_thms, elem_thms)) =
385 val (_, (tname, _, _)) = hd ds;
386 val {induction, ...} = the (Symtab.lookup dt_info tname);
388 fun mk_ind_concl (i, _) =
390 val T = List.nth (recTs, i);
391 val Rep_t = Const (List.nth (all_rep_names, i), T --> Univ_elT);
392 val rep_set_name = List.nth (rep_set_names, i)
393 in (HOLogic.all_const T $ Abs ("y", T, HOLogic.imp $
394 HOLogic.mk_eq (Rep_t $ mk_Free "x" T i, Rep_t $ Bound 0) $
395 HOLogic.mk_eq (mk_Free "x" T i, Bound 0)),
396 Const (rep_set_name, UnivT') $ (Rep_t $ mk_Free "x" T i))
399 val (ind_concl1, ind_concl2) = ListPair.unzip (map mk_ind_concl ds);
401 val rewrites = map mk_meta_eq iso_char_thms;
402 val inj_thms' = map snd newT_iso_inj_thms @
403 map (fn r => r RS injD) inj_thms;
405 val inj_thm = Goal.prove_global thy5 [] []
406 (HOLogic.mk_Trueprop (mk_conj ind_concl1)) (fn _ => EVERY
407 [(indtac induction THEN_ALL_NEW ObjectLogic.atomize_tac) 1,
409 [rtac allI 1, rtac impI 1,
410 exh_tac (exh_thm_of dt_info) 1,
413 rewrite_goals_tac rewrites,
414 REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
415 (eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1)
417 [REPEAT (eresolve_tac (Scons_inject ::
418 map make_elim [Leaf_inject, Inl_inject, Inr_inject]) 1),
419 REPEAT (cong_tac 1), rtac refl 1,
420 REPEAT (atac 1 ORELSE (EVERY
421 [REPEAT (rtac ext 1),
422 REPEAT (eresolve_tac (mp :: allE ::
423 map make_elim (Suml_inject :: Sumr_inject ::
424 Lim_inject :: fun_cong :: inj_thms')) 1),
427 val inj_thms'' = map (fn r => r RS datatype_injI)
428 (split_conj_thm inj_thm);
431 Goal.prove_global thy5 [] [] (HOLogic.mk_Trueprop (mk_conj ind_concl2))
433 EVERY [(indtac induction THEN_ALL_NEW ObjectLogic.atomize_tac) 1,
434 rewrite_goals_tac rewrites,
435 REPEAT ((resolve_tac rep_intrs THEN_ALL_NEW
436 ((REPEAT o etac allE) THEN' ares_tac elem_thms)) 1)]);
438 in (inj_thms'' @ inj_thms, elem_thms @ (split_conj_thm elem_thm))
441 val (iso_inj_thms_unfolded, iso_elem_thms) = foldr prove_iso_thms
442 ([], map #3 newT_iso_axms) (tl descr);
443 val iso_inj_thms = map snd newT_iso_inj_thms @
444 map (fn r => r RS injD) iso_inj_thms_unfolded;
446 (* prove dt_rep_set_i x --> x : range dt_Rep_i *)
448 fun mk_iso_t (((set_name, iso_name), i), T) =
449 let val isoT = T --> Univ_elT
451 (Const (set_name, UnivT') $ mk_Free "x" Univ_elT i) $
452 (if i < length newTs then Const ("True", HOLogic.boolT)
453 else HOLogic.mk_mem (mk_Free "x" Univ_elT i,
454 Const ("image", [isoT, HOLogic.mk_setT T] ---> UnivT) $
455 Const (iso_name, isoT) $ Const ("UNIV", HOLogic.mk_setT T)))
458 val iso_t = HOLogic.mk_Trueprop (mk_conj (map mk_iso_t
459 (rep_set_names ~~ all_rep_names ~~ (0 upto (length descr' - 1)) ~~ recTs)));
461 (* all the theorems are proved by one single simultaneous induction *)
463 val range_eqs = map (fn r => mk_meta_eq (r RS range_ex1_eq))
464 iso_inj_thms_unfolded;
466 val iso_thms = if length descr = 1 then [] else
467 Library.drop (length newTs, split_conj_thm
468 (Goal.prove_global thy5 [] [] iso_t (fn _ => EVERY
469 [(indtac rep_induct THEN_ALL_NEW ObjectLogic.atomize_tac) 1,
470 REPEAT (rtac TrueI 1),
471 rewrite_goals_tac (mk_meta_eq choice_eq ::
472 symmetric (mk_meta_eq expand_fun_eq) :: range_eqs),
473 rewrite_goals_tac (map symmetric range_eqs),
475 [REPEAT (eresolve_tac ([rangeE, ex1_implies_ex RS exE] @
476 List.concat (map (mk_funs_inv o #1) newT_iso_axms)) 1),
477 TRY (hyp_subst_tac 1),
478 rtac (sym RS range_eqI) 1,
479 resolve_tac iso_char_thms 1])])));
481 val Abs_inverse_thms' =
482 map #1 newT_iso_axms @
483 map2 (fn r_inj => fn r => f_myinv_f OF [r_inj, r RS mp])
484 iso_inj_thms_unfolded iso_thms;
486 val Abs_inverse_thms = List.concat (map mk_funs_inv Abs_inverse_thms');
488 (******************* freeness theorems for constructors *******************)
490 val _ = message "Proving freeness of constructors ...";
492 (* prove theorem Rep_i (Constr_j ...) = Inj_j ... *)
494 fun prove_constr_rep_thm eqn =
496 val inj_thms = map fst newT_iso_inj_thms;
497 val rewrites = o_def :: constr_defs @ (map (mk_meta_eq o #2) newT_iso_axms)
498 in Goal.prove_global thy5 [] [] eqn (fn _ => EVERY
499 [resolve_tac inj_thms 1,
500 rewrite_goals_tac rewrites,
502 resolve_tac rep_intrs 2,
503 REPEAT (resolve_tac iso_elem_thms 1)])
506 (*--------------------------------------------------------------*)
507 (* constr_rep_thms and rep_congs are used to prove distinctness *)
508 (* of constructors. *)
509 (*--------------------------------------------------------------*)
511 val constr_rep_thms = map (map prove_constr_rep_thm) constr_rep_eqns;
513 val dist_rewrites = map (fn (rep_thms, dist_lemma) =>
514 dist_lemma::(rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0]))
515 (constr_rep_thms ~~ dist_lemmas);
517 fun prove_distinct_thms (_, []) = []
518 | prove_distinct_thms (dist_rewrites', t::_::ts) =
520 val dist_thm = Goal.prove_global thy5 [] [] t (fn _ =>
521 EVERY [simp_tac (HOL_ss addsimps dist_rewrites') 1])
522 in dist_thm::(standard (dist_thm RS not_sym))::
523 (prove_distinct_thms (dist_rewrites', ts))
526 val distinct_thms = map prove_distinct_thms (dist_rewrites ~~
527 DatatypeProp.make_distincts new_type_names descr sorts thy5);
529 val simproc_dists = map (fn ((((_, (_, _, constrs)), rep_thms), congr), dists) =>
530 if length constrs < !DatatypeProp.dtK then FewConstrs dists
531 else ManyConstrs (congr, HOL_basic_ss addsimps rep_thms)) (hd descr ~~
532 constr_rep_thms ~~ rep_congs ~~ distinct_thms);
534 (* prove injectivity of constructors *)
536 fun prove_constr_inj_thm rep_thms t =
537 let val inj_thms = Scons_inject :: (map make_elim
539 [In0_inject, In1_inject, Leaf_inject, Inl_inject, Inr_inject,
540 Lim_inject, Suml_inject, Sumr_inject]))
541 in Goal.prove_global thy5 [] [] t (fn _ => EVERY
543 REPEAT (etac conjE 2), hyp_subst_tac 2, rtac refl 2,
544 dresolve_tac rep_congs 1, dtac box_equals 1,
545 REPEAT (resolve_tac rep_thms 1),
546 REPEAT (eresolve_tac inj_thms 1),
547 REPEAT (ares_tac [conjI] 1 ORELSE (EVERY [REPEAT (rtac ext 1),
548 REPEAT (eresolve_tac (make_elim fun_cong :: inj_thms) 1),
552 val constr_inject = map (fn (ts, thms) => map (prove_constr_inj_thm thms) ts)
553 ((DatatypeProp.make_injs descr sorts) ~~ constr_rep_thms);
555 val ((constr_inject', distinct_thms'), thy6) =
557 |> parent_path flat_names
558 |> store_thmss "inject" new_type_names constr_inject
559 ||>> store_thmss "distinct" new_type_names distinct_thms;
561 (*************************** induction theorem ****************************)
563 val _ = message "Proving induction rule for datatypes ...";
565 val Rep_inverse_thms = (map (fn (_, iso, _) => iso RS subst) newT_iso_axms) @
566 (map (fn r => r RS myinv_f_f RS subst) iso_inj_thms_unfolded);
567 val Rep_inverse_thms' = map (fn r => r RS myinv_f_f) iso_inj_thms_unfolded;
569 fun mk_indrule_lemma ((prems, concls), ((i, _), T)) =
571 val Rep_t = Const (List.nth (all_rep_names, i), T --> Univ_elT) $
574 val Abs_t = if i < length newTs then
575 Const (Sign.intern_const (Theory.sign_of thy6)
576 ("Abs_" ^ (List.nth (new_type_names, i))), Univ_elT --> T)
577 else Const ("Inductive.myinv", [T --> Univ_elT, Univ_elT] ---> T) $
578 Const (List.nth (all_rep_names, i), T --> Univ_elT)
580 in (prems @ [HOLogic.imp $
581 (Const (List.nth (rep_set_names, i), UnivT') $ Rep_t) $
582 (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))],
583 concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i])
586 val (indrule_lemma_prems, indrule_lemma_concls) =
587 Library.foldl mk_indrule_lemma (([], []), (descr' ~~ recTs));
589 val cert = cterm_of (Theory.sign_of thy6);
591 val indrule_lemma = Goal.prove_global thy6 [] []
593 (HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems),
594 HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY
595 [REPEAT (etac conjE 1),
597 [TRY (rtac conjI 1), resolve_tac Rep_inverse_thms 1,
598 etac mp 1, resolve_tac iso_elem_thms 1])]);
600 val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma)));
601 val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else
602 map (Free o apfst fst o dest_Var) Ps;
603 val indrule_lemma' = cterm_instantiate (map cert Ps ~~ map cert frees) indrule_lemma;
605 val dt_induct_prop = DatatypeProp.make_ind descr sorts;
606 val dt_induct = Goal.prove_global thy6 []
607 (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop)
609 [rtac indrule_lemma' 1,
610 (indtac rep_induct THEN_ALL_NEW ObjectLogic.atomize_tac) 1,
611 EVERY (map (fn (prem, r) => (EVERY
612 [REPEAT (eresolve_tac Abs_inverse_thms 1),
613 simp_tac (HOL_basic_ss addsimps ((symmetric r)::Rep_inverse_thms')) 1,
614 DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)]))
615 (prems ~~ (constr_defs @ (map mk_meta_eq iso_char_thms))))]);
617 val ([dt_induct'], thy7) =
619 |> Theory.add_path big_name
620 |> PureThy.add_thms [(("induct", dt_induct), [case_names_induct])]
621 ||> Theory.parent_path;
624 ((constr_inject', distinct_thms', dist_rewrites, simproc_dists, dt_induct'), thy7)