src/HOL/Codatatype/Tools/bnf_wrap.ML
author blanchet
Tue, 18 Sep 2012 11:42:22 +0200
changeset 50453 5bc80d96241e
parent 50452 c139da00fb4a
child 50473 9321a9465021
permissions -rw-r--r--
group "simps" together
     1 (*  Title:      HOL/Codatatype/Tools/bnf_wrap.ML
     2     Author:     Jasmin Blanchette, TU Muenchen
     3     Copyright   2012
     4 
     5 Wrapping existing datatypes.
     6 *)
     7 
     8 signature BNF_WRAP =
     9 sig
    10   val mk_half_pairss: 'a list -> ('a * 'a) list list
    11   val mk_ctr: typ list -> term -> term
    12   val base_name_of_ctr: term -> string
    13   val wrap_datatype: ({prems: thm list, context: Proof.context} -> tactic) list list ->
    14     ((bool * term list) * term) *
    15       (binding list * (binding list list * (binding * term) list list)) -> local_theory ->
    16     (term list list * thm list * thm list * thm list * thm list * thm list list) * local_theory
    17   val parse_wrap_options: bool parser
    18   val parse_bound_term: (binding * string) parser
    19 end;
    20 
    21 structure BNF_Wrap : BNF_WRAP =
    22 struct
    23 
    24 open BNF_Util
    25 open BNF_Wrap_Tactics
    26 
    27 val isN = "is_";
    28 val unN = "un_";
    29 fun mk_unN 1 1 suf = unN ^ suf
    30   | mk_unN _ l suf = unN ^ suf ^ string_of_int l;
    31 
    32 val case_congN = "case_cong";
    33 val case_eqN = "case_eq";
    34 val casesN = "cases";
    35 val collapseN = "collapse";
    36 val disc_excludeN = "disc_exclude";
    37 val disc_exhaustN = "disc_exhaust";
    38 val discsN = "discs";
    39 val distinctN = "distinct";
    40 val exhaustN = "exhaust";
    41 val injectN = "inject";
    42 val nchotomyN = "nchotomy";
    43 val selsN = "sels";
    44 val splitN = "split";
    45 val split_asmN = "split_asm";
    46 val weak_case_cong_thmsN = "weak_case_cong";
    47 
    48 val std_binding = @{binding _};
    49 
    50 val induct_simp_attrs = @{attributes [induct_simp]};
    51 val cong_attrs = @{attributes [cong]};
    52 val iff_attrs = @{attributes [iff]};
    53 val safe_elim_attrs = @{attributes [elim!]};
    54 val simp_attrs = @{attributes [simp]};
    55 
    56 fun pad_list x n xs = xs @ replicate (n - length xs) x;
    57 
    58 fun unflat_lookup eq ys zs = map (map (fn x => nth zs (find_index (curry eq x) ys)));
    59 
    60 fun mk_half_pairss' _ [] = []
    61   | mk_half_pairss' indent (y :: ys) =
    62     indent @ fold_rev (cons o single o pair y) ys (mk_half_pairss' ([] :: indent) ys);
    63 
    64 fun mk_half_pairss ys = mk_half_pairss' [[]] ys;
    65 
    66 fun mk_undefined T = Const (@{const_name undefined}, T);
    67 
    68 fun mk_ctr Ts ctr =
    69   let val Type (_, Ts0) = body_type (fastype_of ctr) in
    70     Term.subst_atomic_types (Ts0 ~~ Ts) ctr
    71   end;
    72 
    73 fun base_name_of_ctr c =
    74   Long_Name.base_name (case head_of c of
    75       Const (s, _) => s
    76     | Free (s, _) => s
    77     | _ => error "Cannot extract name of constructor");
    78 
    79 fun eta_expand_arg xs f_xs = fold_rev Term.lambda xs f_xs;
    80 
    81 fun prepare_wrap_datatype prep_term (((no_dests, raw_ctrs), raw_case),
    82     (raw_disc_bindings, (raw_sel_bindingss, raw_sel_defaultss))) no_defs_lthy =
    83   let
    84     (* TODO: sanity checks on arguments *)
    85     (* TODO: attributes (simp, case_names, etc.) *)
    86     (* TODO: case syntax *)
    87     (* TODO: integration with function package ("size") *)
    88 
    89     val n = length raw_ctrs;
    90     val ks = 1 upto n;
    91 
    92     val _ = if n > 0 then () else error "No constructors specified";
    93 
    94     val ctrs0 = map (prep_term no_defs_lthy) raw_ctrs;
    95     val case0 = prep_term no_defs_lthy raw_case;
    96     val sel_defaultss =
    97       pad_list [] n (map (map (apsnd (prep_term no_defs_lthy))) raw_sel_defaultss);
    98 
    99     val Type (dataT_name, As0) = body_type (fastype_of (hd ctrs0));
   100     val data_b = Binding.qualified_name dataT_name;
   101 
   102     val (As, B) =
   103       no_defs_lthy
   104       |> mk_TFrees' (map Type.sort_of_atyp As0)
   105       ||> the_single o fst o mk_TFrees 1;
   106 
   107     val dataT = Type (dataT_name, As);
   108     val ctrs = map (mk_ctr As) ctrs0;
   109     val ctr_Tss = map (binder_types o fastype_of) ctrs;
   110 
   111     val ms = map length ctr_Tss;
   112 
   113     val raw_disc_bindings' = pad_list Binding.empty n raw_disc_bindings;
   114 
   115     fun can_really_rely_on_disc k =
   116       not (Binding.eq_name (nth raw_disc_bindings' (k - 1), Binding.empty)) orelse
   117       nth ms (k - 1) = 0;
   118     fun can_rely_on_disc k =
   119       can_really_rely_on_disc k orelse (k = 1 andalso not (can_really_rely_on_disc 2));
   120     fun can_omit_disc_binding k m =
   121       n = 1 orelse m = 0 orelse (n = 2 andalso can_rely_on_disc (3 - k));
   122 
   123     val std_disc_binding =
   124       Binding.qualify false (Binding.name_of data_b) o Binding.name o prefix isN o base_name_of_ctr;
   125 
   126     val disc_bindings =
   127       raw_disc_bindings'
   128       |> map4 (fn k => fn m => fn ctr => fn disc =>
   129         Option.map (Binding.qualify false (Binding.name_of data_b))
   130           (if Binding.eq_name (disc, Binding.empty) then
   131              if can_omit_disc_binding k m then NONE else SOME (std_disc_binding ctr)
   132            else if Binding.eq_name (disc, std_binding) then
   133              SOME (std_disc_binding ctr)
   134            else
   135              SOME disc)) ks ms ctrs0;
   136 
   137     val no_discs = map is_none disc_bindings;
   138     val no_discs_at_all = forall I no_discs;
   139 
   140     fun std_sel_binding m l = Binding.name o mk_unN m l o base_name_of_ctr;
   141 
   142     val sel_bindingss =
   143       pad_list [] n raw_sel_bindingss
   144       |> map3 (fn ctr => fn m => map2 (fn l => fn sel =>
   145         Binding.qualify false (Binding.name_of data_b)
   146           (if Binding.eq_name (sel, Binding.empty) orelse Binding.eq_name (sel, std_binding) then
   147             std_sel_binding m l ctr
   148           else
   149             sel)) (1 upto m) o pad_list Binding.empty m) ctrs0 ms;
   150 
   151     fun mk_case Ts T =
   152       let
   153         val (bindings, body) = strip_type (fastype_of case0)
   154         val Type (_, Ts0) = List.last bindings
   155       in Term.subst_atomic_types ((body, T) :: (Ts0 ~~ Ts)) case0 end;
   156 
   157     val casex = mk_case As B;
   158     val case_Ts = map (fn Ts => Ts ---> B) ctr_Tss;
   159 
   160     val ((((((((xss, xss'), yss), fs), gs), (v, v')), w), (p, p')), names_lthy) = no_defs_lthy |>
   161       mk_Freess' "x" ctr_Tss
   162       ||>> mk_Freess "y" ctr_Tss
   163       ||>> mk_Frees "f" case_Ts
   164       ||>> mk_Frees "g" case_Ts
   165       ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "v") dataT
   166       ||>> yield_singleton (mk_Frees "w") dataT
   167       ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "P") HOLogic.boolT;
   168 
   169     val q = Free (fst p', B --> HOLogic.boolT);
   170 
   171     fun ap_v t = t $ v;
   172     fun mk_v_eq_v () = HOLogic.mk_eq (v, v);
   173 
   174     val xctrs = map2 (curry Term.list_comb) ctrs xss;
   175     val yctrs = map2 (curry Term.list_comb) ctrs yss;
   176 
   177     val xfs = map2 (curry Term.list_comb) fs xss;
   178     val xgs = map2 (curry Term.list_comb) gs xss;
   179 
   180     val eta_fs = map2 eta_expand_arg xss xfs;
   181     val eta_gs = map2 eta_expand_arg xss xgs;
   182 
   183     val fcase = Term.list_comb (casex, eta_fs);
   184     val gcase = Term.list_comb (casex, eta_gs);
   185 
   186     val exist_xs_v_eq_ctrs =
   187       map2 (fn xctr => fn xs => list_exists_free xs (HOLogic.mk_eq (v, xctr))) xctrs xss;
   188 
   189     val unique_disc_no_def = TrueI; (*arbitrary marker*)
   190     val alternate_disc_no_def = FalseE; (*arbitrary marker*)
   191 
   192     fun alternate_disc_lhs get_disc k =
   193       HOLogic.mk_not
   194         (case nth disc_bindings (k - 1) of
   195           NONE => nth exist_xs_v_eq_ctrs (k - 1)
   196         | SOME b => get_disc b (k - 1) $ v);
   197 
   198     val (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy') =
   199       if no_dests then
   200         (true, [], [], [], [], [], no_defs_lthy)
   201       else
   202         let
   203           fun disc_free b = Free (Binding.name_of b, dataT --> HOLogic.boolT);
   204 
   205           fun disc_spec b exist_xs_v_eq_ctr = mk_Trueprop_eq (disc_free b $ v, exist_xs_v_eq_ctr);
   206 
   207           fun alternate_disc k = Term.lambda v (alternate_disc_lhs (K o disc_free) (3 - k));
   208 
   209           fun mk_default T t =
   210             let
   211               val Ts0 = map TFree (Term.add_tfreesT (fastype_of t) []);
   212               val Ts = map TFree (Term.add_tfreesT T []);
   213             in Term.subst_atomic_types (Ts0 ~~ Ts) t end;
   214 
   215           fun mk_sel_case_args b proto_sels T =
   216             map2 (fn Ts => fn k =>
   217               (case AList.lookup (op =) proto_sels k of
   218                 NONE =>
   219                 (case AList.lookup Binding.eq_name (rev (nth sel_defaultss (k - 1))) b of
   220                   NONE => fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_undefined T)
   221                 | SOME t => mk_default (Ts ---> T) t)
   222               | SOME (xs, x) => fold_rev Term.lambda xs x)) ctr_Tss ks;
   223 
   224           fun sel_spec b proto_sels =
   225             let
   226               val _ =
   227                 (case duplicates (op =) (map fst proto_sels) of
   228                    k :: _ => error ("Duplicate selector name " ^ quote (Binding.name_of b) ^
   229                      " for constructor " ^
   230                      quote (Syntax.string_of_term no_defs_lthy (nth ctrs (k - 1))))
   231                  | [] => ())
   232               val T =
   233                 (case distinct (op =) (map (fastype_of o snd o snd) proto_sels) of
   234                   [T] => T
   235                 | T :: T' :: _ => error ("Inconsistent range type for selector " ^
   236                     quote (Binding.name_of b) ^ ": " ^ quote (Syntax.string_of_typ no_defs_lthy T) ^
   237                     " vs. " ^ quote (Syntax.string_of_typ no_defs_lthy T')));
   238             in
   239               mk_Trueprop_eq (Free (Binding.name_of b, dataT --> T) $ v,
   240                 Term.list_comb (mk_case As T, mk_sel_case_args b proto_sels T) $ v)
   241             end;
   242 
   243           val sel_bindings = flat sel_bindingss;
   244           val uniq_sel_bindings = distinct Binding.eq_name sel_bindings;
   245           val all_sels_distinct = (length uniq_sel_bindings = length sel_bindings);
   246 
   247           val sel_binding_index =
   248             if all_sels_distinct then 1 upto length sel_bindings
   249             else map (fn b => find_index (curry Binding.eq_name b) uniq_sel_bindings) sel_bindings;
   250 
   251           val proto_sels = flat (map3 (fn k => fn xs => map (fn x => (k, (xs, x)))) ks xss xss);
   252           val sel_infos =
   253             AList.group (op =) (sel_binding_index ~~ proto_sels)
   254             |> sort (int_ord o pairself fst)
   255             |> map snd |> curry (op ~~) uniq_sel_bindings;
   256           val sel_bindings = map fst sel_infos;
   257 
   258           fun unflat_selss xs = unflat_lookup Binding.eq_name sel_bindings xs sel_bindingss;
   259 
   260           val (((raw_discs, raw_disc_defs), (raw_sels, raw_sel_defs)), (lthy', lthy)) =
   261             no_defs_lthy
   262             |> apfst split_list o fold_map4 (fn k => fn m => fn exist_xs_v_eq_ctr =>
   263               fn NONE =>
   264                  if n = 1 then pair (Term.lambda v (mk_v_eq_v ()), unique_disc_no_def)
   265                  else if m = 0 then pair (Term.lambda v exist_xs_v_eq_ctr, refl)
   266                  else pair (alternate_disc k, alternate_disc_no_def)
   267                | SOME b => Specification.definition (SOME (b, NONE, NoSyn),
   268                    ((Thm.def_binding b, []), disc_spec b exist_xs_v_eq_ctr)) #>> apsnd snd)
   269               ks ms exist_xs_v_eq_ctrs disc_bindings
   270             ||>> apfst split_list o fold_map (fn (b, proto_sels) =>
   271               Specification.definition (SOME (b, NONE, NoSyn),
   272                 ((Thm.def_binding b, []), sel_spec b proto_sels)) #>> apsnd snd) sel_infos
   273             ||> `Local_Theory.restore;
   274 
   275           val phi = Proof_Context.export_morphism lthy lthy';
   276 
   277           val disc_defs = map (Morphism.thm phi) raw_disc_defs;
   278           val sel_defs = map (Morphism.thm phi) raw_sel_defs;
   279           val sel_defss = unflat_selss sel_defs;
   280 
   281           val discs0 = map (Morphism.term phi) raw_discs;
   282           val selss0 = unflat_selss (map (Morphism.term phi) raw_sels);
   283 
   284           fun mk_disc_or_sel Ts c =
   285             Term.subst_atomic_types (snd (Term.dest_Type (domain_type (fastype_of c))) ~~ Ts) c;
   286 
   287           val discs = map (mk_disc_or_sel As) discs0;
   288           val selss = map (map (mk_disc_or_sel As)) selss0;
   289         in
   290           (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy')
   291         end;
   292 
   293     fun mk_imp_p Qs = Logic.list_implies (Qs, HOLogic.mk_Trueprop p);
   294 
   295     val goal_exhaust =
   296       let fun mk_prem xctr xs = fold_rev Logic.all xs (mk_imp_p [mk_Trueprop_eq (v, xctr)]) in
   297         fold_rev Logic.all [p, v] (mk_imp_p (map2 mk_prem xctrs xss))
   298       end;
   299 
   300     val goal_injectss =
   301       let
   302         fun mk_goal _ _ [] [] = []
   303           | mk_goal xctr yctr xs ys =
   304             [fold_rev Logic.all (xs @ ys) (mk_Trueprop_eq (HOLogic.mk_eq (xctr, yctr),
   305               Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) xs ys)))];
   306       in
   307         map4 mk_goal xctrs yctrs xss yss
   308       end;
   309 
   310     val goal_half_distinctss =
   311       let
   312         fun mk_goal ((xs, xc), (xs', xc')) =
   313           fold_rev Logic.all (xs @ xs')
   314             (HOLogic.mk_Trueprop (HOLogic.mk_not (HOLogic.mk_eq (xc, xc'))));
   315       in
   316         map (map mk_goal) (mk_half_pairss (xss ~~ xctrs))
   317       end;
   318 
   319     val goal_cases =
   320       map3 (fn xs => fn xctr => fn xf =>
   321         fold_rev Logic.all (fs @ xs) (mk_Trueprop_eq (fcase $ xctr, xf))) xss xctrs xfs;
   322 
   323     val goalss = [goal_exhaust] :: goal_injectss @ goal_half_distinctss @ [goal_cases];
   324 
   325     fun after_qed thmss lthy =
   326       let
   327         val ([exhaust_thm], (inject_thmss, (half_distinct_thmss, [case_thms]))) =
   328           (hd thmss, apsnd (chop (n * n)) (chop n (tl thmss)));
   329 
   330         val inject_thms = flat inject_thmss;
   331 
   332         val exhaust_thm' =
   333           let val Tinst = map (pairself (certifyT lthy)) (map Logic.varifyT_global As ~~ As) in
   334             Drule.instantiate' [] [SOME (certify lthy v)]
   335               (Thm.instantiate (Tinst, []) (Drule.zero_var_indexes exhaust_thm))
   336           end;
   337 
   338         val exhaust_cases = map base_name_of_ctr ctrs;
   339 
   340         val other_half_distinct_thmss = map (map (fn thm => thm RS not_sym)) half_distinct_thmss;
   341 
   342         val (distinct_thmsss', distinct_thmsss) =
   343           map2 (map2 append) (Library.chop_groups n half_distinct_thmss)
   344             (transpose (Library.chop_groups n other_half_distinct_thmss))
   345           |> `transpose;
   346         val distinct_thms = interleave (flat half_distinct_thmss) (flat other_half_distinct_thmss);
   347 
   348         val nchotomy_thm =
   349           let
   350             val goal =
   351               HOLogic.mk_Trueprop (HOLogic.mk_all (fst v', snd v',
   352                 Library.foldr1 HOLogic.mk_disj exist_xs_v_eq_ctrs));
   353           in
   354             Skip_Proof.prove lthy [] [] goal (fn _ => mk_nchotomy_tac n exhaust_thm)
   355           end;
   356 
   357         val (all_sel_thms, sel_thmss, disc_thms, discI_thms, disc_exclude_thms, disc_exhaust_thms,
   358              collapse_thms, case_eq_thms) =
   359           if no_dests then
   360             ([], [], [], [], [], [], [], [])
   361           else
   362             let
   363               fun make_sel_thm xs' case_thm sel_def =
   364                 zero_var_indexes (Drule.gen_all (Drule.rename_bvars' (map (SOME o fst) xs')
   365                     (Drule.forall_intr_vars (case_thm RS (sel_def RS trans)))));
   366 
   367               fun has_undefined_rhs thm =
   368                 (case snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of thm))) of
   369                   Const (@{const_name undefined}, _) => true
   370                 | _ => false);
   371 
   372               val sel_thmss = map3 (map oo make_sel_thm) xss' case_thms sel_defss;
   373 
   374               val all_sel_thms =
   375                 (if all_sels_distinct andalso forall null sel_defaultss then
   376                    flat sel_thmss
   377                  else
   378                    map_product (fn s => fn (xs', c) => make_sel_thm xs' c s) sel_defs
   379                      (xss' ~~ case_thms))
   380                 |> filter_out has_undefined_rhs;
   381 
   382               fun mk_unique_disc_def () =
   383                 let
   384                   val m = the_single ms;
   385                   val goal = mk_Trueprop_eq (mk_v_eq_v (), the_single exist_xs_v_eq_ctrs);
   386                 in
   387                   Skip_Proof.prove lthy [] [] goal (fn _ => mk_unique_disc_def_tac m exhaust_thm')
   388                   |> singleton (Proof_Context.export names_lthy lthy)
   389                   |> Thm.close_derivation
   390                 end;
   391 
   392               fun mk_alternate_disc_def k =
   393                 let
   394                   val goal =
   395                     mk_Trueprop_eq (alternate_disc_lhs (K (nth discs)) (3 - k),
   396                       nth exist_xs_v_eq_ctrs (k - 1));
   397                 in
   398                   Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
   399                     mk_alternate_disc_def_tac ctxt k (nth disc_defs (2 - k))
   400                       (nth distinct_thms (2 - k)) exhaust_thm')
   401                   |> singleton (Proof_Context.export names_lthy lthy)
   402                   |> Thm.close_derivation
   403                 end;
   404 
   405               val has_alternate_disc_def =
   406                 exists (fn def => Thm.eq_thm_prop (def, alternate_disc_no_def)) disc_defs;
   407 
   408               val disc_defs' =
   409                 map2 (fn k => fn def =>
   410                   if Thm.eq_thm_prop (def, unique_disc_no_def) then mk_unique_disc_def ()
   411                   else if Thm.eq_thm_prop (def, alternate_disc_no_def) then mk_alternate_disc_def k
   412                   else def) ks disc_defs;
   413 
   414               val discD_thms = map (fn def => def RS iffD1) disc_defs';
   415               val discI_thms =
   416                 map2 (fn m => fn def => funpow m (fn thm => exI RS thm) (def RS iffD2)) ms
   417                   disc_defs';
   418               val not_discI_thms =
   419                 map2 (fn m => fn def => funpow m (fn thm => allI RS thm)
   420                     (Local_Defs.unfold lthy @{thms not_ex} (def RS @{thm ssubst[of _ _ Not]})))
   421                   ms disc_defs';
   422 
   423               val (disc_thmss', disc_thmss) =
   424                 let
   425                   fun mk_thm discI _ [] = refl RS discI
   426                     | mk_thm _ not_discI [distinct] = distinct RS not_discI;
   427                   fun mk_thms discI not_discI distinctss = map (mk_thm discI not_discI) distinctss;
   428                 in
   429                   map3 mk_thms discI_thms not_discI_thms distinct_thmsss' |> `transpose
   430                 end;
   431 
   432               val disc_thms = flat (map2 (fn true => K [] | false => I) no_discs disc_thmss);
   433 
   434               val disc_exclude_thms =
   435                 if has_alternate_disc_def then
   436                   []
   437                 else
   438                   let
   439                     fun mk_goal [] = []
   440                       | mk_goal [((_, true), (_, true))] = []
   441                       | mk_goal [(((_, disc), _), ((_, disc'), _))] =
   442                         [Logic.all v (Logic.mk_implies (HOLogic.mk_Trueprop (betapply (disc, v)),
   443                            HOLogic.mk_Trueprop (HOLogic.mk_not (betapply (disc', v)))))];
   444                     fun prove tac goal = Skip_Proof.prove lthy [] [] goal (K tac);
   445 
   446                     val infos = ms ~~ discD_thms ~~ discs ~~ no_discs;
   447                     val half_pairss = mk_half_pairss infos;
   448 
   449                     val goal_halvess = map mk_goal half_pairss;
   450                     val half_thmss =
   451                       map3 (fn [] => K (K []) | [goal] => fn [((((m, discD), _), _), _)] =>
   452                           fn disc_thm => [prove (mk_half_disc_exclude_tac m discD disc_thm) goal])
   453                         goal_halvess half_pairss (flat disc_thmss');
   454 
   455                     val goal_other_halvess = map (mk_goal o map swap) half_pairss;
   456                     val other_half_thmss =
   457                       map2 (map2 (prove o mk_other_half_disc_exclude_tac)) half_thmss
   458                         goal_other_halvess;
   459                   in
   460                     interleave (flat half_thmss) (flat other_half_thmss)
   461                   end;
   462 
   463               val disc_exhaust_thms =
   464                 if has_alternate_disc_def orelse no_discs_at_all then
   465                   []
   466                 else
   467                   let
   468                     fun mk_prem disc = mk_imp_p [HOLogic.mk_Trueprop (betapply (disc, v))];
   469                     val goal = fold_rev Logic.all [p, v] (mk_imp_p (map mk_prem discs));
   470                   in
   471                     [Skip_Proof.prove lthy [] [] goal (fn _ =>
   472                        mk_disc_exhaust_tac n exhaust_thm discI_thms)]
   473                   end;
   474 
   475               val collapse_thms =
   476                 if no_dests then
   477                   []
   478                 else
   479                   let
   480                     fun mk_goal ctr disc sels =
   481                       let
   482                         val prem = HOLogic.mk_Trueprop (betapply (disc, v));
   483                         val concl =
   484                           mk_Trueprop_eq ((null sels ? swap)
   485                             (Term.list_comb (ctr, map ap_v sels), v));
   486                       in
   487                         if prem aconv concl then NONE
   488                         else SOME (Logic.all v (Logic.mk_implies (prem, concl)))
   489                       end;
   490                     val goals = map3 mk_goal ctrs discs selss;
   491                   in
   492                     map4 (fn m => fn discD => fn sel_thms => Option.map (fn goal =>
   493                       Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
   494                         mk_collapse_tac ctxt m discD sel_thms)
   495                       |> perhaps (try (fn thm => refl RS thm)))) ms discD_thms sel_thmss goals
   496                     |> map_filter I
   497                   end;
   498 
   499               val case_eq_thms =
   500                 if no_dests then
   501                   []
   502                 else
   503                   let
   504                     fun mk_body f sels = Term.list_comb (f, map ap_v sels);
   505                     val goal =
   506                       mk_Trueprop_eq (fcase $ v, mk_IfN B (map ap_v discs) (map2 mk_body fs selss));
   507                   in
   508                     [Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
   509                       mk_case_eq_tac ctxt n exhaust_thm' case_thms disc_thmss' sel_thmss)]
   510                     |> Proof_Context.export names_lthy lthy
   511                   end;
   512             in
   513               (all_sel_thms, sel_thmss, disc_thms, discI_thms, disc_exclude_thms, disc_exhaust_thms,
   514                collapse_thms, case_eq_thms)
   515             end;
   516 
   517         val (case_cong_thm, weak_case_cong_thm) =
   518           let
   519             fun mk_prem xctr xs f g =
   520               fold_rev Logic.all xs (Logic.mk_implies (mk_Trueprop_eq (w, xctr),
   521                 mk_Trueprop_eq (f, g)));
   522 
   523             val v_eq_w = mk_Trueprop_eq (v, w);
   524 
   525             val goal =
   526               Logic.list_implies (v_eq_w :: map4 mk_prem xctrs xss fs gs,
   527                  mk_Trueprop_eq (fcase $ v, gcase $ w));
   528             val goal_weak = Logic.mk_implies (v_eq_w, mk_Trueprop_eq (fcase $ v, fcase $ w));
   529           in
   530             (Skip_Proof.prove lthy [] [] goal (fn _ => mk_case_cong_tac exhaust_thm' case_thms),
   531              Skip_Proof.prove lthy [] [] goal_weak (K (etac arg_cong 1)))
   532             |> pairself (singleton (Proof_Context.export names_lthy lthy))
   533           end;
   534 
   535         val (split_thm, split_asm_thm) =
   536           let
   537             fun mk_conjunct xctr xs f_xs =
   538               list_all_free xs (HOLogic.mk_imp (HOLogic.mk_eq (v, xctr), q $ f_xs));
   539             fun mk_disjunct xctr xs f_xs =
   540               list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (v, xctr),
   541                 HOLogic.mk_not (q $ f_xs)));
   542 
   543             val lhs = q $ (fcase $ v);
   544 
   545             val goal =
   546               mk_Trueprop_eq (lhs, Library.foldr1 HOLogic.mk_conj (map3 mk_conjunct xctrs xss xfs));
   547             val goal_asm =
   548               mk_Trueprop_eq (lhs, HOLogic.mk_not (Library.foldr1 HOLogic.mk_disj
   549                 (map3 mk_disjunct xctrs xss xfs)));
   550 
   551             val split_thm =
   552               Skip_Proof.prove lthy [] [] goal
   553                 (fn _ => mk_split_tac exhaust_thm' case_thms inject_thmss distinct_thmsss)
   554               |> singleton (Proof_Context.export names_lthy lthy)
   555             val split_asm_thm =
   556               Skip_Proof.prove lthy [] [] goal_asm (fn {context = ctxt, ...} =>
   557                 mk_split_asm_tac ctxt split_thm)
   558               |> singleton (Proof_Context.export names_lthy lthy)
   559           in
   560             (split_thm, split_asm_thm)
   561           end;
   562 
   563         val exhaust_case_names_attr = Attrib.internal (K (Rule_Cases.case_names exhaust_cases));
   564         val cases_type_attr = Attrib.internal (K (Induct.cases_type dataT_name));
   565 
   566         val notes =
   567           [(case_congN, [case_cong_thm], []),
   568            (case_eqN, case_eq_thms, []),
   569            (casesN, case_thms, simp_attrs),
   570            (collapseN, collapse_thms, simp_attrs),
   571            (discsN, disc_thms, simp_attrs),
   572            (disc_excludeN, disc_exclude_thms, []),
   573            (disc_exhaustN, disc_exhaust_thms, [exhaust_case_names_attr]),
   574            (distinctN, distinct_thms, simp_attrs @ induct_simp_attrs),
   575            (exhaustN, [exhaust_thm], [exhaust_case_names_attr, cases_type_attr]),
   576            (injectN, inject_thms, iff_attrs @ induct_simp_attrs),
   577            (nchotomyN, [nchotomy_thm], []),
   578            (selsN, all_sel_thms, simp_attrs),
   579            (splitN, [split_thm], []),
   580            (split_asmN, [split_asm_thm], []),
   581            (weak_case_cong_thmsN, [weak_case_cong_thm], cong_attrs)]
   582           |> filter_out (null o #2)
   583           |> map (fn (thmN, thms, attrs) =>
   584             ((Binding.qualify true (Binding.name_of data_b) (Binding.name thmN), attrs),
   585              [(thms, [])]));
   586 
   587         val notes' =
   588           [(map (fn th => th RS notE) distinct_thms, safe_elim_attrs)]
   589           |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
   590       in
   591         ((selss, inject_thms, distinct_thms, case_thms, discI_thms, sel_thmss),
   592          lthy |> Local_Theory.notes (notes' @ notes) |> snd)
   593       end;
   594   in
   595     (goalss, after_qed, lthy')
   596   end;
   597 
   598 fun wrap_datatype tacss = (fn (goalss, after_qed, lthy) =>
   599   map2 (map2 (Skip_Proof.prove lthy [] [])) goalss tacss
   600   |> (fn thms => after_qed thms lthy)) oo prepare_wrap_datatype (K I);
   601 
   602 val wrap_datatype_cmd = (fn (goalss, after_qed, lthy) =>
   603   Proof.theorem NONE (snd oo after_qed) (map (map (rpair [])) goalss) lthy) oo
   604   prepare_wrap_datatype Syntax.read_term;
   605 
   606 fun parse_bracket_list parser = @{keyword "["} |-- Parse.list parser --|  @{keyword "]"};
   607 
   608 val parse_bindings = parse_bracket_list Parse.binding;
   609 val parse_bindingss = parse_bracket_list parse_bindings;
   610 
   611 val parse_bound_term = (Parse.binding --| @{keyword ":"}) -- Parse.term;
   612 val parse_bound_terms = parse_bracket_list parse_bound_term;
   613 val parse_bound_termss = parse_bracket_list parse_bound_terms;
   614 
   615 val parse_wrap_options =
   616   Scan.optional (@{keyword "("} |-- (@{keyword "no_dests"} >> K true) --| @{keyword ")"}) false;
   617 
   618 val _ =
   619   Outer_Syntax.local_theory_to_proof @{command_spec "wrap_data"} "wraps an existing datatype"
   620     ((parse_wrap_options -- (@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}) --
   621       Parse.term -- Scan.optional (parse_bindings -- Scan.optional (parse_bindingss --
   622         Scan.optional parse_bound_termss []) ([], [])) ([], ([], [])))
   623      >> wrap_datatype_cmd);
   624 
   625 end;