src/HOL/BNF/Tools/ctr_sugar.ML
changeset 56013 d64a4ef26edb
parent 55145 b15cfc2864de
child 55597 297d1c603999
equal deleted inserted replaced
56012:cfb21e03fe2a 56013:d64a4ef26edb
     1 (*  Title:      HOL/BNF/Tools/ctr_sugar.ML
       
     2     Author:     Jasmin Blanchette, TU Muenchen
       
     3     Copyright   2012
       
     4 
       
     5 Wrapping existing freely generated type's constructors.
       
     6 *)
       
     7 
       
     8 signature CTR_SUGAR =
       
     9 sig
       
    10   type ctr_sugar =
       
    11     {ctrs: term list,
       
    12      casex: term,
       
    13      discs: term list,
       
    14      selss: term list list,
       
    15      exhaust: thm,
       
    16      nchotomy: thm,
       
    17      injects: thm list,
       
    18      distincts: thm list,
       
    19      case_thms: thm list,
       
    20      case_cong: thm,
       
    21      weak_case_cong: thm,
       
    22      split: thm,
       
    23      split_asm: thm,
       
    24      disc_thmss: thm list list,
       
    25      discIs: thm list,
       
    26      sel_thmss: thm list list,
       
    27      disc_exhausts: thm list,
       
    28      sel_exhausts: thm list,
       
    29      collapses: thm list,
       
    30      expands: thm list,
       
    31      sel_splits: thm list,
       
    32      sel_split_asms: thm list,
       
    33      case_conv_ifs: thm list};
       
    34 
       
    35   val morph_ctr_sugar: morphism -> ctr_sugar -> ctr_sugar
       
    36   val ctr_sugar_of: Proof.context -> string -> ctr_sugar option
       
    37   val ctr_sugars_of: Proof.context -> ctr_sugar list
       
    38 
       
    39   val rep_compat_prefix: string
       
    40 
       
    41   val mk_half_pairss: 'a list * 'a list -> ('a * 'a) list list
       
    42   val join_halves: int -> 'a list list -> 'a list list -> 'a list * 'a list list list
       
    43 
       
    44   val mk_ctr: typ list -> term -> term
       
    45   val mk_case: typ list -> typ -> term -> term
       
    46   val mk_disc_or_sel: typ list -> term -> term
       
    47   val name_of_ctr: term -> string
       
    48   val name_of_disc: term -> string
       
    49   val dest_ctr: Proof.context -> string -> term -> term * term list
       
    50   val dest_case: Proof.context -> string -> typ list -> term -> (term list * term list) option
       
    51 
       
    52   val wrap_free_constructors: ({prems: thm list, context: Proof.context} -> tactic) list list ->
       
    53     (((bool * bool) * term list) * binding) *
       
    54       (binding list * (binding list list * (binding * term) list list)) -> local_theory ->
       
    55     ctr_sugar * local_theory
       
    56   val parse_wrap_free_constructors_options: (bool * bool) parser
       
    57   val parse_bound_term: (binding * string) parser
       
    58 end;
       
    59 
       
    60 structure Ctr_Sugar : CTR_SUGAR =
       
    61 struct
       
    62 
       
    63 open Ctr_Sugar_Util
       
    64 open Ctr_Sugar_Tactics
       
    65 
       
    66 type ctr_sugar =
       
    67   {ctrs: term list,
       
    68    casex: term,
       
    69    discs: term list,
       
    70    selss: term list list,
       
    71    exhaust: thm,
       
    72    nchotomy: thm,
       
    73    injects: thm list,
       
    74    distincts: thm list,
       
    75    case_thms: thm list,
       
    76    case_cong: thm,
       
    77    weak_case_cong: thm,
       
    78    split: thm,
       
    79    split_asm: thm,
       
    80    disc_thmss: thm list list,
       
    81    discIs: thm list,
       
    82    sel_thmss: thm list list,
       
    83    disc_exhausts: thm list,
       
    84    sel_exhausts: thm list,
       
    85    collapses: thm list,
       
    86    expands: thm list,
       
    87    sel_splits: thm list,
       
    88    sel_split_asms: thm list,
       
    89    case_conv_ifs: thm list};
       
    90 
       
    91 fun eq_ctr_sugar ({ctrs = ctrs1, casex = case1, discs = discs1, selss = selss1, ...} : ctr_sugar,
       
    92     {ctrs = ctrs2, casex = case2, discs = discs2, selss = selss2, ...} : ctr_sugar) =
       
    93   ctrs1 = ctrs2 andalso case1 = case2 andalso discs1 = discs2 andalso selss1 = selss2;
       
    94 
       
    95 fun morph_ctr_sugar phi {ctrs, casex, discs, selss, exhaust, nchotomy, injects, distincts,
       
    96     case_thms, case_cong, weak_case_cong, split, split_asm, disc_thmss, discIs, sel_thmss,
       
    97     disc_exhausts, sel_exhausts, collapses, expands, sel_splits, sel_split_asms, case_conv_ifs} =
       
    98   {ctrs = map (Morphism.term phi) ctrs,
       
    99    casex = Morphism.term phi casex,
       
   100    discs = map (Morphism.term phi) discs,
       
   101    selss = map (map (Morphism.term phi)) selss,
       
   102    exhaust = Morphism.thm phi exhaust,
       
   103    nchotomy = Morphism.thm phi nchotomy,
       
   104    injects = map (Morphism.thm phi) injects,
       
   105    distincts = map (Morphism.thm phi) distincts,
       
   106    case_thms = map (Morphism.thm phi) case_thms,
       
   107    case_cong = Morphism.thm phi case_cong,
       
   108    weak_case_cong = Morphism.thm phi weak_case_cong,
       
   109    split = Morphism.thm phi split,
       
   110    split_asm = Morphism.thm phi split_asm,
       
   111    disc_thmss = map (map (Morphism.thm phi)) disc_thmss,
       
   112    discIs = map (Morphism.thm phi) discIs,
       
   113    sel_thmss = map (map (Morphism.thm phi)) sel_thmss,
       
   114    disc_exhausts = map (Morphism.thm phi) disc_exhausts,
       
   115    sel_exhausts = map (Morphism.thm phi) sel_exhausts,
       
   116    collapses = map (Morphism.thm phi) collapses,
       
   117    expands = map (Morphism.thm phi) expands,
       
   118    sel_splits = map (Morphism.thm phi) sel_splits,
       
   119    sel_split_asms = map (Morphism.thm phi) sel_split_asms,
       
   120    case_conv_ifs = map (Morphism.thm phi) case_conv_ifs};
       
   121 
       
   122 val transfer_ctr_sugar =
       
   123   morph_ctr_sugar o Morphism.thm_morphism o Thm.transfer o Proof_Context.theory_of;
       
   124 
       
   125 structure Data = Generic_Data
       
   126 (
       
   127   type T = ctr_sugar Symtab.table;
       
   128   val empty = Symtab.empty;
       
   129   val extend = I;
       
   130   val merge = Symtab.merge eq_ctr_sugar;
       
   131 );
       
   132 
       
   133 fun ctr_sugar_of ctxt =
       
   134   Symtab.lookup (Data.get (Context.Proof ctxt))
       
   135   #> Option.map (transfer_ctr_sugar ctxt);
       
   136 
       
   137 fun ctr_sugars_of ctxt =
       
   138   Symtab.fold (cons o transfer_ctr_sugar ctxt o snd) (Data.get (Context.Proof ctxt)) [];
       
   139 
       
   140 fun register_ctr_sugar key ctr_sugar =
       
   141   Local_Theory.declaration {syntax = false, pervasive = true}
       
   142     (fn phi => Data.map (Symtab.default (key, morph_ctr_sugar phi ctr_sugar)));
       
   143 
       
   144 val rep_compat_prefix = "new";
       
   145 
       
   146 val isN = "is_";
       
   147 val unN = "un_";
       
   148 fun mk_unN 1 1 suf = unN ^ suf
       
   149   | mk_unN _ l suf = unN ^ suf ^ string_of_int l;
       
   150 
       
   151 val caseN = "case";
       
   152 val case_congN = "case_cong";
       
   153 val case_conv_ifN = "case_conv_if";
       
   154 val collapseN = "collapse";
       
   155 val disc_excludeN = "disc_exclude";
       
   156 val disc_exhaustN = "disc_exhaust";
       
   157 val discN = "disc";
       
   158 val discIN = "discI";
       
   159 val distinctN = "distinct";
       
   160 val exhaustN = "exhaust";
       
   161 val expandN = "expand";
       
   162 val injectN = "inject";
       
   163 val nchotomyN = "nchotomy";
       
   164 val selN = "sel";
       
   165 val sel_exhaustN = "sel_exhaust";
       
   166 val sel_splitN = "sel_split";
       
   167 val sel_split_asmN = "sel_split_asm";
       
   168 val splitN = "split";
       
   169 val splitsN = "splits";
       
   170 val split_asmN = "split_asm";
       
   171 val weak_case_cong_thmsN = "weak_case_cong";
       
   172 
       
   173 val cong_attrs = @{attributes [cong]};
       
   174 val dest_attrs = @{attributes [dest]};
       
   175 val safe_elim_attrs = @{attributes [elim!]};
       
   176 val iff_attrs = @{attributes [iff]};
       
   177 val induct_simp_attrs = @{attributes [induct_simp]};
       
   178 val nitpick_attrs = @{attributes [nitpick_simp]};
       
   179 val simp_attrs = @{attributes [simp]};
       
   180 val code_nitpick_simp_simp_attrs = Code.add_default_eqn_attrib :: nitpick_attrs @ simp_attrs;
       
   181 
       
   182 fun unflat_lookup eq xs ys = map (fn xs' => permute_like eq xs xs' ys);
       
   183 
       
   184 fun mk_half_pairss' _ ([], []) = []
       
   185   | mk_half_pairss' indent (x :: xs, _ :: ys) =
       
   186     indent @ fold_rev (cons o single o pair x) ys (mk_half_pairss' ([] :: indent) (xs, ys));
       
   187 
       
   188 fun mk_half_pairss p = mk_half_pairss' [[]] p;
       
   189 
       
   190 fun join_halves n half_xss other_half_xss =
       
   191   let
       
   192     val xsss =
       
   193       map2 (map2 append) (Library.chop_groups n half_xss)
       
   194         (transpose (Library.chop_groups n other_half_xss))
       
   195     val xs = splice (flat half_xss) (flat other_half_xss);
       
   196   in (xs, xsss) end;
       
   197 
       
   198 fun mk_undefined T = Const (@{const_name undefined}, T);
       
   199 
       
   200 fun mk_ctr Ts t =
       
   201   let val Type (_, Ts0) = body_type (fastype_of t) in
       
   202     Term.subst_atomic_types (Ts0 ~~ Ts) t
       
   203   end;
       
   204 
       
   205 fun mk_case Ts T t =
       
   206   let val (Type (_, Ts0), body) = strip_type (fastype_of t) |>> List.last in
       
   207     Term.subst_atomic_types ((body, T) :: (Ts0 ~~ Ts)) t
       
   208   end;
       
   209 
       
   210 fun mk_disc_or_sel Ts t =
       
   211   Term.subst_atomic_types (snd (Term.dest_Type (domain_type (fastype_of t))) ~~ Ts) t;
       
   212 
       
   213 fun name_of_const what t =
       
   214   (case head_of t of
       
   215     Const (s, _) => s
       
   216   | Free (s, _) => s
       
   217   | _ => error ("Cannot extract name of " ^ what));
       
   218 
       
   219 val name_of_ctr = name_of_const "constructor";
       
   220 
       
   221 val notN = "not_";
       
   222 val eqN = "eq_";
       
   223 val neqN = "neq_";
       
   224 
       
   225 fun name_of_disc t =
       
   226   (case head_of t of
       
   227     Abs (_, _, @{const Not} $ (t' $ Bound 0)) =>
       
   228     Long_Name.map_base_name (prefix notN) (name_of_disc t')
       
   229   | Abs (_, _, Const (@{const_name HOL.eq}, _) $ Bound 0 $ t') =>
       
   230     Long_Name.map_base_name (prefix eqN) (name_of_disc t')
       
   231   | Abs (_, _, @{const Not} $ (Const (@{const_name HOL.eq}, _) $ Bound 0 $ t')) =>
       
   232     Long_Name.map_base_name (prefix neqN) (name_of_disc t')
       
   233   | t' => name_of_const "destructor" t');
       
   234 
       
   235 val base_name_of_ctr = Long_Name.base_name o name_of_ctr;
       
   236 
       
   237 fun dest_ctr ctxt s t =
       
   238   let
       
   239     val (f, args) = Term.strip_comb t;
       
   240   in
       
   241     (case ctr_sugar_of ctxt s of
       
   242       SOME {ctrs, ...} =>
       
   243       (case find_first (can (fo_match ctxt f)) ctrs of
       
   244         SOME f' => (f', args)
       
   245       | NONE => raise Fail "dest_ctr")
       
   246     | NONE => raise Fail "dest_ctr")
       
   247   end;
       
   248 
       
   249 fun dest_case ctxt s Ts t =
       
   250   (case Term.strip_comb t of
       
   251     (Const (c, _), args as _ :: _) =>
       
   252     (case ctr_sugar_of ctxt s of
       
   253       SOME {casex = Const (case_name, _), discs = discs0, selss = selss0, ...} =>
       
   254       if case_name = c then
       
   255         let val n = length discs0 in
       
   256           if n < length args then
       
   257             let
       
   258               val (branches, obj :: leftovers) = chop n args;
       
   259               val discs = map (mk_disc_or_sel Ts) discs0;
       
   260               val selss = map (map (mk_disc_or_sel Ts)) selss0;
       
   261               val conds = map (rapp obj) discs;
       
   262               val branch_argss = map (fn sels => map (rapp obj) sels @ leftovers) selss;
       
   263               val branches' = map2 (curry Term.betapplys) branches branch_argss;
       
   264             in
       
   265               SOME (conds, branches')
       
   266             end
       
   267           else
       
   268             NONE
       
   269         end
       
   270       else
       
   271         NONE
       
   272     | _ => NONE)
       
   273   | _ => NONE);
       
   274 
       
   275 fun eta_expand_arg xs f_xs = fold_rev Term.lambda xs f_xs;
       
   276 
       
   277 fun prepare_wrap_free_constructors prep_term ((((no_discs_sels, rep_compat), raw_ctrs),
       
   278     raw_case_binding), (raw_disc_bindings, (raw_sel_bindingss, raw_sel_defaultss))) no_defs_lthy =
       
   279   let
       
   280     (* TODO: sanity checks on arguments *)
       
   281 
       
   282     val n = length raw_ctrs;
       
   283     val ks = 1 upto n;
       
   284 
       
   285     val _ = if n > 0 then () else error "No constructors specified";
       
   286 
       
   287     val ctrs0 = map (prep_term no_defs_lthy) raw_ctrs;
       
   288     val sel_defaultss =
       
   289       pad_list [] n (map (map (apsnd (prep_term no_defs_lthy))) raw_sel_defaultss);
       
   290 
       
   291     val Type (fcT_name, As0) = body_type (fastype_of (hd ctrs0));
       
   292     val fc_b_name = Long_Name.base_name fcT_name;
       
   293     val fc_b = Binding.name fc_b_name;
       
   294 
       
   295     fun qualify mandatory =
       
   296       Binding.qualify mandatory fc_b_name o (rep_compat ? Binding.qualify false rep_compat_prefix);
       
   297 
       
   298     fun dest_TFree_or_TVar (TFree p) = p
       
   299       | dest_TFree_or_TVar (TVar ((s, _), S)) = (s, S)
       
   300       | dest_TFree_or_TVar _ = error "Invalid type argument";
       
   301 
       
   302     val (unsorted_As, B) =
       
   303       no_defs_lthy
       
   304       |> variant_tfrees (map (fst o dest_TFree_or_TVar) As0)
       
   305       ||> the_single o fst o mk_TFrees 1;
       
   306 
       
   307     val As = map2 (resort_tfree o snd o dest_TFree_or_TVar) As0 unsorted_As;
       
   308 
       
   309     val fcT = Type (fcT_name, As);
       
   310     val ctrs = map (mk_ctr As) ctrs0;
       
   311     val ctr_Tss = map (binder_types o fastype_of) ctrs;
       
   312 
       
   313     val ms = map length ctr_Tss;
       
   314 
       
   315     val raw_disc_bindings' = pad_list Binding.empty n raw_disc_bindings;
       
   316 
       
   317     fun can_definitely_rely_on_disc k = not (Binding.is_empty (nth raw_disc_bindings' (k - 1)));
       
   318     fun can_rely_on_disc k =
       
   319       can_definitely_rely_on_disc k orelse (k = 1 andalso not (can_definitely_rely_on_disc 2));
       
   320     fun should_omit_disc_binding k = n = 1 orelse (n = 2 andalso can_rely_on_disc (3 - k));
       
   321 
       
   322     fun is_disc_binding_valid b =
       
   323       not (Binding.is_empty b orelse Binding.eq_name (b, equal_binding));
       
   324 
       
   325     val standard_disc_binding = Binding.name o prefix isN o base_name_of_ctr;
       
   326 
       
   327     val disc_bindings =
       
   328       raw_disc_bindings'
       
   329       |> map4 (fn k => fn m => fn ctr => fn disc =>
       
   330         qualify false
       
   331           (if Binding.is_empty disc then
       
   332              if should_omit_disc_binding k then disc else standard_disc_binding ctr
       
   333            else if Binding.eq_name (disc, equal_binding) then
       
   334              if m = 0 then disc
       
   335              else error "Cannot use \"=\" syntax for discriminating nonnullary constructor"
       
   336            else if Binding.eq_name (disc, standard_binding) then
       
   337              standard_disc_binding ctr
       
   338            else
       
   339              disc)) ks ms ctrs0;
       
   340 
       
   341     fun standard_sel_binding m l = Binding.name o mk_unN m l o base_name_of_ctr;
       
   342 
       
   343     val sel_bindingss =
       
   344       pad_list [] n raw_sel_bindingss
       
   345       |> map3 (fn ctr => fn m => map2 (fn l => fn sel =>
       
   346         qualify false
       
   347           (if Binding.is_empty sel orelse Binding.eq_name (sel, standard_binding) then
       
   348             standard_sel_binding m l ctr
       
   349           else
       
   350             sel)) (1 upto m) o pad_list Binding.empty m) ctrs0 ms;
       
   351 
       
   352     val case_Ts = map (fn Ts => Ts ---> B) ctr_Tss;
       
   353 
       
   354     val ((((((((xss, xss'), yss), fs), gs), [u', v']), [w]), (p, p')), names_lthy) = no_defs_lthy |>
       
   355       mk_Freess' "x" ctr_Tss
       
   356       ||>> mk_Freess "y" ctr_Tss
       
   357       ||>> mk_Frees "f" case_Ts
       
   358       ||>> mk_Frees "g" case_Ts
       
   359       ||>> (apfst (map (rpair fcT)) oo Variable.variant_fixes) [fc_b_name, fc_b_name ^ "'"]
       
   360       ||>> mk_Frees "z" [B]
       
   361       ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "P") HOLogic.boolT;
       
   362 
       
   363     val u = Free u';
       
   364     val v = Free v';
       
   365     val q = Free (fst p', mk_pred1T B);
       
   366 
       
   367     val xctrs = map2 (curry Term.list_comb) ctrs xss;
       
   368     val yctrs = map2 (curry Term.list_comb) ctrs yss;
       
   369 
       
   370     val xfs = map2 (curry Term.list_comb) fs xss;
       
   371     val xgs = map2 (curry Term.list_comb) gs xss;
       
   372 
       
   373     (* TODO: Eta-expension is for compatibility with the old datatype package (but it also provides
       
   374        nicer names). Consider removing. *)
       
   375     val eta_fs = map2 eta_expand_arg xss xfs;
       
   376     val eta_gs = map2 eta_expand_arg xss xgs;
       
   377 
       
   378     val case_binding =
       
   379       qualify false
       
   380         (if Binding.is_empty raw_case_binding orelse
       
   381             Binding.eq_name (raw_case_binding, standard_binding) then
       
   382            Binding.suffix_name ("_" ^ caseN) fc_b
       
   383          else
       
   384            raw_case_binding);
       
   385 
       
   386     fun mk_case_disj xctr xf xs =
       
   387       list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr), HOLogic.mk_eq (w, xf)));
       
   388 
       
   389     val case_rhs = fold_rev (fold_rev Term.lambda) [fs, [u]]
       
   390       (Const (@{const_name The}, (B --> HOLogic.boolT) --> B) $
       
   391          Term.lambda w (Library.foldr1 HOLogic.mk_disj (map3 mk_case_disj xctrs xfs xss)));
       
   392 
       
   393     val ((raw_case, (_, raw_case_def)), (lthy', lthy)) = no_defs_lthy
       
   394       |> Local_Theory.define ((case_binding, NoSyn), ((Thm.def_binding case_binding, []), case_rhs))
       
   395       ||> `Local_Theory.restore;
       
   396 
       
   397     val phi = Proof_Context.export_morphism lthy lthy';
       
   398 
       
   399     val case_def = Morphism.thm phi raw_case_def;
       
   400 
       
   401     val case0 = Morphism.term phi raw_case;
       
   402     val casex = mk_case As B case0;
       
   403 
       
   404     val fcase = Term.list_comb (casex, fs);
       
   405 
       
   406     val ufcase = fcase $ u;
       
   407     val vfcase = fcase $ v;
       
   408 
       
   409     val eta_fcase = Term.list_comb (casex, eta_fs);
       
   410     val eta_gcase = Term.list_comb (casex, eta_gs);
       
   411 
       
   412     val eta_ufcase = eta_fcase $ u;
       
   413     val eta_vgcase = eta_gcase $ v;
       
   414 
       
   415     fun mk_uu_eq () = HOLogic.mk_eq (u, u);
       
   416 
       
   417     val uv_eq = mk_Trueprop_eq (u, v);
       
   418 
       
   419     val exist_xs_u_eq_ctrs =
       
   420       map2 (fn xctr => fn xs => list_exists_free xs (HOLogic.mk_eq (u, xctr))) xctrs xss;
       
   421 
       
   422     val unique_disc_no_def = TrueI; (*arbitrary marker*)
       
   423     val alternate_disc_no_def = FalseE; (*arbitrary marker*)
       
   424 
       
   425     fun alternate_disc_lhs get_udisc k =
       
   426       HOLogic.mk_not
       
   427         (let val b = nth disc_bindings (k - 1) in
       
   428            if is_disc_binding_valid b then get_udisc b (k - 1) else nth exist_xs_u_eq_ctrs (k - 1)
       
   429          end);
       
   430 
       
   431     val (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy') =
       
   432       if no_discs_sels then
       
   433         (true, [], [], [], [], [], lthy)
       
   434       else
       
   435         let
       
   436           fun disc_free b = Free (Binding.name_of b, mk_pred1T fcT);
       
   437 
       
   438           fun disc_spec b exist_xs_u_eq_ctr = mk_Trueprop_eq (disc_free b $ u, exist_xs_u_eq_ctr);
       
   439 
       
   440           fun alternate_disc k =
       
   441             Term.lambda u (alternate_disc_lhs (K o rapp u o disc_free) (3 - k));
       
   442 
       
   443           fun mk_sel_case_args b proto_sels T =
       
   444             map2 (fn Ts => fn k =>
       
   445               (case AList.lookup (op =) proto_sels k of
       
   446                 NONE =>
       
   447                 (case AList.lookup Binding.eq_name (rev (nth sel_defaultss (k - 1))) b of
       
   448                   NONE => fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_undefined T)
       
   449                 | SOME t => t |> Type.constraint (Ts ---> T) |> Syntax.check_term lthy)
       
   450               | SOME (xs, x) => fold_rev Term.lambda xs x)) ctr_Tss ks;
       
   451 
       
   452           fun sel_spec b proto_sels =
       
   453             let
       
   454               val _ =
       
   455                 (case duplicates (op =) (map fst proto_sels) of
       
   456                    k :: _ => error ("Duplicate selector name " ^ quote (Binding.name_of b) ^
       
   457                      " for constructor " ^
       
   458                      quote (Syntax.string_of_term lthy (nth ctrs (k - 1))))
       
   459                  | [] => ())
       
   460               val T =
       
   461                 (case distinct (op =) (map (fastype_of o snd o snd) proto_sels) of
       
   462                   [T] => T
       
   463                 | T :: T' :: _ => error ("Inconsistent range type for selector " ^
       
   464                     quote (Binding.name_of b) ^ ": " ^ quote (Syntax.string_of_typ lthy T) ^ " vs. "
       
   465                     ^ quote (Syntax.string_of_typ lthy T')));
       
   466             in
       
   467               mk_Trueprop_eq (Free (Binding.name_of b, fcT --> T) $ u,
       
   468                 Term.list_comb (mk_case As T case0, mk_sel_case_args b proto_sels T) $ u)
       
   469             end;
       
   470 
       
   471           val sel_bindings = flat sel_bindingss;
       
   472           val uniq_sel_bindings = distinct Binding.eq_name sel_bindings;
       
   473           val all_sels_distinct = (length uniq_sel_bindings = length sel_bindings);
       
   474 
       
   475           val sel_binding_index =
       
   476             if all_sels_distinct then 1 upto length sel_bindings
       
   477             else map (fn b => find_index (curry Binding.eq_name b) uniq_sel_bindings) sel_bindings;
       
   478 
       
   479           val proto_sels = flat (map3 (fn k => fn xs => map (fn x => (k, (xs, x)))) ks xss xss);
       
   480           val sel_infos =
       
   481             AList.group (op =) (sel_binding_index ~~ proto_sels)
       
   482             |> sort (int_ord o pairself fst)
       
   483             |> map snd |> curry (op ~~) uniq_sel_bindings;
       
   484           val sel_bindings = map fst sel_infos;
       
   485 
       
   486           fun unflat_selss xs = unflat_lookup Binding.eq_name sel_bindings xs sel_bindingss;
       
   487 
       
   488           val (((raw_discs, raw_disc_defs), (raw_sels, raw_sel_defs)), (lthy', lthy)) =
       
   489             lthy
       
   490             |> apfst split_list o fold_map3 (fn k => fn exist_xs_u_eq_ctr => fn b =>
       
   491                 if Binding.is_empty b then
       
   492                   if n = 1 then pair (Term.lambda u (mk_uu_eq ()), unique_disc_no_def)
       
   493                   else pair (alternate_disc k, alternate_disc_no_def)
       
   494                 else if Binding.eq_name (b, equal_binding) then
       
   495                   pair (Term.lambda u exist_xs_u_eq_ctr, refl)
       
   496                 else
       
   497                   Specification.definition (SOME (b, NONE, NoSyn),
       
   498                     ((Thm.def_binding b, []), disc_spec b exist_xs_u_eq_ctr)) #>> apsnd snd)
       
   499               ks exist_xs_u_eq_ctrs disc_bindings
       
   500             ||>> apfst split_list o fold_map (fn (b, proto_sels) =>
       
   501               Specification.definition (SOME (b, NONE, NoSyn),
       
   502                 ((Thm.def_binding b, []), sel_spec b proto_sels)) #>> apsnd snd) sel_infos
       
   503             ||> `Local_Theory.restore;
       
   504 
       
   505           val phi = Proof_Context.export_morphism lthy lthy';
       
   506 
       
   507           val disc_defs = map (Morphism.thm phi) raw_disc_defs;
       
   508           val sel_defs = map (Morphism.thm phi) raw_sel_defs;
       
   509           val sel_defss = unflat_selss sel_defs;
       
   510 
       
   511           val discs0 = map (Morphism.term phi) raw_discs;
       
   512           val selss0 = unflat_selss (map (Morphism.term phi) raw_sels);
       
   513 
       
   514           val discs = map (mk_disc_or_sel As) discs0;
       
   515           val selss = map (map (mk_disc_or_sel As)) selss0;
       
   516         in
       
   517           (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy')
       
   518         end;
       
   519 
       
   520     fun mk_imp_p Qs = Logic.list_implies (Qs, HOLogic.mk_Trueprop p);
       
   521 
       
   522     val exhaust_goal =
       
   523       let fun mk_prem xctr xs = fold_rev Logic.all xs (mk_imp_p [mk_Trueprop_eq (u, xctr)]) in
       
   524         fold_rev Logic.all [p, u] (mk_imp_p (map2 mk_prem xctrs xss))
       
   525       end;
       
   526 
       
   527     val inject_goalss =
       
   528       let
       
   529         fun mk_goal _ _ [] [] = []
       
   530           | mk_goal xctr yctr xs ys =
       
   531             [fold_rev Logic.all (xs @ ys) (mk_Trueprop_eq (HOLogic.mk_eq (xctr, yctr),
       
   532               Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) xs ys)))];
       
   533       in
       
   534         map4 mk_goal xctrs yctrs xss yss
       
   535       end;
       
   536 
       
   537     val half_distinct_goalss =
       
   538       let
       
   539         fun mk_goal ((xs, xc), (xs', xc')) =
       
   540           fold_rev Logic.all (xs @ xs')
       
   541             (HOLogic.mk_Trueprop (HOLogic.mk_not (HOLogic.mk_eq (xc, xc'))));
       
   542       in
       
   543         map (map mk_goal) (mk_half_pairss (`I (xss ~~ xctrs)))
       
   544       end;
       
   545 
       
   546     val goalss = [exhaust_goal] :: inject_goalss @ half_distinct_goalss;
       
   547 
       
   548     fun after_qed thmss lthy =
       
   549       let
       
   550         val ([exhaust_thm], (inject_thmss, half_distinct_thmss)) = (hd thmss, chop n (tl thmss));
       
   551 
       
   552         val inject_thms = flat inject_thmss;
       
   553 
       
   554         val rho_As = map (pairself (certifyT lthy)) (map Logic.varifyT_global As ~~ As);
       
   555 
       
   556         fun inst_thm t thm =
       
   557           Drule.instantiate' [] [SOME (certify lthy t)]
       
   558             (Thm.instantiate (rho_As, []) (Drule.zero_var_indexes thm));
       
   559 
       
   560         val uexhaust_thm = inst_thm u exhaust_thm;
       
   561 
       
   562         val exhaust_cases = map base_name_of_ctr ctrs;
       
   563 
       
   564         val other_half_distinct_thmss = map (map (fn thm => thm RS not_sym)) half_distinct_thmss;
       
   565 
       
   566         val (distinct_thms, (distinct_thmsss', distinct_thmsss)) =
       
   567           join_halves n half_distinct_thmss other_half_distinct_thmss ||> `transpose;
       
   568 
       
   569         val nchotomy_thm =
       
   570           let
       
   571             val goal =
       
   572               HOLogic.mk_Trueprop (HOLogic.mk_all (fst u', snd u',
       
   573                 Library.foldr1 HOLogic.mk_disj exist_xs_u_eq_ctrs));
       
   574           in
       
   575             Goal.prove_sorry lthy [] [] goal (fn _ => mk_nchotomy_tac n exhaust_thm)
       
   576             |> Thm.close_derivation
       
   577           end;
       
   578 
       
   579         val case_thms =
       
   580           let
       
   581             val goals =
       
   582               map3 (fn xctr => fn xf => fn xs =>
       
   583                 fold_rev Logic.all (fs @ xs) (mk_Trueprop_eq (fcase $ xctr, xf))) xctrs xfs xss;
       
   584           in
       
   585             map4 (fn k => fn goal => fn injects => fn distinctss =>
       
   586                 Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
       
   587                   mk_case_tac ctxt n k case_def injects distinctss)
       
   588                 |> Thm.close_derivation)
       
   589               ks goals inject_thmss distinct_thmsss
       
   590           end;
       
   591 
       
   592         val (case_cong_thm, weak_case_cong_thm) =
       
   593           let
       
   594             fun mk_prem xctr xs xf xg =
       
   595               fold_rev Logic.all xs (Logic.mk_implies (mk_Trueprop_eq (v, xctr),
       
   596                 mk_Trueprop_eq (xf, xg)));
       
   597 
       
   598             val goal =
       
   599               Logic.list_implies (uv_eq :: map4 mk_prem xctrs xss xfs xgs,
       
   600                  mk_Trueprop_eq (eta_ufcase, eta_vgcase));
       
   601             val weak_goal = Logic.mk_implies (uv_eq, mk_Trueprop_eq (ufcase, vfcase));
       
   602           in
       
   603             (Goal.prove_sorry lthy [] [] goal (fn _ => mk_case_cong_tac lthy uexhaust_thm case_thms),
       
   604              Goal.prove_sorry lthy [] [] weak_goal (K (etac arg_cong 1)))
       
   605             |> pairself (Thm.close_derivation #> singleton (Proof_Context.export names_lthy lthy))
       
   606           end;
       
   607 
       
   608         val split_lhs = q $ ufcase;
       
   609 
       
   610         fun mk_split_conjunct xctr xs f_xs =
       
   611           list_all_free xs (HOLogic.mk_imp (HOLogic.mk_eq (u, xctr), q $ f_xs));
       
   612         fun mk_split_disjunct xctr xs f_xs =
       
   613           list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr),
       
   614             HOLogic.mk_not (q $ f_xs)));
       
   615 
       
   616         fun mk_split_goal xctrs xss xfs =
       
   617           mk_Trueprop_eq (split_lhs, Library.foldr1 HOLogic.mk_conj
       
   618             (map3 mk_split_conjunct xctrs xss xfs));
       
   619         fun mk_split_asm_goal xctrs xss xfs =
       
   620           mk_Trueprop_eq (split_lhs, HOLogic.mk_not (Library.foldr1 HOLogic.mk_disj
       
   621             (map3 mk_split_disjunct xctrs xss xfs)));
       
   622 
       
   623         fun prove_split selss goal =
       
   624           Goal.prove_sorry lthy [] [] goal (fn _ =>
       
   625             mk_split_tac lthy uexhaust_thm case_thms selss inject_thmss distinct_thmsss)
       
   626           |> Thm.close_derivation
       
   627           |> singleton (Proof_Context.export names_lthy lthy);
       
   628 
       
   629         fun prove_split_asm asm_goal split_thm =
       
   630           Goal.prove_sorry lthy [] [] asm_goal (fn {context = ctxt, ...} =>
       
   631             mk_split_asm_tac ctxt split_thm)
       
   632           |> Thm.close_derivation
       
   633           |> singleton (Proof_Context.export names_lthy lthy);
       
   634 
       
   635         val (split_thm, split_asm_thm) =
       
   636           let
       
   637             val goal = mk_split_goal xctrs xss xfs;
       
   638             val asm_goal = mk_split_asm_goal xctrs xss xfs;
       
   639 
       
   640             val thm = prove_split (replicate n []) goal;
       
   641             val asm_thm = prove_split_asm asm_goal thm;
       
   642           in
       
   643             (thm, asm_thm)
       
   644           end;
       
   645 
       
   646         val (all_sel_thms, sel_thmss, disc_thmss, nontriv_disc_thms, discI_thms, nontriv_discI_thms,
       
   647              disc_exclude_thms, disc_exhaust_thms, sel_exhaust_thms, all_collapse_thms,
       
   648              safe_collapse_thms, expand_thms, sel_split_thms, sel_split_asm_thms,
       
   649              case_conv_if_thms) =
       
   650           if no_discs_sels then
       
   651             ([], [], [], [], [], [], [], [], [], [], [], [], [], [], [])
       
   652           else
       
   653             let
       
   654               val udiscs = map (rapp u) discs;
       
   655               val uselss = map (map (rapp u)) selss;
       
   656               val usel_ctrs = map2 (curry Term.list_comb) ctrs uselss;
       
   657               val usel_fs = map2 (curry Term.list_comb) fs uselss;
       
   658 
       
   659               val vdiscs = map (rapp v) discs;
       
   660               val vselss = map (map (rapp v)) selss;
       
   661 
       
   662               fun make_sel_thm xs' case_thm sel_def =
       
   663                 zero_var_indexes (Drule.gen_all (Drule.rename_bvars' (map (SOME o fst) xs')
       
   664                     (Drule.forall_intr_vars (case_thm RS (sel_def RS trans)))));
       
   665 
       
   666               val sel_thmss = map3 (map oo make_sel_thm) xss' case_thms sel_defss;
       
   667 
       
   668               fun has_undefined_rhs thm =
       
   669                 (case snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of thm))) of
       
   670                   Const (@{const_name undefined}, _) => true
       
   671                 | _ => false);
       
   672 
       
   673               val all_sel_thms =
       
   674                 (if all_sels_distinct andalso forall null sel_defaultss then
       
   675                    flat sel_thmss
       
   676                  else
       
   677                    map_product (fn s => fn (xs', c) => make_sel_thm xs' c s) sel_defs
       
   678                      (xss' ~~ case_thms))
       
   679                 |> filter_out has_undefined_rhs;
       
   680 
       
   681               fun mk_unique_disc_def () =
       
   682                 let
       
   683                   val m = the_single ms;
       
   684                   val goal = mk_Trueprop_eq (mk_uu_eq (), the_single exist_xs_u_eq_ctrs);
       
   685                 in
       
   686                   Goal.prove_sorry lthy [] [] goal (fn _ => mk_unique_disc_def_tac m uexhaust_thm)
       
   687                   |> Thm.close_derivation
       
   688                   |> singleton (Proof_Context.export names_lthy lthy)
       
   689                 end;
       
   690 
       
   691               fun mk_alternate_disc_def k =
       
   692                 let
       
   693                   val goal =
       
   694                     mk_Trueprop_eq (alternate_disc_lhs (K (nth udiscs)) (3 - k),
       
   695                       nth exist_xs_u_eq_ctrs (k - 1));
       
   696                 in
       
   697                   Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
       
   698                     mk_alternate_disc_def_tac ctxt k (nth disc_defs (2 - k))
       
   699                       (nth distinct_thms (2 - k)) uexhaust_thm)
       
   700                   |> Thm.close_derivation
       
   701                   |> singleton (Proof_Context.export names_lthy lthy)
       
   702                 end;
       
   703 
       
   704               val has_alternate_disc_def =
       
   705                 exists (fn def => Thm.eq_thm_prop (def, alternate_disc_no_def)) disc_defs;
       
   706 
       
   707               val disc_defs' =
       
   708                 map2 (fn k => fn def =>
       
   709                   if Thm.eq_thm_prop (def, unique_disc_no_def) then mk_unique_disc_def ()
       
   710                   else if Thm.eq_thm_prop (def, alternate_disc_no_def) then mk_alternate_disc_def k
       
   711                   else def) ks disc_defs;
       
   712 
       
   713               val discD_thms = map (fn def => def RS iffD1) disc_defs';
       
   714               val discI_thms =
       
   715                 map2 (fn m => fn def => funpow m (fn thm => exI RS thm) (def RS iffD2)) ms
       
   716                   disc_defs';
       
   717               val not_discI_thms =
       
   718                 map2 (fn m => fn def => funpow m (fn thm => allI RS thm)
       
   719                     (unfold_thms lthy @{thms not_ex} (def RS @{thm ssubst[of _ _ Not]})))
       
   720                   ms disc_defs';
       
   721 
       
   722               val (disc_thmss', disc_thmss) =
       
   723                 let
       
   724                   fun mk_thm discI _ [] = refl RS discI
       
   725                     | mk_thm _ not_discI [distinct] = distinct RS not_discI;
       
   726                   fun mk_thms discI not_discI distinctss = map (mk_thm discI not_discI) distinctss;
       
   727                 in
       
   728                   map3 mk_thms discI_thms not_discI_thms distinct_thmsss' |> `transpose
       
   729                 end;
       
   730 
       
   731               val nontriv_disc_thms =
       
   732                 flat (map2 (fn b => if is_disc_binding_valid b then I else K [])
       
   733                   disc_bindings disc_thmss);
       
   734 
       
   735               fun is_discI_boring b =
       
   736                 (n = 1 andalso Binding.is_empty b) orelse Binding.eq_name (b, equal_binding);
       
   737 
       
   738               val nontriv_discI_thms =
       
   739                 flat (map2 (fn b => if is_discI_boring b then K [] else single) disc_bindings
       
   740                   discI_thms);
       
   741 
       
   742               val (disc_exclude_thms, (disc_exclude_thmsss', disc_exclude_thmsss)) =
       
   743                 let
       
   744                   fun mk_goal [] = []
       
   745                     | mk_goal [((_, udisc), (_, udisc'))] =
       
   746                       [Logic.all u (Logic.mk_implies (HOLogic.mk_Trueprop udisc,
       
   747                          HOLogic.mk_Trueprop (HOLogic.mk_not udisc')))];
       
   748 
       
   749                   fun prove tac goal =
       
   750                     Goal.prove_sorry lthy [] [] goal (K tac)
       
   751                     |> Thm.close_derivation;
       
   752 
       
   753                   val half_pairss = mk_half_pairss (`I (ms ~~ discD_thms ~~ udiscs));
       
   754 
       
   755                   val half_goalss = map mk_goal half_pairss;
       
   756                   val half_thmss =
       
   757                     map3 (fn [] => K (K []) | [goal] => fn [(((m, discD), _), _)] =>
       
   758                         fn disc_thm => [prove (mk_half_disc_exclude_tac lthy m discD disc_thm) goal])
       
   759                       half_goalss half_pairss (flat disc_thmss');
       
   760 
       
   761                   val other_half_goalss = map (mk_goal o map swap) half_pairss;
       
   762                   val other_half_thmss =
       
   763                     map2 (map2 (prove o mk_other_half_disc_exclude_tac)) half_thmss
       
   764                       other_half_goalss;
       
   765                 in
       
   766                   join_halves n half_thmss other_half_thmss ||> `transpose
       
   767                   |>> has_alternate_disc_def ? K []
       
   768                 end;
       
   769 
       
   770               val disc_exhaust_thm =
       
   771                 let
       
   772                   fun mk_prem udisc = mk_imp_p [HOLogic.mk_Trueprop udisc];
       
   773                   val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem udiscs));
       
   774                 in
       
   775                   Goal.prove_sorry lthy [] [] goal (fn _ =>
       
   776                     mk_disc_exhaust_tac n exhaust_thm discI_thms)
       
   777                   |> Thm.close_derivation
       
   778                 end;
       
   779 
       
   780               val (safe_collapse_thms, all_collapse_thms) =
       
   781                 let
       
   782                   fun mk_goal m udisc usel_ctr =
       
   783                     let
       
   784                       val prem = HOLogic.mk_Trueprop udisc;
       
   785                       val concl = mk_Trueprop_eq ((usel_ctr, u) |> m = 0 ? swap);
       
   786                     in
       
   787                       (prem aconv concl, Logic.all u (Logic.mk_implies (prem, concl)))
       
   788                     end;
       
   789                   val (trivs, goals) = map3 mk_goal ms udiscs usel_ctrs |> split_list;
       
   790                   val thms =
       
   791                     map5 (fn m => fn discD => fn sel_thms => fn triv => fn goal =>
       
   792                         Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
       
   793                           mk_collapse_tac ctxt m discD sel_thms ORELSE HEADGOAL atac)
       
   794                         |> Thm.close_derivation
       
   795                         |> not triv ? perhaps (try (fn thm => refl RS thm)))
       
   796                       ms discD_thms sel_thmss trivs goals;
       
   797                 in
       
   798                   (map_filter (fn (true, _) => NONE | (false, thm) => SOME thm) (trivs ~~ thms),
       
   799                    thms)
       
   800                 end;
       
   801 
       
   802               val swapped_all_collapse_thms =
       
   803                 map2 (fn m => fn thm => if m = 0 then thm else thm RS sym) ms all_collapse_thms;
       
   804 
       
   805               val sel_exhaust_thm =
       
   806                 let
       
   807                   fun mk_prem usel_ctr = mk_imp_p [mk_Trueprop_eq (u, usel_ctr)];
       
   808                   val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem usel_ctrs));
       
   809                 in
       
   810                   Goal.prove_sorry lthy [] [] goal (fn _ =>
       
   811                     mk_sel_exhaust_tac n disc_exhaust_thm swapped_all_collapse_thms)
       
   812                   |> Thm.close_derivation
       
   813                 end;
       
   814 
       
   815               val expand_thm =
       
   816                 let
       
   817                   fun mk_prems k udisc usels vdisc vsels =
       
   818                     (if k = n then [] else [mk_Trueprop_eq (udisc, vdisc)]) @
       
   819                     (if null usels then
       
   820                        []
       
   821                      else
       
   822                        [Logic.list_implies
       
   823                           (if n = 1 then [] else map HOLogic.mk_Trueprop [udisc, vdisc],
       
   824                              HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
       
   825                                (map2 (curry HOLogic.mk_eq) usels vsels)))]);
       
   826 
       
   827                   val goal =
       
   828                     Library.foldr Logic.list_implies
       
   829                       (map5 mk_prems ks udiscs uselss vdiscs vselss, uv_eq);
       
   830                   val uncollapse_thms =
       
   831                     map2 (fn thm => fn [] => thm | _ => thm RS sym) all_collapse_thms uselss;
       
   832                 in
       
   833                   Goal.prove_sorry lthy [] [] goal (fn _ =>
       
   834                     mk_expand_tac lthy n ms (inst_thm u disc_exhaust_thm)
       
   835                       (inst_thm v disc_exhaust_thm) uncollapse_thms disc_exclude_thmsss
       
   836                       disc_exclude_thmsss')
       
   837                   |> Thm.close_derivation
       
   838                   |> singleton (Proof_Context.export names_lthy lthy)
       
   839                 end;
       
   840 
       
   841               val (sel_split_thm, sel_split_asm_thm) =
       
   842                 let
       
   843                   val zss = map (K []) xss;
       
   844                   val goal = mk_split_goal usel_ctrs zss usel_fs;
       
   845                   val asm_goal = mk_split_asm_goal usel_ctrs zss usel_fs;
       
   846 
       
   847                   val thm = prove_split sel_thmss goal;
       
   848                   val asm_thm = prove_split_asm asm_goal thm;
       
   849                 in
       
   850                   (thm, asm_thm)
       
   851                 end;
       
   852 
       
   853               val case_conv_if_thm =
       
   854                 let
       
   855                   val goal = mk_Trueprop_eq (ufcase, mk_IfN B udiscs usel_fs);
       
   856                 in
       
   857                   Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
       
   858                     mk_case_conv_if_tac ctxt n uexhaust_thm case_thms disc_thmss' sel_thmss)
       
   859                   |> Thm.close_derivation
       
   860                   |> singleton (Proof_Context.export names_lthy lthy)
       
   861                 end;
       
   862             in
       
   863               (all_sel_thms, sel_thmss, disc_thmss, nontriv_disc_thms, discI_thms,
       
   864                nontriv_discI_thms, disc_exclude_thms, [disc_exhaust_thm], [sel_exhaust_thm],
       
   865                all_collapse_thms, safe_collapse_thms, [expand_thm], [sel_split_thm],
       
   866                [sel_split_asm_thm], [case_conv_if_thm])
       
   867             end;
       
   868 
       
   869         val exhaust_case_names_attr = Attrib.internal (K (Rule_Cases.case_names exhaust_cases));
       
   870         val cases_type_attr = Attrib.internal (K (Induct.cases_type fcT_name));
       
   871 
       
   872         val notes =
       
   873           [(caseN, case_thms, code_nitpick_simp_simp_attrs),
       
   874            (case_congN, [case_cong_thm], []),
       
   875            (case_conv_ifN, case_conv_if_thms, []),
       
   876            (collapseN, safe_collapse_thms, simp_attrs),
       
   877            (discN, nontriv_disc_thms, simp_attrs),
       
   878            (discIN, nontriv_discI_thms, []),
       
   879            (disc_excludeN, disc_exclude_thms, dest_attrs),
       
   880            (disc_exhaustN, disc_exhaust_thms, [exhaust_case_names_attr]),
       
   881            (distinctN, distinct_thms, simp_attrs @ induct_simp_attrs),
       
   882            (exhaustN, [exhaust_thm], [exhaust_case_names_attr, cases_type_attr]),
       
   883            (expandN, expand_thms, []),
       
   884            (injectN, inject_thms, iff_attrs @ induct_simp_attrs),
       
   885            (nchotomyN, [nchotomy_thm], []),
       
   886            (selN, all_sel_thms, code_nitpick_simp_simp_attrs),
       
   887            (sel_exhaustN, sel_exhaust_thms, [exhaust_case_names_attr]),
       
   888            (sel_splitN, sel_split_thms, []),
       
   889            (sel_split_asmN, sel_split_asm_thms, []),
       
   890            (splitN, [split_thm], []),
       
   891            (split_asmN, [split_asm_thm], []),
       
   892            (splitsN, [split_thm, split_asm_thm], []),
       
   893            (weak_case_cong_thmsN, [weak_case_cong_thm], cong_attrs)]
       
   894           |> filter_out (null o #2)
       
   895           |> map (fn (thmN, thms, attrs) =>
       
   896             ((qualify true (Binding.name thmN), attrs), [(thms, [])]));
       
   897 
       
   898         val notes' =
       
   899           [(map (fn th => th RS notE) distinct_thms, safe_elim_attrs)]
       
   900           |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
       
   901 
       
   902         val ctr_sugar =
       
   903           {ctrs = ctrs, casex = casex, discs = discs, selss = selss, exhaust = exhaust_thm,
       
   904            nchotomy = nchotomy_thm, injects = inject_thms, distincts = distinct_thms,
       
   905            case_thms = case_thms, case_cong = case_cong_thm, weak_case_cong = weak_case_cong_thm,
       
   906            split = split_thm, split_asm = split_asm_thm, disc_thmss = disc_thmss,
       
   907            discIs = discI_thms, sel_thmss = sel_thmss, disc_exhausts = disc_exhaust_thms,
       
   908            sel_exhausts = sel_exhaust_thms, collapses = all_collapse_thms, expands = expand_thms,
       
   909            sel_splits = sel_split_thms, sel_split_asms = sel_split_asm_thms,
       
   910            case_conv_ifs = case_conv_if_thms};
       
   911       in
       
   912         (ctr_sugar,
       
   913          lthy
       
   914          |> not rep_compat ?
       
   915             (Local_Theory.declaration {syntax = false, pervasive = true}
       
   916                (fn phi => Case_Translation.register
       
   917                   (Morphism.term phi casex) (map (Morphism.term phi) ctrs)))
       
   918          |> Local_Theory.notes (notes' @ notes) |> snd
       
   919          |> register_ctr_sugar fcT_name ctr_sugar)
       
   920       end;
       
   921   in
       
   922     (goalss, after_qed, lthy')
       
   923   end;
       
   924 
       
   925 fun wrap_free_constructors tacss = (fn (goalss, after_qed, lthy) =>
       
   926   map2 (map2 (Thm.close_derivation oo Goal.prove_sorry lthy [] [])) goalss tacss
       
   927   |> (fn thms => after_qed thms lthy)) oo prepare_wrap_free_constructors (K I);
       
   928 
       
   929 val wrap_free_constructors_cmd = (fn (goalss, after_qed, lthy) =>
       
   930   Proof.theorem NONE (snd oo after_qed) (map (map (rpair [])) goalss) lthy) oo
       
   931   prepare_wrap_free_constructors Syntax.read_term;
       
   932 
       
   933 fun parse_bracket_list parser = @{keyword "["} |-- Parse.list parser --|  @{keyword "]"};
       
   934 
       
   935 val parse_bindings = parse_bracket_list parse_binding;
       
   936 val parse_bindingss = parse_bracket_list parse_bindings;
       
   937 
       
   938 val parse_bound_term = (parse_binding --| @{keyword ":"}) -- Parse.term;
       
   939 val parse_bound_terms = parse_bracket_list parse_bound_term;
       
   940 val parse_bound_termss = parse_bracket_list parse_bound_terms;
       
   941 
       
   942 val parse_wrap_free_constructors_options =
       
   943   Scan.optional (@{keyword "("} |-- Parse.list1 ((@{keyword "no_discs_sels"} >> K (true, false)) ||
       
   944       (@{keyword "rep_compat"} >> K (false, true))) --| @{keyword ")"}
       
   945     >> (pairself (exists I) o split_list)) (false, false);
       
   946 
       
   947 val _ =
       
   948   Outer_Syntax.local_theory_to_proof @{command_spec "wrap_free_constructors"}
       
   949     "wrap an existing freely generated type's constructors"
       
   950     ((parse_wrap_free_constructors_options -- (@{keyword "["} |-- Parse.list Parse.term --|
       
   951         @{keyword "]"}) --
       
   952       parse_binding -- Scan.optional (parse_bindings -- Scan.optional (parse_bindingss --
       
   953         Scan.optional parse_bound_termss []) ([], [])) ([], ([], [])))
       
   954      >> wrap_free_constructors_cmd);
       
   955 
       
   956 end;