src/HOL/BNF/Tools/bnf_fp_def_sugar.ML
changeset 49636 b7256a88a84b
parent 49633 5b5450bc544c
child 49670 c7a034d01936
equal deleted inserted replaced
49635:fc0777f04205 49636:b7256a88a84b
       
     1 (*  Title:      HOL/BNF/Tools/bnf_fp_def_sugar.ML
       
     2     Author:     Jasmin Blanchette, TU Muenchen
       
     3     Copyright   2012
       
     4 
       
     5 Sugared datatype and codatatype constructions.
       
     6 *)
       
     7 
       
     8 signature BNF_FP_DEF_SUGAR =
       
     9 sig
       
    10   val datatypes: bool ->
       
    11     (mixfix list -> (string * sort) list option -> binding list -> typ list * typ list list ->
       
    12       BNF_Def.BNF list -> local_theory -> BNF_FP.fp_result * local_theory) ->
       
    13     (bool * bool) * ((((typ * sort) list * binding) * mixfix) * ((((binding * binding) *
       
    14       (binding * typ) list) * (binding * term) list) * mixfix) list) list ->
       
    15     local_theory -> local_theory
       
    16   val parse_datatype_cmd: bool ->
       
    17     (mixfix list -> (string * sort) list option -> binding list -> typ list * typ list list ->
       
    18       BNF_Def.BNF list -> local_theory -> BNF_FP.fp_result * local_theory) ->
       
    19     (local_theory -> local_theory) parser
       
    20 end;
       
    21 
       
    22 structure BNF_FP_Def_Sugar : BNF_FP_DEF_SUGAR =
       
    23 struct
       
    24 
       
    25 open BNF_Util
       
    26 open BNF_Wrap
       
    27 open BNF_Def
       
    28 open BNF_FP
       
    29 open BNF_FP_Def_Sugar_Tactics
       
    30 
       
    31 (* This function could produce clashes in contrived examples (e.g., "x.A", "x.x_A", "y.A") *)
       
    32 fun quasi_unambiguous_case_names names =
       
    33   let
       
    34     val ps = map (`Long_Name.base_name) names;
       
    35     val dups = Library.duplicates (op =) (map fst ps);
       
    36     fun underscore s =
       
    37       let val ss = space_explode Long_Name.separator s in
       
    38         space_implode "_" (drop (length ss - 2) ss)
       
    39       end;
       
    40   in
       
    41     map (fn (base, full) => if member (op =) dups base then underscore full else base) ps
       
    42   end;
       
    43 
       
    44 val mp_conj = @{thm mp_conj};
       
    45 
       
    46 val simp_attrs = @{attributes [simp]};
       
    47 val code_simp_attrs = Code.add_default_eqn_attrib :: simp_attrs;
       
    48 
       
    49 fun split_list4 xs = (map #1 xs, map #2 xs, map #3 xs, map #4 xs);
       
    50 
       
    51 fun resort_tfree S (TFree (s, _)) = TFree (s, S);
       
    52 
       
    53 fun typ_subst inst (T as Type (s, Ts)) =
       
    54     (case AList.lookup (op =) inst T of
       
    55       NONE => Type (s, map (typ_subst inst) Ts)
       
    56     | SOME T' => T')
       
    57   | typ_subst inst T = the_default T (AList.lookup (op =) inst T);
       
    58 
       
    59 fun variant_types ss Ss ctxt =
       
    60   let
       
    61     val (tfrees, _) =
       
    62       fold_map2 (fn s => fn S => Name.variant s #> apfst (rpair S)) ss Ss (Variable.names_of ctxt);
       
    63     val ctxt' = fold (Variable.declare_constraints o Logic.mk_type o TFree) tfrees ctxt;
       
    64   in (tfrees, ctxt') end;
       
    65 
       
    66 val lists_bmoc = fold (fn xs => fn t => Term.list_comb (t, xs));
       
    67 
       
    68 fun mk_tupled_fun x f xs = HOLogic.tupled_lambda x (Term.list_comb (f, xs));
       
    69 fun mk_uncurried_fun f xs = mk_tupled_fun (HOLogic.mk_tuple xs) f xs;
       
    70 fun mk_uncurried2_fun f xss =
       
    71   mk_tupled_fun (HOLogic.mk_tuple (map HOLogic.mk_tuple xss)) f (flat xss);
       
    72 
       
    73 fun mk_flip (x, Type (_, [T1, Type (_, [T2, T3])])) =
       
    74   Abs ("x", T1, Abs ("y", T2, Var (x, T2 --> T1 --> T3) $ Bound 0 $ Bound 1));
       
    75 
       
    76 fun flip_rels lthy n thm =
       
    77   let
       
    78     val Rs = Term.add_vars (prop_of thm) [];
       
    79     val Rs' = rev (drop (length Rs - n) Rs);
       
    80     val cRs = map (fn f => (certify lthy (Var f), certify lthy (mk_flip f))) Rs';
       
    81   in
       
    82     Drule.cterm_instantiate cRs thm
       
    83   end;
       
    84 
       
    85 fun mk_ctor_or_dtor get_T Ts t =
       
    86   let val Type (_, Ts0) = get_T (fastype_of t) in
       
    87     Term.subst_atomic_types (Ts0 ~~ Ts) t
       
    88   end;
       
    89 
       
    90 val mk_ctor = mk_ctor_or_dtor range_type;
       
    91 val mk_dtor = mk_ctor_or_dtor domain_type;
       
    92 
       
    93 fun mk_rec_like lfp Ts Us t =
       
    94   let
       
    95     val (bindings, body) = strip_type (fastype_of t);
       
    96     val (f_Us, prebody) = split_last bindings;
       
    97     val Type (_, Ts0) = if lfp then prebody else body;
       
    98     val Us0 = distinct (op =) (map (if lfp then body_type else domain_type) f_Us);
       
    99   in
       
   100     Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
       
   101   end;
       
   102 
       
   103 fun mk_map live Ts Us t =
       
   104   let val (Type (_, Ts0), Type (_, Us0)) = strip_typeN (live + 1) (fastype_of t) |>> List.last in
       
   105     Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
       
   106   end;
       
   107 
       
   108 fun mk_rel live Ts Us t =
       
   109   let val [Type (_, Ts0), Type (_, Us0)] = binder_types (snd (strip_typeN live (fastype_of t))) in
       
   110     Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
       
   111   end;
       
   112 
       
   113 fun liveness_of_fp_bnf n bnf =
       
   114   (case T_of_bnf bnf of
       
   115     Type (_, Ts) => map (not o member (op =) (deads_of_bnf bnf)) Ts
       
   116   | _ => replicate n false);
       
   117 
       
   118 fun tick u f = Term.lambda u (HOLogic.mk_prod (u, f $ u));
       
   119 
       
   120 fun tack z_name (c, u) f =
       
   121   let val z = Free (z_name, mk_sumT (fastype_of u, fastype_of c)) in
       
   122     Term.lambda z (mk_sum_case (Term.lambda u u, Term.lambda c (f $ c)) $ z)
       
   123   end;
       
   124 
       
   125 fun cannot_merge_types () = error "Mutually recursive types must have the same type parameters";
       
   126 
       
   127 fun merge_type_arg T T' = if T = T' then T else cannot_merge_types ();
       
   128 
       
   129 fun merge_type_args (As, As') =
       
   130   if length As = length As' then map2 merge_type_arg As As' else cannot_merge_types ();
       
   131 
       
   132 fun reassoc_conjs thm =
       
   133   reassoc_conjs (thm RS @{thm conj_assoc[THEN iffD1]})
       
   134   handle THM _ => thm;
       
   135 
       
   136 fun type_args_constrained_of (((cAs, _), _), _) = cAs;
       
   137 fun type_binding_of (((_, b), _), _) = b;
       
   138 fun mixfix_of ((_, mx), _) = mx;
       
   139 fun ctr_specs_of (_, ctr_specs) = ctr_specs;
       
   140 
       
   141 fun disc_of ((((disc, _), _), _), _) = disc;
       
   142 fun ctr_of ((((_, ctr), _), _), _) = ctr;
       
   143 fun args_of (((_, args), _), _) = args;
       
   144 fun defaults_of ((_, ds), _) = ds;
       
   145 fun ctr_mixfix_of (_, mx) = mx;
       
   146 
       
   147 fun define_datatypes prepare_constraint prepare_typ prepare_term lfp construct_fp
       
   148     (wrap_opts as (no_dests, rep_compat), specs) no_defs_lthy0 =
       
   149   let
       
   150     (* TODO: sanity checks on arguments *)
       
   151     (* TODO: integration with function package ("size") *)
       
   152 
       
   153     val _ = if not lfp andalso no_dests then error "Cannot define destructor-less codatatypes"
       
   154       else ();
       
   155 
       
   156     fun qualify mandatory fp_b_name =
       
   157       Binding.qualify mandatory fp_b_name o (rep_compat ? Binding.qualify false rep_compat_prefix);
       
   158 
       
   159     val nn = length specs;
       
   160     val fp_bs = map type_binding_of specs;
       
   161     val fp_b_names = map Binding.name_of fp_bs;
       
   162     val fp_common_name = mk_common_name fp_b_names;
       
   163 
       
   164     fun prepare_type_arg (ty, c) =
       
   165       let val TFree (s, _) = prepare_typ no_defs_lthy0 ty in
       
   166         TFree (s, prepare_constraint no_defs_lthy0 c)
       
   167       end;
       
   168 
       
   169     val Ass0 = map (map prepare_type_arg o type_args_constrained_of) specs;
       
   170     val unsorted_Ass0 = map (map (resort_tfree HOLogic.typeS)) Ass0;
       
   171     val unsorted_As = Library.foldr1 merge_type_args unsorted_Ass0;
       
   172 
       
   173     val (((Bs0, Cs), Xs), no_defs_lthy) =
       
   174       no_defs_lthy0
       
   175       |> fold (Variable.declare_typ o resort_tfree dummyS) unsorted_As
       
   176       |> mk_TFrees (length unsorted_As)
       
   177       ||>> mk_TFrees nn
       
   178       ||>> apfst (map TFree) o
       
   179         variant_types (map (prefix "'") fp_b_names) (replicate nn HOLogic.typeS);
       
   180 
       
   181     (* TODO: cleaner handling of fake contexts, without "background_theory" *)
       
   182     (*the "perhaps o try" below helps gracefully handles the case where the new type is defined in a
       
   183       locale and shadows an existing global type*)
       
   184     val fake_thy =
       
   185       Theory.copy #> fold (fn spec => perhaps (try (Sign.add_type no_defs_lthy
       
   186         (type_binding_of spec, length (type_args_constrained_of spec), mixfix_of spec)))) specs;
       
   187     val fake_lthy = Proof_Context.background_theory fake_thy no_defs_lthy;
       
   188 
       
   189     fun mk_fake_T b =
       
   190       Type (fst (Term.dest_Type (Proof_Context.read_type_name fake_lthy true (Binding.name_of b))),
       
   191         unsorted_As);
       
   192 
       
   193     val fake_Ts = map mk_fake_T fp_bs;
       
   194 
       
   195     val mixfixes = map mixfix_of specs;
       
   196 
       
   197     val _ = (case duplicates Binding.eq_name fp_bs of [] => ()
       
   198       | b :: _ => error ("Duplicate type name declaration " ^ quote (Binding.name_of b)));
       
   199 
       
   200     val ctr_specss = map ctr_specs_of specs;
       
   201 
       
   202     val disc_bindingss = map (map disc_of) ctr_specss;
       
   203     val ctr_bindingss =
       
   204       map2 (fn fp_b_name => map (qualify false fp_b_name o ctr_of)) fp_b_names ctr_specss;
       
   205     val ctr_argsss = map (map args_of) ctr_specss;
       
   206     val ctr_mixfixess = map (map ctr_mixfix_of) ctr_specss;
       
   207 
       
   208     val sel_bindingsss = map (map (map fst)) ctr_argsss;
       
   209     val fake_ctr_Tsss0 = map (map (map (prepare_typ fake_lthy o snd))) ctr_argsss;
       
   210     val raw_sel_defaultsss = map (map defaults_of) ctr_specss;
       
   211 
       
   212     val (As :: _) :: fake_ctr_Tsss =
       
   213       burrow (burrow (Syntax.check_typs fake_lthy)) (Ass0 :: fake_ctr_Tsss0);
       
   214 
       
   215     val _ = (case duplicates (op =) unsorted_As of [] => ()
       
   216       | A :: _ => error ("Duplicate type parameter " ^
       
   217           quote (Syntax.string_of_typ no_defs_lthy A)));
       
   218 
       
   219     val rhs_As' = fold (fold (fold Term.add_tfreesT)) fake_ctr_Tsss [];
       
   220     val _ = (case subtract (op =) (map dest_TFree As) rhs_As' of
       
   221         [] => ()
       
   222       | A' :: _ => error ("Extra type variable on right-hand side: " ^
       
   223           quote (Syntax.string_of_typ no_defs_lthy (TFree A'))));
       
   224 
       
   225     fun eq_fpT_check (T as Type (s, Us)) (Type (s', Us')) =
       
   226         s = s' andalso (Us = Us' orelse error ("Illegal occurrence of recursive type " ^
       
   227           quote (Syntax.string_of_typ fake_lthy T)))
       
   228       | eq_fpT_check _ _ = false;
       
   229 
       
   230     fun freeze_fp (T as Type (s, Us)) =
       
   231         (case find_index (eq_fpT_check T) fake_Ts of
       
   232           ~1 => Type (s, map freeze_fp Us)
       
   233         | kk => nth Xs kk)
       
   234       | freeze_fp T = T;
       
   235 
       
   236     val ctr_TsssXs = map (map (map freeze_fp)) fake_ctr_Tsss;
       
   237     val ctr_sum_prod_TsXs = map (mk_sumTN_balanced o map HOLogic.mk_tupleT) ctr_TsssXs;
       
   238 
       
   239     val fp_eqs =
       
   240       map dest_TFree Xs ~~ map (Term.typ_subst_atomic (As ~~ unsorted_As)) ctr_sum_prod_TsXs;
       
   241 
       
   242     (* TODO: clean up list *)
       
   243     val (pre_bnfs, ((fp_bnfs as any_fp_bnf :: _, dtors0, ctors0, fp_folds0, fp_recs0, fp_induct,
       
   244            fp_strong_induct, dtor_ctors, ctor_dtors, ctor_injects, fp_map_thms, fp_set_thmss,
       
   245            fp_rel_thms, fp_fold_thms, fp_rec_thms), lthy)) =
       
   246       fp_bnf construct_fp fp_bs mixfixes (map dest_TFree unsorted_As) fp_eqs no_defs_lthy0;
       
   247 
       
   248     val timer = time (Timer.startRealTimer ());
       
   249 
       
   250     fun add_nesty_bnf_names Us =
       
   251       let
       
   252         fun add (Type (s, Ts)) ss =
       
   253             let val (needs, ss') = fold_map add Ts ss in
       
   254               if exists I needs then (true, insert (op =) s ss') else (false, ss')
       
   255             end
       
   256           | add T ss = (member (op =) Us T, ss);
       
   257       in snd oo add end;
       
   258 
       
   259     fun nesty_bnfs Us =
       
   260       map_filter (bnf_of lthy) (fold (fold (fold (add_nesty_bnf_names Us))) ctr_TsssXs []);
       
   261 
       
   262     val nesting_bnfs = nesty_bnfs As;
       
   263     val nested_bnfs = nesty_bnfs Xs;
       
   264 
       
   265     val pre_map_defs = map map_def_of_bnf pre_bnfs;
       
   266     val pre_set_defss = map set_defs_of_bnf pre_bnfs;
       
   267     val pre_rel_defs = map rel_def_of_bnf pre_bnfs;
       
   268     val nested_set_natural's = maps set_natural'_of_bnf nested_bnfs;
       
   269     val nesting_map_ids = map map_id_of_bnf nesting_bnfs;
       
   270     val nesting_set_natural's = maps set_natural'_of_bnf nesting_bnfs;
       
   271 
       
   272     val live = live_of_bnf any_fp_bnf;
       
   273 
       
   274     val Bs =
       
   275       map3 (fn alive => fn A as TFree (_, S) => fn B => if alive then resort_tfree S B else A)
       
   276         (liveness_of_fp_bnf (length As) any_fp_bnf) As Bs0;
       
   277 
       
   278     val B_ify = Term.typ_subst_atomic (As ~~ Bs);
       
   279 
       
   280     val ctors = map (mk_ctor As) ctors0;
       
   281     val dtors = map (mk_dtor As) dtors0;
       
   282 
       
   283     val fpTs = map (domain_type o fastype_of) dtors;
       
   284 
       
   285     val exists_fp_subtype = exists_subtype (member (op =) fpTs);
       
   286 
       
   287     val ctr_Tsss = map (map (map (Term.typ_subst_atomic (Xs ~~ fpTs)))) ctr_TsssXs;
       
   288     val ns = map length ctr_Tsss;
       
   289     val kss = map (fn n => 1 upto n) ns;
       
   290     val mss = map (map length) ctr_Tsss;
       
   291     val Css = map2 replicate ns Cs;
       
   292 
       
   293     val fp_folds as any_fp_fold :: _ = map (mk_rec_like lfp As Cs) fp_folds0;
       
   294     val fp_recs as any_fp_rec :: _ = map (mk_rec_like lfp As Cs) fp_recs0;
       
   295 
       
   296     val fp_fold_fun_Ts = fst (split_last (binder_types (fastype_of any_fp_fold)));
       
   297     val fp_rec_fun_Ts = fst (split_last (binder_types (fastype_of any_fp_rec)));
       
   298 
       
   299     val (((fold_only as (gss, _, _), rec_only as (hss, _, _)),
       
   300           (zs, cs, cpss, unfold_only as ((pgss, crgsss), _), corec_only as ((phss, cshsss), _))),
       
   301          names_lthy0) =
       
   302       if lfp then
       
   303         let
       
   304           val y_Tsss =
       
   305             map3 (fn n => fn ms => map2 dest_tupleT ms o dest_sumTN_balanced n o domain_type)
       
   306               ns mss fp_fold_fun_Ts;
       
   307           val g_Tss = map2 (map2 (curry (op --->))) y_Tsss Css;
       
   308 
       
   309           val ((gss, ysss), lthy) =
       
   310             lthy
       
   311             |> mk_Freess "f" g_Tss
       
   312             ||>> mk_Freesss "x" y_Tsss;
       
   313           val yssss = map (map (map single)) ysss;
       
   314 
       
   315           fun dest_rec_prodT (T as Type (@{type_name prod}, Us as [_, U])) =
       
   316               if member (op =) Cs U then Us else [T]
       
   317             | dest_rec_prodT T = [T];
       
   318 
       
   319           val z_Tssss =
       
   320             map3 (fn n => fn ms => map2 (map dest_rec_prodT oo dest_tupleT) ms o
       
   321               dest_sumTN_balanced n o domain_type) ns mss fp_rec_fun_Ts;
       
   322           val h_Tss = map2 (map2 (fold_rev (curry (op --->)))) z_Tssss Css;
       
   323 
       
   324           val hss = map2 (map2 retype_free) h_Tss gss;
       
   325           val zssss_hd = map2 (map2 (map2 (retype_free o hd))) z_Tssss ysss;
       
   326           val (zssss_tl, lthy) =
       
   327             lthy
       
   328             |> mk_Freessss "y" (map (map (map tl)) z_Tssss);
       
   329           val zssss = map2 (map2 (map2 cons)) zssss_hd zssss_tl;
       
   330         in
       
   331           ((((gss, g_Tss, yssss), (hss, h_Tss, zssss)),
       
   332             ([], [], [], (([], []), ([], [])), (([], []), ([], [])))), lthy)
       
   333         end
       
   334       else
       
   335         let
       
   336           (*avoid "'a itself" arguments in coiterators and corecursors*)
       
   337           val mss' =  map (fn [0] => [1] | ms => ms) mss;
       
   338 
       
   339           val p_Tss = map2 (fn n => replicate (Int.max (0, n - 1)) o mk_pred1T) ns Cs;
       
   340 
       
   341           fun flat_predss_getterss qss fss = maps (op @) (qss ~~ fss);
       
   342 
       
   343           fun flat_preds_predsss_gettersss [] [qss] [fss] = flat_predss_getterss qss fss
       
   344             | flat_preds_predsss_gettersss (p :: ps) (qss :: qsss) (fss :: fsss) =
       
   345               p :: flat_predss_getterss qss fss @ flat_preds_predsss_gettersss ps qsss fsss;
       
   346 
       
   347           fun mk_types maybe_dest_sumT fun_Ts =
       
   348             let
       
   349               val f_sum_prod_Ts = map range_type fun_Ts;
       
   350               val f_prod_Tss = map2 dest_sumTN_balanced ns f_sum_prod_Ts;
       
   351               val f_Tssss =
       
   352                 map3 (fn C => map2 (map (map (curry (op -->) C) o maybe_dest_sumT) oo dest_tupleT))
       
   353                   Cs mss' f_prod_Tss;
       
   354               val q_Tssss =
       
   355                 map (map (map (fn [_] => [] | [_, C] => [mk_pred1T (domain_type C)]))) f_Tssss;
       
   356               val pf_Tss = map3 flat_preds_predsss_gettersss p_Tss q_Tssss f_Tssss;
       
   357             in (q_Tssss, f_sum_prod_Ts, f_Tssss, pf_Tss) end;
       
   358 
       
   359           val (r_Tssss, g_sum_prod_Ts, g_Tssss, pg_Tss) = mk_types single fp_fold_fun_Ts;
       
   360 
       
   361           val ((((Free (z, _), cs), pss), gssss), lthy) =
       
   362             lthy
       
   363             |> yield_singleton (mk_Frees "z") dummyT
       
   364             ||>> mk_Frees "a" Cs
       
   365             ||>> mk_Freess "p" p_Tss
       
   366             ||>> mk_Freessss "g" g_Tssss;
       
   367           val rssss = map (map (map (fn [] => []))) r_Tssss;
       
   368 
       
   369           fun dest_corec_sumT (T as Type (@{type_name sum}, Us as [_, U])) =
       
   370               if member (op =) Cs U then Us else [T]
       
   371             | dest_corec_sumT T = [T];
       
   372 
       
   373           val (s_Tssss, h_sum_prod_Ts, h_Tssss, ph_Tss) = mk_types dest_corec_sumT fp_rec_fun_Ts;
       
   374 
       
   375           val hssss_hd = map2 (map2 (map2 (fn T :: _ => fn [g] => retype_free T g))) h_Tssss gssss;
       
   376           val ((sssss, hssss_tl), lthy) =
       
   377             lthy
       
   378             |> mk_Freessss "q" s_Tssss
       
   379             ||>> mk_Freessss "h" (map (map (map tl)) h_Tssss);
       
   380           val hssss = map2 (map2 (map2 cons)) hssss_hd hssss_tl;
       
   381 
       
   382           val cpss = map2 (fn c => map (fn p => p $ c)) cs pss;
       
   383 
       
   384           fun mk_preds_getters_join [] [cf] = cf
       
   385             | mk_preds_getters_join [cq] [cf, cf'] =
       
   386               mk_If cq (mk_Inl (fastype_of cf') cf) (mk_Inr (fastype_of cf) cf');
       
   387 
       
   388           fun mk_terms qssss fssss =
       
   389             let
       
   390               val pfss = map3 flat_preds_predsss_gettersss pss qssss fssss;
       
   391               val cqssss = map2 (fn c => map (map (map (fn f => f $ c)))) cs qssss;
       
   392               val cfssss = map2 (fn c => map (map (map (fn f => f $ c)))) cs fssss;
       
   393               val cqfsss = map2 (map2 (map2 mk_preds_getters_join)) cqssss cfssss;
       
   394             in (pfss, cqfsss) end;
       
   395         in
       
   396           (((([], [], []), ([], [], [])),
       
   397             ([z], cs, cpss, (mk_terms rssss gssss, (g_sum_prod_Ts, pg_Tss)),
       
   398              (mk_terms sssss hssss, (h_sum_prod_Ts, ph_Tss)))), lthy)
       
   399         end;
       
   400 
       
   401     fun define_ctrs_case_for_type (((((((((((((((((((((((((fp_bnf, fp_b), fpT), C), ctor), dtor),
       
   402             fp_fold), fp_rec), ctor_dtor), dtor_ctor), ctor_inject), pre_map_def), pre_set_defs),
       
   403           pre_rel_def), fp_map_thm), fp_set_thms), fp_rel_thm), n), ks), ms), ctr_bindings),
       
   404         ctr_mixfixes), ctr_Tss), disc_bindings), sel_bindingss), raw_sel_defaultss) no_defs_lthy =
       
   405       let
       
   406         val fp_b_name = Binding.name_of fp_b;
       
   407 
       
   408         val dtorT = domain_type (fastype_of ctor);
       
   409         val ctr_prod_Ts = map HOLogic.mk_tupleT ctr_Tss;
       
   410         val ctr_sum_prod_T = mk_sumTN_balanced ctr_prod_Ts;
       
   411         val case_Ts = map (fn Ts => Ts ---> C) ctr_Tss;
       
   412 
       
   413         val (((((w, fs), xss), yss), u'), names_lthy) =
       
   414           no_defs_lthy
       
   415           |> yield_singleton (mk_Frees "w") dtorT
       
   416           ||>> mk_Frees "f" case_Ts
       
   417           ||>> mk_Freess "x" ctr_Tss
       
   418           ||>> mk_Freess "y" (map (map B_ify) ctr_Tss)
       
   419           ||>> yield_singleton Variable.variant_fixes fp_b_name;
       
   420 
       
   421         val u = Free (u', fpT);
       
   422 
       
   423         val tuple_xs = map HOLogic.mk_tuple xss;
       
   424         val tuple_ys = map HOLogic.mk_tuple yss;
       
   425 
       
   426         val ctr_rhss =
       
   427           map3 (fn k => fn xs => fn tuple_x => fold_rev Term.lambda xs (ctor $
       
   428             mk_InN_balanced ctr_sum_prod_T n tuple_x k)) ks xss tuple_xs;
       
   429 
       
   430         val case_binding = qualify false fp_b_name (Binding.suffix_name ("_" ^ caseN) fp_b);
       
   431 
       
   432         val case_rhs =
       
   433           fold_rev Term.lambda (fs @ [u])
       
   434             (mk_sum_caseN_balanced (map2 mk_uncurried_fun fs xss) $ (dtor $ u));
       
   435 
       
   436         val ((raw_case :: raw_ctrs, raw_case_def :: raw_ctr_defs), (lthy', lthy)) = no_defs_lthy
       
   437           |> apfst split_list o fold_map3 (fn b => fn mx => fn rhs =>
       
   438               Local_Theory.define ((b, mx), ((Thm.def_binding b, []), rhs)) #>> apsnd snd)
       
   439             (case_binding :: ctr_bindings) (NoSyn :: ctr_mixfixes) (case_rhs :: ctr_rhss)
       
   440           ||> `Local_Theory.restore;
       
   441 
       
   442         val phi = Proof_Context.export_morphism lthy lthy';
       
   443 
       
   444         val ctr_defs = map (Morphism.thm phi) raw_ctr_defs;
       
   445         val ctr_defs' =
       
   446           map2 (fn m => fn def => mk_unabs_def m (def RS meta_eq_to_obj_eq)) ms ctr_defs;
       
   447         val case_def = Morphism.thm phi raw_case_def;
       
   448 
       
   449         val ctrs0 = map (Morphism.term phi) raw_ctrs;
       
   450         val casex0 = Morphism.term phi raw_case;
       
   451 
       
   452         val ctrs = map (mk_ctr As) ctrs0;
       
   453 
       
   454         fun wrap lthy =
       
   455           let
       
   456             fun exhaust_tac {context = ctxt, ...} =
       
   457               let
       
   458                 val ctor_iff_dtor_thm =
       
   459                   let
       
   460                     val goal =
       
   461                       fold_rev Logic.all [w, u]
       
   462                         (mk_Trueprop_eq (HOLogic.mk_eq (u, ctor $ w), HOLogic.mk_eq (dtor $ u, w)));
       
   463                   in
       
   464                     Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
       
   465                       mk_ctor_iff_dtor_tac ctxt (map (SOME o certifyT lthy) [dtorT, fpT])
       
   466                         (certify lthy ctor) (certify lthy dtor) ctor_dtor dtor_ctor)
       
   467                     |> Thm.close_derivation
       
   468                     |> Morphism.thm phi
       
   469                   end;
       
   470 
       
   471                 val sumEN_thm' =
       
   472                   unfold_thms lthy @{thms all_unit_eq}
       
   473                     (Drule.instantiate' (map (SOME o certifyT lthy) ctr_prod_Ts) []
       
   474                        (mk_sumEN_balanced n))
       
   475                   |> Morphism.thm phi;
       
   476               in
       
   477                 mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor_thm sumEN_thm'
       
   478               end;
       
   479 
       
   480             val inject_tacss =
       
   481               map2 (fn 0 => K [] | _ => fn ctr_def => [fn {context = ctxt, ...} =>
       
   482                   mk_inject_tac ctxt ctr_def ctor_inject]) ms ctr_defs;
       
   483 
       
   484             val half_distinct_tacss =
       
   485               map (map (fn (def, def') => fn {context = ctxt, ...} =>
       
   486                 mk_half_distinct_tac ctxt ctor_inject [def, def'])) (mk_half_pairss (`I ctr_defs));
       
   487 
       
   488             val case_tacs =
       
   489               map3 (fn k => fn m => fn ctr_def => fn {context = ctxt, ...} =>
       
   490                 mk_case_tac ctxt n k m case_def ctr_def dtor_ctor) ks ms ctr_defs;
       
   491 
       
   492             val tacss = [exhaust_tac] :: inject_tacss @ half_distinct_tacss @ [case_tacs];
       
   493 
       
   494             val sel_defaultss = map (map (apsnd (prepare_term lthy))) raw_sel_defaultss
       
   495           in
       
   496             wrap_datatype tacss (((wrap_opts, ctrs0), casex0), (disc_bindings, (sel_bindingss,
       
   497               sel_defaultss))) lthy
       
   498           end;
       
   499 
       
   500         fun derive_maps_sets_rels (wrap_res, lthy) =
       
   501           let
       
   502             val rel_flip = rel_flip_of_bnf fp_bnf;
       
   503             val nones = replicate live NONE;
       
   504 
       
   505             val ctor_cong =
       
   506               if lfp then Drule.dummy_thm
       
   507               else cterm_instantiate_pos [NONE, NONE, SOME (certify lthy ctor)] arg_cong;
       
   508 
       
   509             fun mk_cIn ify =
       
   510               certify lthy o (not lfp ? curry (op $) (map_types ify ctor)) oo
       
   511               mk_InN_balanced (ify ctr_sum_prod_T) n;
       
   512 
       
   513             val cxIns = map2 (mk_cIn I) tuple_xs ks;
       
   514             val cyIns = map2 (mk_cIn B_ify) tuple_ys ks;
       
   515 
       
   516             fun mk_map_thm ctr_def' xs cxIn =
       
   517               fold_thms lthy [ctr_def']
       
   518                 (unfold_thms lthy (pre_map_def ::
       
   519                      (if lfp then [] else [ctor_dtor, dtor_ctor]) @ sum_prod_thms_map)
       
   520                    (cterm_instantiate_pos (nones @ [SOME cxIn])
       
   521                       (if lfp then fp_map_thm else fp_map_thm RS ctor_cong)))
       
   522               |> singleton (Proof_Context.export names_lthy no_defs_lthy);
       
   523 
       
   524             fun mk_set_thm fp_set_thm ctr_def' xs cxIn =
       
   525               fold_thms lthy [ctr_def']
       
   526                 (unfold_thms lthy (pre_set_defs @ nested_set_natural's @ nesting_set_natural's @
       
   527                      (if lfp then [] else [dtor_ctor]) @ sum_prod_thms_set)
       
   528                    (cterm_instantiate_pos [SOME cxIn] fp_set_thm))
       
   529               |> singleton (Proof_Context.export names_lthy no_defs_lthy);
       
   530 
       
   531             fun mk_set_thms fp_set_thm = map3 (mk_set_thm fp_set_thm) ctr_defs' xss cxIns;
       
   532 
       
   533             val map_thms = map3 mk_map_thm ctr_defs' xss cxIns;
       
   534             val set_thmss = map mk_set_thms fp_set_thms;
       
   535 
       
   536             val rel_infos = (ctr_defs' ~~ xss ~~ cxIns, ctr_defs' ~~ yss ~~ cyIns);
       
   537 
       
   538             fun mk_rel_thm postproc ctr_defs' xs cxIn ys cyIn =
       
   539               fold_thms lthy ctr_defs'
       
   540                  (unfold_thms lthy (pre_rel_def :: (if lfp then [] else [dtor_ctor]) @
       
   541                       sum_prod_thms_rel)
       
   542                     (cterm_instantiate_pos (nones @ [SOME cxIn, SOME cyIn]) fp_rel_thm))
       
   543               |> postproc
       
   544               |> singleton (Proof_Context.export names_lthy no_defs_lthy);
       
   545 
       
   546             fun mk_rel_inject_thm (((ctr_def', xs), cxIn), ((_, ys), cyIn)) =
       
   547               mk_rel_thm (unfold_thms lthy @{thms eq_sym_Unity_conv}) [ctr_def'] xs cxIn ys cyIn;
       
   548 
       
   549             val rel_inject_thms = map mk_rel_inject_thm (op ~~ rel_infos);
       
   550 
       
   551             fun mk_half_rel_distinct_thm (((xctr_def', xs), cxIn), ((yctr_def', ys), cyIn)) =
       
   552               mk_rel_thm (fn thm => thm RS @{thm eq_False[THEN iffD1]}) [xctr_def', yctr_def']
       
   553                 xs cxIn ys cyIn;
       
   554 
       
   555             fun mk_other_half_rel_distinct_thm thm =
       
   556               flip_rels lthy live thm RS (rel_flip RS sym RS @{thm arg_cong[of _ _ Not]} RS iffD2);
       
   557 
       
   558             val half_rel_distinct_thmss =
       
   559               map (map mk_half_rel_distinct_thm) (mk_half_pairss rel_infos);
       
   560             val other_half_rel_distinct_thmss =
       
   561               map (map mk_other_half_rel_distinct_thm) half_rel_distinct_thmss;
       
   562             val (rel_distinct_thms, _) =
       
   563               join_halves n half_rel_distinct_thmss other_half_rel_distinct_thmss;
       
   564 
       
   565             val notes =
       
   566               [(mapN, map_thms, code_simp_attrs),
       
   567                (rel_distinctN, rel_distinct_thms, code_simp_attrs),
       
   568                (rel_injectN, rel_inject_thms, code_simp_attrs),
       
   569                (setsN, flat set_thmss, code_simp_attrs)]
       
   570               |> filter_out (null o #2)
       
   571               |> map (fn (thmN, thms, attrs) =>
       
   572                 ((qualify true fp_b_name (Binding.name thmN), attrs), [(thms, [])]));
       
   573           in
       
   574             (wrap_res, lthy |> Local_Theory.notes notes |> snd)
       
   575           end;
       
   576 
       
   577         fun define_fold_rec no_defs_lthy =
       
   578           let
       
   579             val fpT_to_C = fpT --> C;
       
   580 
       
   581             fun generate_rec_like (suf, fp_rec_like, (fss, f_Tss, xssss)) =
       
   582               let
       
   583                 val res_T = fold_rev (curry (op --->)) f_Tss fpT_to_C;
       
   584                 val binding = qualify false fp_b_name (Binding.suffix_name ("_" ^ suf) fp_b);
       
   585                 val spec =
       
   586                   mk_Trueprop_eq (lists_bmoc fss (Free (Binding.name_of binding, res_T)),
       
   587                     Term.list_comb (fp_rec_like,
       
   588                       map2 (mk_sum_caseN_balanced oo map2 mk_uncurried2_fun) fss xssss));
       
   589               in (binding, spec) end;
       
   590 
       
   591             val rec_like_infos =
       
   592               [(foldN, fp_fold, fold_only),
       
   593                (recN, fp_rec, rec_only)];
       
   594 
       
   595             val (bindings, specs) = map generate_rec_like rec_like_infos |> split_list;
       
   596 
       
   597             val ((csts, defs), (lthy', lthy)) = no_defs_lthy
       
   598               |> apfst split_list o fold_map2 (fn b => fn spec =>
       
   599                 Specification.definition (SOME (b, NONE, NoSyn), ((Thm.def_binding b, []), spec))
       
   600                 #>> apsnd snd) bindings specs
       
   601               ||> `Local_Theory.restore;
       
   602 
       
   603             val phi = Proof_Context.export_morphism lthy lthy';
       
   604 
       
   605             val [fold_def, rec_def] = map (Morphism.thm phi) defs;
       
   606 
       
   607             val [foldx, recx] = map (mk_rec_like lfp As Cs o Morphism.term phi) csts;
       
   608           in
       
   609             ((foldx, recx, fold_def, rec_def), lthy)
       
   610           end;
       
   611 
       
   612         fun define_unfold_corec no_defs_lthy =
       
   613           let
       
   614             val B_to_fpT = C --> fpT;
       
   615 
       
   616             fun mk_preds_getterss_join c n cps sum_prod_T cqfss =
       
   617               Term.lambda c (mk_IfN sum_prod_T cps
       
   618                 (map2 (mk_InN_balanced sum_prod_T n) (map HOLogic.mk_tuple cqfss) (1 upto n)));
       
   619 
       
   620             fun generate_corec_like (suf, fp_rec_like, ((pfss, cqfsss), (f_sum_prod_Ts,
       
   621                 pf_Tss))) =
       
   622               let
       
   623                 val res_T = fold_rev (curry (op --->)) pf_Tss B_to_fpT;
       
   624                 val binding = qualify false fp_b_name (Binding.suffix_name ("_" ^ suf) fp_b);
       
   625                 val spec =
       
   626                   mk_Trueprop_eq (lists_bmoc pfss (Free (Binding.name_of binding, res_T)),
       
   627                     Term.list_comb (fp_rec_like,
       
   628                       map5 mk_preds_getterss_join cs ns cpss f_sum_prod_Ts cqfsss));
       
   629               in (binding, spec) end;
       
   630 
       
   631             val corec_like_infos =
       
   632               [(unfoldN, fp_fold, unfold_only),
       
   633                (corecN, fp_rec, corec_only)];
       
   634 
       
   635             val (bindings, specs) = map generate_corec_like corec_like_infos |> split_list;
       
   636 
       
   637             val ((csts, defs), (lthy', lthy)) = no_defs_lthy
       
   638               |> apfst split_list o fold_map2 (fn b => fn spec =>
       
   639                 Specification.definition (SOME (b, NONE, NoSyn), ((Thm.def_binding b, []), spec))
       
   640                 #>> apsnd snd) bindings specs
       
   641               ||> `Local_Theory.restore;
       
   642 
       
   643             val phi = Proof_Context.export_morphism lthy lthy';
       
   644 
       
   645             val [unfold_def, corec_def] = map (Morphism.thm phi) defs;
       
   646 
       
   647             val [unfold, corec] = map (mk_rec_like lfp As Cs o Morphism.term phi) csts;
       
   648           in
       
   649             ((unfold, corec, unfold_def, corec_def), lthy)
       
   650           end;
       
   651 
       
   652         val define_rec_likes = if lfp then define_fold_rec else define_unfold_corec;
       
   653 
       
   654         fun massage_res ((wrap_res, rec_like_res), lthy) =
       
   655           (((ctrs, xss, ctr_defs, wrap_res), rec_like_res), lthy);
       
   656       in
       
   657         (wrap #> (live > 0 ? derive_maps_sets_rels) ##>> define_rec_likes #> massage_res, lthy')
       
   658       end;
       
   659 
       
   660     fun wrap_types_and_more (wrap_types_and_mores, lthy) =
       
   661       fold_map I wrap_types_and_mores lthy
       
   662       |>> apsnd split_list4 o apfst split_list4 o split_list;
       
   663 
       
   664     fun build_map build_arg (Type (s, Ts)) (Type (_, Us)) =
       
   665       let
       
   666         val bnf = the (bnf_of lthy s);
       
   667         val live = live_of_bnf bnf;
       
   668         val mapx = mk_map live Ts Us (map_of_bnf bnf);
       
   669         val TUs' = map dest_funT (fst (strip_typeN live (fastype_of mapx)));
       
   670       in Term.list_comb (mapx, map build_arg TUs') end;
       
   671 
       
   672     (* TODO: Add map, sets, rel simps *)
       
   673     val mk_simp_thmss =
       
   674       map3 (fn (_, _, _, injects, distincts, cases, _, _, _) => fn rec_likes => fn fold_likes =>
       
   675         injects @ distincts @ cases @ rec_likes @ fold_likes);
       
   676 
       
   677     fun derive_induct_fold_rec_thms_for_types (((ctrss, xsss, ctr_defss, wrap_ress), (folds, recs,
       
   678         fold_defs, rec_defs)), lthy) =
       
   679       let
       
   680         val (((ps, ps'), us'), names_lthy) =
       
   681           lthy
       
   682           |> mk_Frees' "P" (map mk_pred1T fpTs)
       
   683           ||>> Variable.variant_fixes fp_b_names;
       
   684 
       
   685         val us = map2 (curry Free) us' fpTs;
       
   686 
       
   687         fun mk_sets_nested bnf =
       
   688           let
       
   689             val Type (T_name, Us) = T_of_bnf bnf;
       
   690             val lives = lives_of_bnf bnf;
       
   691             val sets = sets_of_bnf bnf;
       
   692             fun mk_set U =
       
   693               (case find_index (curry (op =) U) lives of
       
   694                 ~1 => Term.dummy
       
   695               | i => nth sets i);
       
   696           in
       
   697             (T_name, map mk_set Us)
       
   698           end;
       
   699 
       
   700         val setss_nested = map mk_sets_nested nested_bnfs;
       
   701 
       
   702         val (induct_thms, induct_thm) =
       
   703           let
       
   704             fun mk_set Ts t =
       
   705               let val Type (_, Ts0) = domain_type (fastype_of t) in
       
   706                 Term.subst_atomic_types (Ts0 ~~ Ts) t
       
   707               end;
       
   708 
       
   709             fun mk_raw_prem_prems names_lthy (x as Free (s, T as Type (T_name, Ts0))) =
       
   710                 (case find_index (curry (op =) T) fpTs of
       
   711                   ~1 =>
       
   712                   (case AList.lookup (op =) setss_nested T_name of
       
   713                     NONE => []
       
   714                   | SOME raw_sets0 =>
       
   715                     let
       
   716                       val (Ts, raw_sets) =
       
   717                         split_list (filter (exists_fp_subtype o fst) (Ts0 ~~ raw_sets0));
       
   718                       val sets = map (mk_set Ts0) raw_sets;
       
   719                       val (ys, names_lthy') = names_lthy |> mk_Frees s Ts;
       
   720                       val xysets = map (pair x) (ys ~~ sets);
       
   721                       val ppremss = map (mk_raw_prem_prems names_lthy') ys;
       
   722                     in
       
   723                       flat (map2 (map o apfst o cons) xysets ppremss)
       
   724                     end)
       
   725                 | kk => [([], (kk + 1, x))])
       
   726               | mk_raw_prem_prems _ _ = [];
       
   727 
       
   728             fun close_prem_prem xs t =
       
   729               fold_rev Logic.all (map Free (drop (nn + length xs)
       
   730                 (rev (Term.add_frees t (map dest_Free xs @ ps'))))) t;
       
   731 
       
   732             fun mk_prem_prem xs (xysets, (j, x)) =
       
   733               close_prem_prem xs (Logic.list_implies (map (fn (x', (y, set)) =>
       
   734                   HOLogic.mk_Trueprop (HOLogic.mk_mem (y, set $ x'))) xysets,
       
   735                 HOLogic.mk_Trueprop (nth ps (j - 1) $ x)));
       
   736 
       
   737             fun mk_raw_prem phi ctr ctr_Ts =
       
   738               let
       
   739                 val (xs, names_lthy') = names_lthy |> mk_Frees "x" ctr_Ts;
       
   740                 val pprems = maps (mk_raw_prem_prems names_lthy') xs;
       
   741               in (xs, pprems, HOLogic.mk_Trueprop (phi $ Term.list_comb (ctr, xs))) end;
       
   742 
       
   743             fun mk_prem (xs, raw_pprems, concl) =
       
   744               fold_rev Logic.all xs (Logic.list_implies (map (mk_prem_prem xs) raw_pprems, concl));
       
   745 
       
   746             val raw_premss = map3 (map2 o mk_raw_prem) ps ctrss ctr_Tsss;
       
   747 
       
   748             val goal =
       
   749               Library.foldr (Logic.list_implies o apfst (map mk_prem)) (raw_premss,
       
   750                 HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry (op $)) ps us)));
       
   751 
       
   752             val kksss = map (map (map (fst o snd) o #2)) raw_premss;
       
   753 
       
   754             val ctor_induct' = fp_induct OF (map mk_sumEN_tupled_balanced mss);
       
   755 
       
   756             val thm =
       
   757               Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
       
   758                 mk_induct_tac ctxt nn ns mss kksss (flat ctr_defss) ctor_induct'
       
   759                   nested_set_natural's pre_set_defss)
       
   760               |> singleton (Proof_Context.export names_lthy lthy)
       
   761               |> Thm.close_derivation;
       
   762           in
       
   763             `(conj_dests nn) thm
       
   764           end;
       
   765 
       
   766         val induct_cases = quasi_unambiguous_case_names (maps (map name_of_ctr) ctrss);
       
   767 
       
   768         val (fold_thmss, rec_thmss) =
       
   769           let
       
   770             val xctrss = map2 (map2 (curry Term.list_comb)) ctrss xsss;
       
   771             val gfolds = map (lists_bmoc gss) folds;
       
   772             val hrecs = map (lists_bmoc hss) recs;
       
   773 
       
   774             fun mk_goal fss frec_like xctr f xs fxs =
       
   775               fold_rev (fold_rev Logic.all) (xs :: fss)
       
   776                 (mk_Trueprop_eq (frec_like $ xctr, Term.list_comb (f, fxs)));
       
   777 
       
   778             fun build_rec_like frec_likes maybe_tick (T, U) =
       
   779               if T = U then
       
   780                 id_const T
       
   781               else
       
   782                 (case find_index (curry (op =) T) fpTs of
       
   783                   ~1 => build_map (build_rec_like frec_likes maybe_tick) T U
       
   784                 | kk => maybe_tick (nth us kk) (nth frec_likes kk));
       
   785 
       
   786             fun mk_U maybe_mk_prodT =
       
   787               typ_subst (map2 (fn fpT => fn C => (fpT, maybe_mk_prodT fpT C)) fpTs Cs);
       
   788 
       
   789             fun intr_rec_likes frec_likes maybe_cons maybe_tick maybe_mk_prodT (x as Free (_, T)) =
       
   790               if member (op =) fpTs T then
       
   791                 maybe_cons x [build_rec_like frec_likes (K I) (T, mk_U (K I) T) $ x]
       
   792               else if exists_fp_subtype T then
       
   793                 [build_rec_like frec_likes maybe_tick (T, mk_U maybe_mk_prodT T) $ x]
       
   794               else
       
   795                 [x];
       
   796 
       
   797             val gxsss = map (map (maps (intr_rec_likes gfolds (K I) (K I) (K I)))) xsss;
       
   798             val hxsss =
       
   799               map (map (maps (intr_rec_likes hrecs cons tick (curry HOLogic.mk_prodT)))) xsss;
       
   800 
       
   801             val fold_goalss = map5 (map4 o mk_goal gss) gfolds xctrss gss xsss gxsss;
       
   802             val rec_goalss = map5 (map4 o mk_goal hss) hrecs xctrss hss xsss hxsss;
       
   803 
       
   804             val fold_tacss =
       
   805               map2 (map o mk_rec_like_tac pre_map_defs nesting_map_ids fold_defs) fp_fold_thms
       
   806                 ctr_defss;
       
   807             val rec_tacss =
       
   808               map2 (map o mk_rec_like_tac pre_map_defs nesting_map_ids rec_defs) fp_rec_thms
       
   809                 ctr_defss;
       
   810 
       
   811             fun prove goal tac =
       
   812               Skip_Proof.prove lthy [] [] goal (tac o #context)
       
   813               |> Thm.close_derivation;
       
   814           in
       
   815             (map2 (map2 prove) fold_goalss fold_tacss, map2 (map2 prove) rec_goalss rec_tacss)
       
   816           end;
       
   817 
       
   818         val simp_thmss = mk_simp_thmss wrap_ress rec_thmss fold_thmss;
       
   819 
       
   820         val induct_case_names_attr = Attrib.internal (K (Rule_Cases.case_names induct_cases));
       
   821         fun induct_type_attr T_name = Attrib.internal (K (Induct.induct_type T_name));
       
   822 
       
   823         val common_notes =
       
   824           (if nn > 1 then [(inductN, [induct_thm], [induct_case_names_attr])] else [])
       
   825           |> map (fn (thmN, thms, attrs) =>
       
   826             ((qualify true fp_common_name (Binding.name thmN), attrs), [(thms, [])]));
       
   827 
       
   828         val notes =
       
   829           [(foldN, fold_thmss, K code_simp_attrs),
       
   830            (inductN, map single induct_thms,
       
   831             fn T_name => [induct_case_names_attr, induct_type_attr T_name]),
       
   832            (recN, rec_thmss, K code_simp_attrs),
       
   833            (simpsN, simp_thmss, K [])]
       
   834           |> maps (fn (thmN, thmss, attrs) =>
       
   835             map3 (fn fp_b_name => fn Type (T_name, _) => fn thms =>
       
   836               ((qualify true fp_b_name (Binding.name thmN), attrs T_name),
       
   837                [(thms, [])])) fp_b_names fpTs thmss);
       
   838       in
       
   839         lthy |> Local_Theory.notes (common_notes @ notes) |> snd
       
   840       end;
       
   841 
       
   842     fun derive_coinduct_unfold_corec_thms_for_types (((ctrss, _, ctr_defss, wrap_ress), (unfolds,
       
   843         corecs, unfold_defs, corec_defs)), lthy) =
       
   844       let
       
   845         val nesting_rel_eqs = map rel_eq_of_bnf nesting_bnfs;
       
   846 
       
   847         val discss = map (map (mk_disc_or_sel As) o #1) wrap_ress;
       
   848         val selsss = map (map (map (mk_disc_or_sel As)) o #2) wrap_ress;
       
   849         val exhaust_thms = map #3 wrap_ress;
       
   850         val disc_thmsss = map #7 wrap_ress;
       
   851         val discIss = map #8 wrap_ress;
       
   852         val sel_thmsss = map #9 wrap_ress;
       
   853 
       
   854         val (((rs, us'), vs'), names_lthy) =
       
   855           lthy
       
   856           |> mk_Frees "R" (map (fn T => mk_pred2T T T) fpTs)
       
   857           ||>> Variable.variant_fixes fp_b_names
       
   858           ||>> Variable.variant_fixes (map (suffix "'") fp_b_names);
       
   859 
       
   860         val us = map2 (curry Free) us' fpTs;
       
   861         val udiscss = map2 (map o rapp) us discss;
       
   862         val uselsss = map2 (map o map o rapp) us selsss;
       
   863 
       
   864         val vs = map2 (curry Free) vs' fpTs;
       
   865         val vdiscss = map2 (map o rapp) vs discss;
       
   866         val vselsss = map2 (map o map o rapp) vs selsss;
       
   867 
       
   868         val ((coinduct_thms, coinduct_thm), (strong_coinduct_thms, strong_coinduct_thm)) =
       
   869           let
       
   870             val uvrs = map3 (fn r => fn u => fn v => r $ u $ v) rs us vs;
       
   871             val uv_eqs = map2 (curry HOLogic.mk_eq) us vs;
       
   872             val strong_rs =
       
   873               map4 (fn u => fn v => fn uvr => fn uv_eq =>
       
   874                 fold_rev Term.lambda [u, v] (HOLogic.mk_disj (uvr, uv_eq))) us vs uvrs uv_eqs;
       
   875 
       
   876             fun build_rel_step build_arg (Type (s, Ts)) =
       
   877               let
       
   878                 val bnf = the (bnf_of lthy s);
       
   879                 val live = live_of_bnf bnf;
       
   880                 val rel = mk_rel live Ts Ts (rel_of_bnf bnf);
       
   881                 val Ts' = map domain_type (fst (strip_typeN live (fastype_of rel)));
       
   882               in Term.list_comb (rel, map build_arg Ts') end;
       
   883 
       
   884             fun build_rel rs' T =
       
   885               (case find_index (curry (op =) T) fpTs of
       
   886                 ~1 =>
       
   887                 if exists_fp_subtype T then build_rel_step (build_rel rs') T
       
   888                 else HOLogic.eq_const T
       
   889               | kk => nth rs' kk);
       
   890 
       
   891             fun build_rel_app rs' usel vsel =
       
   892               fold rapp [usel, vsel] (build_rel rs' (fastype_of usel));
       
   893 
       
   894             fun mk_prem_ctr_concls rs' n k udisc usels vdisc vsels =
       
   895               (if k = n then [] else [HOLogic.mk_eq (udisc, vdisc)]) @
       
   896               (if null usels then
       
   897                  []
       
   898                else
       
   899                  [Library.foldr HOLogic.mk_imp (if n = 1 then [] else [udisc, vdisc],
       
   900                     Library.foldr1 HOLogic.mk_conj (map2 (build_rel_app rs') usels vsels))]);
       
   901 
       
   902             fun mk_prem_concl rs' n udiscs uselss vdiscs vselss =
       
   903               Library.foldr1 HOLogic.mk_conj
       
   904                 (flat (map5 (mk_prem_ctr_concls rs' n) (1 upto n) udiscs uselss vdiscs vselss))
       
   905               handle List.Empty => @{term True};
       
   906 
       
   907             fun mk_prem rs' uvr u v n udiscs uselss vdiscs vselss =
       
   908               fold_rev Logic.all [u, v] (Logic.mk_implies (HOLogic.mk_Trueprop uvr,
       
   909                 HOLogic.mk_Trueprop (mk_prem_concl rs' n udiscs uselss vdiscs vselss)));
       
   910 
       
   911             val concl =
       
   912               HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
       
   913                 (map3 (fn uvr => fn u => fn v => HOLogic.mk_imp (uvr, HOLogic.mk_eq (u, v)))
       
   914                    uvrs us vs));
       
   915 
       
   916             fun mk_goal rs' =
       
   917               Logic.list_implies (map8 (mk_prem rs') uvrs us vs ns udiscss uselsss vdiscss vselsss,
       
   918                 concl);
       
   919 
       
   920             val goal = mk_goal rs;
       
   921             val strong_goal = mk_goal strong_rs;
       
   922 
       
   923             fun prove dtor_coinduct' goal =
       
   924               Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
       
   925                 mk_coinduct_tac ctxt nesting_rel_eqs nn ns dtor_coinduct' pre_rel_defs dtor_ctors
       
   926                   exhaust_thms ctr_defss disc_thmsss sel_thmsss)
       
   927               |> singleton (Proof_Context.export names_lthy lthy)
       
   928               |> Thm.close_derivation;
       
   929 
       
   930             fun postproc nn thm =
       
   931               Thm.permute_prems 0 nn
       
   932                 (if nn = 1 then thm RS mp
       
   933                  else funpow nn (fn thm => reassoc_conjs (thm RS mp_conj)) thm)
       
   934               |> Drule.zero_var_indexes
       
   935               |> `(conj_dests nn);
       
   936           in
       
   937             (postproc nn (prove fp_induct goal), postproc nn (prove fp_strong_induct strong_goal))
       
   938           end;
       
   939 
       
   940         fun mk_maybe_not pos = not pos ? HOLogic.mk_not;
       
   941 
       
   942         val z = the_single zs;
       
   943         val gunfolds = map (lists_bmoc pgss) unfolds;
       
   944         val hcorecs = map (lists_bmoc phss) corecs;
       
   945 
       
   946         val (unfold_thmss, corec_thmss, safe_unfold_thmss, safe_corec_thmss) =
       
   947           let
       
   948             fun mk_goal pfss c cps fcorec_like n k ctr m cfs' =
       
   949               fold_rev (fold_rev Logic.all) ([c] :: pfss)
       
   950                 (Logic.list_implies (seq_conds (HOLogic.mk_Trueprop oo mk_maybe_not) n k cps,
       
   951                    mk_Trueprop_eq (fcorec_like $ c, Term.list_comb (ctr, take m cfs'))));
       
   952 
       
   953             fun build_corec_like fcorec_likes maybe_tack (T, U) =
       
   954               if T = U then
       
   955                 id_const T
       
   956               else
       
   957                 (case find_index (curry (op =) U) fpTs of
       
   958                   ~1 => build_map (build_corec_like fcorec_likes maybe_tack) T U
       
   959                 | kk => maybe_tack (nth cs kk, nth us kk) (nth fcorec_likes kk));
       
   960 
       
   961             fun mk_U maybe_mk_sumT =
       
   962               typ_subst (map2 (fn C => fn fpT => (maybe_mk_sumT fpT C, fpT)) Cs fpTs);
       
   963 
       
   964             fun intr_corec_likes fcorec_likes maybe_mk_sumT maybe_tack cqf =
       
   965               let val T = fastype_of cqf in
       
   966                 if exists_subtype (member (op =) Cs) T then
       
   967                   build_corec_like fcorec_likes maybe_tack (T, mk_U maybe_mk_sumT T) $ cqf
       
   968                 else
       
   969                   cqf
       
   970               end;
       
   971 
       
   972             val crgsss' = map (map (map (intr_corec_likes gunfolds (K I) (K I)))) crgsss;
       
   973             val cshsss' =
       
   974               map (map (map (intr_corec_likes hcorecs (curry mk_sumT) (tack z)))) cshsss;
       
   975 
       
   976             val unfold_goalss =
       
   977               map8 (map4 oooo mk_goal pgss) cs cpss gunfolds ns kss ctrss mss crgsss';
       
   978             val corec_goalss =
       
   979               map8 (map4 oooo mk_goal phss) cs cpss hcorecs ns kss ctrss mss cshsss';
       
   980 
       
   981             val unfold_tacss =
       
   982               map3 (map oo mk_corec_like_tac unfold_defs nesting_map_ids) fp_fold_thms pre_map_defs
       
   983                 ctr_defss;
       
   984             val corec_tacss =
       
   985               map3 (map oo mk_corec_like_tac corec_defs nesting_map_ids) fp_rec_thms pre_map_defs
       
   986                 ctr_defss;
       
   987 
       
   988             fun prove goal tac =
       
   989               Skip_Proof.prove lthy [] [] goal (tac o #context) |> Thm.close_derivation;
       
   990 
       
   991             val unfold_thmss = map2 (map2 prove) unfold_goalss unfold_tacss;
       
   992             val corec_thmss =
       
   993               map2 (map2 prove) corec_goalss corec_tacss
       
   994               |> map (map (unfold_thms lthy @{thms sum_case_if}));
       
   995 
       
   996             val unfold_safesss = map2 (map2 (map2 (curry (op =)))) crgsss' crgsss;
       
   997             val corec_safesss = map2 (map2 (map2 (curry (op =)))) cshsss' cshsss;
       
   998 
       
   999             val filter_safesss =
       
  1000               map2 (map_filter (fn (safes, thm) => if forall I safes then SOME thm else NONE) oo
       
  1001                 curry (op ~~));
       
  1002 
       
  1003             val safe_unfold_thmss = filter_safesss unfold_safesss unfold_thmss;
       
  1004             val safe_corec_thmss = filter_safesss corec_safesss corec_thmss;
       
  1005           in
       
  1006             (unfold_thmss, corec_thmss, safe_unfold_thmss, safe_corec_thmss)
       
  1007           end;
       
  1008 
       
  1009         val (disc_unfold_iff_thmss, disc_corec_iff_thmss) =
       
  1010           let
       
  1011             fun mk_goal c cps fcorec_like n k disc =
       
  1012               mk_Trueprop_eq (disc $ (fcorec_like $ c),
       
  1013                 if n = 1 then @{const True}
       
  1014                 else Library.foldr1 HOLogic.mk_conj (seq_conds mk_maybe_not n k cps));
       
  1015 
       
  1016             val unfold_goalss = map6 (map2 oooo mk_goal) cs cpss gunfolds ns kss discss;
       
  1017             val corec_goalss = map6 (map2 oooo mk_goal) cs cpss hcorecs ns kss discss;
       
  1018 
       
  1019             fun mk_case_split' cp =
       
  1020               Drule.instantiate' [] [SOME (certify lthy cp)] @{thm case_split};
       
  1021 
       
  1022             val case_splitss' = map (map mk_case_split') cpss;
       
  1023 
       
  1024             val unfold_tacss =
       
  1025               map3 (map oo mk_disc_corec_like_iff_tac) case_splitss' unfold_thmss disc_thmsss;
       
  1026             val corec_tacss =
       
  1027               map3 (map oo mk_disc_corec_like_iff_tac) case_splitss' corec_thmss disc_thmsss;
       
  1028 
       
  1029             fun prove goal tac =
       
  1030               Skip_Proof.prove lthy [] [] goal (tac o #context)
       
  1031               |> singleton (Proof_Context.export names_lthy0 no_defs_lthy)
       
  1032               |> Thm.close_derivation;
       
  1033 
       
  1034             fun proves [_] [_] = []
       
  1035               | proves goals tacs = map2 prove goals tacs;
       
  1036           in
       
  1037             (map2 proves unfold_goalss unfold_tacss,
       
  1038              map2 proves corec_goalss corec_tacss)
       
  1039           end;
       
  1040 
       
  1041         val is_triv_discI = is_triv_implies orf is_concl_refl;
       
  1042 
       
  1043         fun mk_disc_corec_like_thms corec_likes discIs =
       
  1044           map (op RS) (filter_out (is_triv_discI o snd) (corec_likes ~~ discIs));
       
  1045 
       
  1046         val disc_unfold_thmss = map2 mk_disc_corec_like_thms unfold_thmss discIss;
       
  1047         val disc_corec_thmss = map2 mk_disc_corec_like_thms corec_thmss discIss;
       
  1048 
       
  1049         fun mk_sel_corec_like_thm corec_like_thm sel sel_thm =
       
  1050           let
       
  1051             val (domT, ranT) = dest_funT (fastype_of sel);
       
  1052             val arg_cong' =
       
  1053               Drule.instantiate' (map (SOME o certifyT lthy) [domT, ranT])
       
  1054                 [NONE, NONE, SOME (certify lthy sel)] arg_cong
       
  1055               |> Thm.varifyT_global;
       
  1056             val sel_thm' = sel_thm RSN (2, trans);
       
  1057           in
       
  1058             corec_like_thm RS arg_cong' RS sel_thm'
       
  1059           end;
       
  1060 
       
  1061         fun mk_sel_corec_like_thms corec_likess =
       
  1062           map3 (map3 (map2 o mk_sel_corec_like_thm)) corec_likess selsss sel_thmsss |> map flat;
       
  1063 
       
  1064         val sel_unfold_thmss = mk_sel_corec_like_thms unfold_thmss;
       
  1065         val sel_corec_thmss = mk_sel_corec_like_thms corec_thmss;
       
  1066 
       
  1067         fun flat_corec_like_thms corec_likes disc_corec_likes sel_corec_likes =
       
  1068           corec_likes @ disc_corec_likes @ sel_corec_likes;
       
  1069 
       
  1070         val simp_thmss =
       
  1071           mk_simp_thmss wrap_ress
       
  1072             (map3 flat_corec_like_thms safe_corec_thmss disc_corec_thmss sel_corec_thmss)
       
  1073             (map3 flat_corec_like_thms safe_unfold_thmss disc_unfold_thmss sel_unfold_thmss);
       
  1074 
       
  1075         val anonymous_notes =
       
  1076           [(flat safe_unfold_thmss @ flat safe_corec_thmss, simp_attrs)]
       
  1077           |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
       
  1078 
       
  1079         val common_notes =
       
  1080           (if nn > 1 then
       
  1081              (* FIXME: attribs *)
       
  1082              [(coinductN, [coinduct_thm], []),
       
  1083               (strong_coinductN, [strong_coinduct_thm], [])]
       
  1084            else
       
  1085              [])
       
  1086           |> map (fn (thmN, thms, attrs) =>
       
  1087             ((qualify true fp_common_name (Binding.name thmN), attrs), [(thms, [])]));
       
  1088 
       
  1089         val notes =
       
  1090           [(coinductN, map single coinduct_thms, []), (* FIXME: attribs *)
       
  1091            (corecN, corec_thmss, []),
       
  1092            (disc_corecN, disc_corec_thmss, simp_attrs),
       
  1093            (disc_corec_iffN, disc_corec_iff_thmss, simp_attrs),
       
  1094            (disc_unfoldN, disc_unfold_thmss, simp_attrs),
       
  1095            (disc_unfold_iffN, disc_unfold_iff_thmss, simp_attrs),
       
  1096            (sel_corecN, sel_corec_thmss, simp_attrs),
       
  1097            (sel_unfoldN, sel_unfold_thmss, simp_attrs),
       
  1098            (simpsN, simp_thmss, []),
       
  1099            (strong_coinductN, map single strong_coinduct_thms, []), (* FIXME: attribs *)
       
  1100            (unfoldN, unfold_thmss, [])]
       
  1101           |> maps (fn (thmN, thmss, attrs) =>
       
  1102             map_filter (fn (_, []) => NONE | (fp_b_name, thms) =>
       
  1103               SOME ((qualify true fp_b_name (Binding.name thmN), attrs),
       
  1104                 [(thms, [])])) (fp_b_names ~~ thmss));
       
  1105       in
       
  1106         lthy |> Local_Theory.notes (anonymous_notes @ common_notes @ notes) |> snd
       
  1107       end;
       
  1108 
       
  1109     val lthy' = lthy
       
  1110       |> fold_map define_ctrs_case_for_type (fp_bnfs ~~ fp_bs ~~ fpTs ~~ Cs ~~ ctors ~~ dtors ~~
       
  1111         fp_folds ~~ fp_recs ~~ ctor_dtors ~~ dtor_ctors ~~ ctor_injects ~~ pre_map_defs ~~
       
  1112         pre_set_defss ~~ pre_rel_defs ~~ fp_map_thms ~~ fp_set_thmss ~~ fp_rel_thms ~~ ns ~~ kss ~~
       
  1113         mss ~~ ctr_bindingss ~~ ctr_mixfixess ~~ ctr_Tsss ~~ disc_bindingss ~~ sel_bindingsss ~~
       
  1114         raw_sel_defaultsss)
       
  1115       |> wrap_types_and_more
       
  1116       |> (if lfp then derive_induct_fold_rec_thms_for_types
       
  1117           else derive_coinduct_unfold_corec_thms_for_types);
       
  1118 
       
  1119     val timer = time (timer ("Constructors, discriminators, selectors, etc., for the new " ^
       
  1120       (if lfp then "" else "co") ^ "datatype"));
       
  1121   in
       
  1122     timer; lthy'
       
  1123   end;
       
  1124 
       
  1125 val datatypes = define_datatypes (K I) (K I) (K I);
       
  1126 
       
  1127 val datatype_cmd = define_datatypes Typedecl.read_constraint Syntax.parse_typ Syntax.read_term;
       
  1128 
       
  1129 val parse_ctr_arg =
       
  1130   @{keyword "("} |-- parse_binding_colon -- Parse.typ --| @{keyword ")"} ||
       
  1131   (Parse.typ >> pair Binding.empty);
       
  1132 
       
  1133 val parse_defaults =
       
  1134   @{keyword "("} |-- @{keyword "defaults"} |-- Scan.repeat parse_bound_term --| @{keyword ")"};
       
  1135 
       
  1136 val parse_single_spec =
       
  1137   Parse.type_args_constrained -- Parse.binding -- Parse.opt_mixfix --
       
  1138   (@{keyword "="} |-- Parse.enum1 "|" (parse_opt_binding_colon -- Parse.binding --
       
  1139     Scan.repeat parse_ctr_arg -- Scan.optional parse_defaults [] -- Parse.opt_mixfix));
       
  1140 
       
  1141 val parse_datatype = parse_wrap_options -- Parse.and_list1 parse_single_spec;
       
  1142 
       
  1143 fun parse_datatype_cmd lfp construct_fp = parse_datatype >> datatype_cmd lfp construct_fp;
       
  1144 
       
  1145 end;