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