src/HOL/BNF/Tools/bnf_lfp.ML
changeset 55058 4e700eb471d4
parent 55057 6b0fcbeebaba
child 55059 ef2e0fb783c6
equal deleted inserted replaced
55057:6b0fcbeebaba 55058:4e700eb471d4
     1 (*  Title:      HOL/BNF/Tools/bnf_lfp.ML
       
     2     Author:     Dmitriy Traytel, TU Muenchen
       
     3     Author:     Andrei Popescu, TU Muenchen
       
     4     Copyright   2012
       
     5 
       
     6 Datatype construction.
       
     7 *)
       
     8 
       
     9 signature BNF_LFP =
       
    10 sig
       
    11   val construct_lfp: mixfix list -> binding list -> binding list -> binding list list ->
       
    12     binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list ->
       
    13     local_theory -> BNF_FP_Util.fp_result * local_theory
       
    14 end;
       
    15 
       
    16 structure BNF_LFP : BNF_LFP =
       
    17 struct
       
    18 
       
    19 open BNF_Def
       
    20 open BNF_Util
       
    21 open BNF_Tactics
       
    22 open BNF_Comp
       
    23 open BNF_FP_Util
       
    24 open BNF_FP_Def_Sugar
       
    25 open BNF_LFP_Rec_Sugar
       
    26 open BNF_LFP_Util
       
    27 open BNF_LFP_Tactics
       
    28 
       
    29 (*all BNFs have the same lives*)
       
    30 fun construct_lfp mixfixes map_bs rel_bs set_bss0 bs resBs (resDs, Dss) bnfs lthy =
       
    31   let
       
    32     val time = time lthy;
       
    33     val timer = time (Timer.startRealTimer ());
       
    34 
       
    35     val live = live_of_bnf (hd bnfs);
       
    36     val n = length bnfs; (*active*)
       
    37     val ks = 1 upto n;
       
    38     val m = live - n; (*passive, if 0 don't generate a new BNF*)
       
    39 
       
    40     val note_all = Config.get lthy bnf_note_all;
       
    41     val b_names = map Binding.name_of bs;
       
    42     val b_name = mk_common_name b_names;
       
    43     val b = Binding.name b_name;
       
    44     val mk_internal_b = Binding.name #> Binding.prefix true b_name #> Binding.conceal;
       
    45     fun mk_internal_bs name =
       
    46       map (fn b =>
       
    47         Binding.prefix true b_name (Binding.prefix_name (name ^ "_") b) |> Binding.conceal) bs;
       
    48     val external_bs = map2 (Binding.prefix false) b_names bs
       
    49       |> note_all = false ? map Binding.conceal;
       
    50 
       
    51     (* TODO: check if m, n, etc., are sane *)
       
    52 
       
    53     val deads = fold (union (op =)) Dss resDs;
       
    54     val names_lthy = fold Variable.declare_typ deads lthy;
       
    55     val passives = map fst (subtract (op = o apsnd TFree) deads resBs);
       
    56 
       
    57     (* tvars *)
       
    58     val (((((passiveAs, activeAs), passiveBs), activeBs), passiveCs), activeCs) =
       
    59       names_lthy
       
    60       |> variant_tfrees passives
       
    61       ||>> mk_TFrees n
       
    62       ||>> variant_tfrees passives
       
    63       ||>> mk_TFrees n
       
    64       ||>> variant_tfrees passives
       
    65       ||>> mk_TFrees n
       
    66       |> fst;
       
    67 
       
    68     val allAs = passiveAs @ activeAs;
       
    69     val allBs' = passiveBs @ activeBs;
       
    70     val Ass = replicate n allAs;
       
    71     val allBs = passiveAs @ activeBs;
       
    72     val Bss = replicate n allBs;
       
    73     val allCs = passiveAs @ activeCs;
       
    74     val allCs' = passiveBs @ activeCs;
       
    75     val Css' = replicate n allCs';
       
    76 
       
    77     (* types *)
       
    78     val dead_poss =
       
    79       map (fn x => if member (op =) deads (TFree x) then SOME (TFree x) else NONE) resBs;
       
    80     fun mk_param NONE passive = (hd passive, tl passive)
       
    81       | mk_param (SOME a) passive = (a, passive);
       
    82     val mk_params = fold_map mk_param dead_poss #> fst;
       
    83 
       
    84     fun mk_FTs Ts = map2 (fn Ds => mk_T_of_bnf Ds Ts) Dss bnfs;
       
    85     val (params, params') = `(map Term.dest_TFree) (mk_params passiveAs);
       
    86     val FTsAs = mk_FTs allAs;
       
    87     val FTsBs = mk_FTs allBs;
       
    88     val FTsCs = mk_FTs allCs;
       
    89     val ATs = map HOLogic.mk_setT passiveAs;
       
    90     val BTs = map HOLogic.mk_setT activeAs;
       
    91     val B'Ts = map HOLogic.mk_setT activeBs;
       
    92     val B''Ts = map HOLogic.mk_setT activeCs;
       
    93     val sTs = map2 (curry op -->) FTsAs activeAs;
       
    94     val s'Ts = map2 (curry op -->) FTsBs activeBs;
       
    95     val s''Ts = map2 (curry op -->) FTsCs activeCs;
       
    96     val fTs = map2 (curry op -->) activeAs activeBs;
       
    97     val inv_fTs = map2 (curry op -->) activeBs activeAs;
       
    98     val self_fTs = map2 (curry op -->) activeAs activeAs;
       
    99     val gTs = map2 (curry op -->) activeBs activeCs;
       
   100     val all_gTs = map2 (curry op -->) allBs allCs';
       
   101     val prodBsAs = map2 (curry HOLogic.mk_prodT) activeBs activeAs;
       
   102     val prodFTs = mk_FTs (passiveAs @ prodBsAs);
       
   103     val prod_sTs = map2 (curry op -->) prodFTs activeAs;
       
   104 
       
   105     (* terms *)
       
   106     val mapsAsAs = map4 mk_map_of_bnf Dss Ass Ass bnfs;
       
   107     val mapsAsBs = map4 mk_map_of_bnf Dss Ass Bss bnfs;
       
   108     val mapsBsAs = map4 mk_map_of_bnf Dss Bss Ass bnfs;
       
   109     val mapsBsCs' = map4 mk_map_of_bnf Dss Bss Css' bnfs;
       
   110     val mapsAsCs' = map4 mk_map_of_bnf Dss Ass Css' bnfs;
       
   111     val map_fsts = map4 mk_map_of_bnf Dss (replicate n (passiveAs @ prodBsAs)) Bss bnfs;
       
   112     val map_fsts_rev = map4 mk_map_of_bnf Dss Bss (replicate n (passiveAs @ prodBsAs)) bnfs;
       
   113     fun mk_setss Ts = map3 mk_sets_of_bnf (map (replicate live) Dss)
       
   114       (map (replicate live) (replicate n Ts)) bnfs;
       
   115     val setssAs = mk_setss allAs;
       
   116     val bd0s = map3 mk_bd_of_bnf Dss Ass bnfs;
       
   117     val bds =
       
   118       map3 (fn bd0 => fn Ds => fn bnf => mk_csum bd0
       
   119         (mk_card_of (HOLogic.mk_UNIV
       
   120           (mk_T_of_bnf Ds (replicate live (fst (dest_relT (fastype_of bd0)))) bnf))))
       
   121       bd0s Dss bnfs;
       
   122     val witss = map wits_of_bnf bnfs;
       
   123 
       
   124     val (((((((((((((((((((zs, zs'), As), Bs), Bs_copy), B's), B''s), ss), prod_ss), s's), s''s),
       
   125       fs), fs_copy), inv_fs), self_fs), gs), all_gs), (xFs, xFs')), (yFs, yFs')),
       
   126       names_lthy) = lthy
       
   127       |> mk_Frees' "z" activeAs
       
   128       ||>> mk_Frees "A" ATs
       
   129       ||>> mk_Frees "B" BTs
       
   130       ||>> mk_Frees "B" BTs
       
   131       ||>> mk_Frees "B'" B'Ts
       
   132       ||>> mk_Frees "B''" B''Ts
       
   133       ||>> mk_Frees "s" sTs
       
   134       ||>> mk_Frees "prods" prod_sTs
       
   135       ||>> mk_Frees "s'" s'Ts
       
   136       ||>> mk_Frees "s''" s''Ts
       
   137       ||>> mk_Frees "f" fTs
       
   138       ||>> mk_Frees "f" fTs
       
   139       ||>> mk_Frees "f" inv_fTs
       
   140       ||>> mk_Frees "f" self_fTs
       
   141       ||>> mk_Frees "g" gTs
       
   142       ||>> mk_Frees "g" all_gTs
       
   143       ||>> mk_Frees' "x" FTsAs
       
   144       ||>> mk_Frees' "y" FTsBs;
       
   145 
       
   146     val passive_UNIVs = map HOLogic.mk_UNIV passiveAs;
       
   147     val active_UNIVs = map HOLogic.mk_UNIV activeAs;
       
   148     val prod_UNIVs = map HOLogic.mk_UNIV prodBsAs;
       
   149     val passive_ids = map HOLogic.id_const passiveAs;
       
   150     val active_ids = map HOLogic.id_const activeAs;
       
   151     val fsts = map fst_const prodBsAs;
       
   152 
       
   153     (* thms *)
       
   154     val bd0_card_orders = map bd_card_order_of_bnf bnfs;
       
   155     val bd0_Card_orders = map bd_Card_order_of_bnf bnfs;
       
   156     val bd0_Cinfinites = map bd_Cinfinite_of_bnf bnfs;
       
   157     val set_bd0ss = map set_bd_of_bnf bnfs;
       
   158 
       
   159     val bd_card_orders =
       
   160       map (fn thm => @{thm card_order_csum} OF [thm, @{thm card_of_card_order_on}]) bd0_card_orders;
       
   161     val bd_Card_order = @{thm Card_order_csum};
       
   162     val bd_Card_orders = replicate n bd_Card_order;
       
   163     val bd_Cinfinites = map (fn thm => thm RS @{thm Cinfinite_csum1}) bd0_Cinfinites;
       
   164     val bd_Cnotzeros = map (fn thm => thm RS @{thm Cinfinite_Cnotzero}) bd_Cinfinites;
       
   165     val bd_Cinfinite = hd bd_Cinfinites;
       
   166     val bd_Cnotzero = hd bd_Cnotzeros;
       
   167     val set_bdss =
       
   168       map2 (fn set_bd0s => fn bd0_Card_order =>
       
   169         map (fn thm => ctrans OF [thm, bd0_Card_order RS @{thm ordLeq_csum1}]) set_bd0s)
       
   170       set_bd0ss bd0_Card_orders;
       
   171     val in_bds = map in_bd_of_bnf bnfs;
       
   172     val sym_map_comps = map (fn bnf => map_comp0_of_bnf bnf RS sym) bnfs;
       
   173     val map_comps = map map_comp_of_bnf bnfs;
       
   174     val map_cong0s = map map_cong0_of_bnf bnfs;
       
   175     val map_id0s = map map_id0_of_bnf bnfs;
       
   176     val map_ids = map map_id_of_bnf bnfs;
       
   177     val set_mapss = map set_map_of_bnf bnfs;
       
   178     val rel_mono_strongs = map rel_mono_strong_of_bnf bnfs;
       
   179     val rel_OOs = map rel_OO_of_bnf bnfs;
       
   180 
       
   181     val timer = time (timer "Extracted terms & thms");
       
   182 
       
   183     (* nonemptiness check *)
       
   184     fun new_wit X (wit: nonemptiness_witness) = subset (op =) (#I wit, (0 upto m - 1) @ map snd X);
       
   185 
       
   186     val all = m upto m + n - 1;
       
   187 
       
   188     fun enrich X = map_filter (fn i =>
       
   189       (case find_first (fn (_, i') => i = i') X of
       
   190         NONE =>
       
   191           (case find_index (new_wit X) (nth witss (i - m)) of
       
   192             ~1 => NONE
       
   193           | j => SOME (j, i))
       
   194       | SOME ji => SOME ji)) all;
       
   195     val reachable = fixpoint (op =) enrich [];
       
   196     val _ = (case subtract (op =) (map snd reachable) all of
       
   197         [] => ()
       
   198       | i :: _ => error ("Cannot define empty datatype " ^ quote (Binding.name_of (nth bs (i - m)))));
       
   199 
       
   200     val wit_thms = flat (map2 (fn bnf => fn (j, _) => nth (wit_thmss_of_bnf bnf) j) bnfs reachable);
       
   201 
       
   202     val timer = time (timer "Checked nonemptiness");
       
   203 
       
   204     (* derived thms *)
       
   205 
       
   206     (*map g1 ... gm g(m+1) ... g(m+n) (map id ... id f(m+1) ... f(m+n) x) =
       
   207       map g1 ... gm (g(m+1) o f(m+1)) ... (g(m+n) o f(m+n)) x*)
       
   208     fun mk_map_comp_id x mapAsBs mapBsCs mapAsCs map_comp0 =
       
   209       let
       
   210         val lhs = Term.list_comb (mapBsCs, all_gs) $
       
   211           (Term.list_comb (mapAsBs, passive_ids @ fs) $ x);
       
   212         val rhs = Term.list_comb (mapAsCs,
       
   213           take m all_gs @ map HOLogic.mk_comp (drop m all_gs ~~ fs)) $ x;
       
   214       in
       
   215         Goal.prove_sorry lthy [] []
       
   216           (fold_rev Logic.all (x :: fs @ all_gs) (mk_Trueprop_eq (lhs, rhs)))
       
   217           (K (mk_map_comp_id_tac map_comp0))
       
   218         |> Thm.close_derivation
       
   219       end;
       
   220 
       
   221     val map_comp_id_thms = map5 mk_map_comp_id xFs mapsAsBs mapsBsCs' mapsAsCs' map_comps;
       
   222 
       
   223     (*forall a : set(m+1) x. f(m+1) a = a; ...; forall a : set(m+n) x. f(m+n) a = a ==>
       
   224       map id ... id f(m+1) ... f(m+n) x = x*)
       
   225     fun mk_map_cong0L x mapAsAs sets map_cong0 map_id =
       
   226       let
       
   227         fun mk_prem set f z z' = HOLogic.mk_Trueprop
       
   228           (mk_Ball (set $ x) (Term.absfree z' (HOLogic.mk_eq (f $ z, z))));
       
   229         val prems = map4 mk_prem (drop m sets) self_fs zs zs';
       
   230         val goal = mk_Trueprop_eq (Term.list_comb (mapAsAs, passive_ids @ self_fs) $ x, x);
       
   231       in
       
   232         Goal.prove_sorry lthy [] []
       
   233           (fold_rev Logic.all (x :: self_fs) (Logic.list_implies (prems, goal)))
       
   234           (K (mk_map_cong0L_tac m map_cong0 map_id))
       
   235         |> Thm.close_derivation
       
   236       end;
       
   237 
       
   238     val map_cong0L_thms = map5 mk_map_cong0L xFs mapsAsAs setssAs map_cong0s map_ids;
       
   239     val in_mono'_thms = map (fn bnf => in_mono_of_bnf bnf OF (replicate m subset_refl)) bnfs;
       
   240     val in_cong'_thms = map (fn bnf => in_cong_of_bnf bnf OF (replicate m refl)) bnfs;
       
   241 
       
   242     val timer = time (timer "Derived simple theorems");
       
   243 
       
   244     (* algebra *)
       
   245 
       
   246     val alg_bind = mk_internal_b algN;
       
   247     val alg_name = Binding.name_of alg_bind;
       
   248     val alg_def_bind = (Thm.def_binding alg_bind, []);
       
   249 
       
   250     (*forall i = 1 ... n: (\<forall>x \<in> Fi_in A1 .. Am B1 ... Bn. si x \<in> Bi)*)
       
   251     val alg_spec =
       
   252       let
       
   253         val algT = Library.foldr (op -->) (ATs @ BTs @ sTs, HOLogic.boolT);
       
   254 
       
   255         val ins = map3 mk_in (replicate n (As @ Bs)) setssAs FTsAs;
       
   256         fun mk_alg_conjunct B s X x x' =
       
   257           mk_Ball X (Term.absfree x' (HOLogic.mk_mem (s $ x, B)));
       
   258 
       
   259         val lhs = Term.list_comb (Free (alg_name, algT), As @ Bs @ ss);
       
   260         val rhs = Library.foldr1 HOLogic.mk_conj (map5 mk_alg_conjunct Bs ss ins xFs xFs')
       
   261       in
       
   262         mk_Trueprop_eq (lhs, rhs)
       
   263       end;
       
   264 
       
   265     val ((alg_free, (_, alg_def_free)), (lthy, lthy_old)) =
       
   266         lthy
       
   267         |> Specification.definition (SOME (alg_bind, NONE, NoSyn), (alg_def_bind, alg_spec))
       
   268         ||> `Local_Theory.restore;
       
   269 
       
   270     val phi = Proof_Context.export_morphism lthy_old lthy;
       
   271     val alg = fst (Term.dest_Const (Morphism.term phi alg_free));
       
   272     val alg_def = Morphism.thm phi alg_def_free;
       
   273 
       
   274     fun mk_alg As Bs ss =
       
   275       let
       
   276         val args = As @ Bs @ ss;
       
   277         val Ts = map fastype_of args;
       
   278         val algT = Library.foldr (op -->) (Ts, HOLogic.boolT);
       
   279       in
       
   280         Term.list_comb (Const (alg, algT), args)
       
   281       end;
       
   282 
       
   283     val alg_set_thms =
       
   284       let
       
   285         val alg_prem = HOLogic.mk_Trueprop (mk_alg As Bs ss);
       
   286         fun mk_prem x set B = HOLogic.mk_Trueprop (mk_leq (set $ x) B);
       
   287         fun mk_concl s x B = HOLogic.mk_Trueprop (HOLogic.mk_mem (s $ x, B));
       
   288         val premss = map2 ((fn x => fn sets =>  map2 (mk_prem x) sets (As @ Bs))) xFs setssAs;
       
   289         val concls = map3 mk_concl ss xFs Bs;
       
   290         val goals = map3 (fn x => fn prems => fn concl =>
       
   291           fold_rev Logic.all (x :: As @ Bs @ ss)
       
   292             (Logic.list_implies (alg_prem :: prems, concl))) xFs premss concls;
       
   293       in
       
   294         map (fn goal =>
       
   295           Goal.prove_sorry lthy [] [] goal (K (mk_alg_set_tac alg_def)) |> Thm.close_derivation)
       
   296         goals
       
   297       end;
       
   298 
       
   299     fun mk_talg ATs BTs = mk_alg (map HOLogic.mk_UNIV ATs) (map HOLogic.mk_UNIV BTs);
       
   300 
       
   301     val talg_thm =
       
   302       let
       
   303         val goal = fold_rev Logic.all ss
       
   304           (HOLogic.mk_Trueprop (mk_talg passiveAs activeAs ss))
       
   305       in
       
   306         Goal.prove_sorry lthy [] [] goal
       
   307           (K (stac alg_def 1 THEN CONJ_WRAP (K (EVERY' [rtac ballI, rtac UNIV_I] 1)) ss))
       
   308         |> Thm.close_derivation
       
   309       end;
       
   310 
       
   311     val timer = time (timer "Algebra definition & thms");
       
   312 
       
   313     val alg_not_empty_thms =
       
   314       let
       
   315         val alg_prem =
       
   316           HOLogic.mk_Trueprop (mk_alg passive_UNIVs Bs ss);
       
   317         val concls = map (HOLogic.mk_Trueprop o mk_not_empty) Bs;
       
   318         val goals =
       
   319           map (fn concl =>
       
   320             fold_rev Logic.all (Bs @ ss) (Logic.mk_implies (alg_prem, concl))) concls;
       
   321       in
       
   322         map2 (fn goal => fn alg_set =>
       
   323           Goal.prove_sorry lthy [] []
       
   324             goal (K (mk_alg_not_empty_tac lthy alg_set alg_set_thms wit_thms))
       
   325           |> Thm.close_derivation)
       
   326         goals alg_set_thms
       
   327       end;
       
   328 
       
   329     val timer = time (timer "Proved nonemptiness");
       
   330 
       
   331     (* morphism *)
       
   332 
       
   333     val mor_bind = mk_internal_b morN;
       
   334     val mor_name = Binding.name_of mor_bind;
       
   335     val mor_def_bind = (Thm.def_binding mor_bind, []);
       
   336 
       
   337     (*fbetw) forall i = 1 ... n: (\<forall>x \<in> Bi. f x \<in> B'i)*)
       
   338     (*mor) forall i = 1 ... n: (\<forall>x \<in> Fi_in UNIV ... UNIV B1 ... Bn.
       
   339        f (s1 x) = s1' (Fi_map id ... id f1 ... fn x))*)
       
   340     val mor_spec =
       
   341       let
       
   342         val morT = Library.foldr (op -->) (BTs @ sTs @ B'Ts @ s'Ts @ fTs, HOLogic.boolT);
       
   343 
       
   344         fun mk_fbetw f B1 B2 z z' =
       
   345           mk_Ball B1 (Term.absfree z' (HOLogic.mk_mem (f $ z, B2)));
       
   346         fun mk_mor sets mapAsBs f s s' T x x' =
       
   347           mk_Ball (mk_in (passive_UNIVs @ Bs) sets T)
       
   348             (Term.absfree x' (HOLogic.mk_eq (f $ (s $ x), s' $
       
   349               (Term.list_comb (mapAsBs, passive_ids @ fs) $ x))));
       
   350         val lhs = Term.list_comb (Free (mor_name, morT), Bs @ ss @ B's @ s's @ fs);
       
   351         val rhs = HOLogic.mk_conj
       
   352           (Library.foldr1 HOLogic.mk_conj (map5 mk_fbetw fs Bs B's zs zs'),
       
   353           Library.foldr1 HOLogic.mk_conj
       
   354             (map8 mk_mor setssAs mapsAsBs fs ss s's FTsAs xFs xFs'))
       
   355       in
       
   356         mk_Trueprop_eq (lhs, rhs)
       
   357       end;
       
   358 
       
   359     val ((mor_free, (_, mor_def_free)), (lthy, lthy_old)) =
       
   360         lthy
       
   361         |> Specification.definition (SOME (mor_bind, NONE, NoSyn), (mor_def_bind, mor_spec))
       
   362         ||> `Local_Theory.restore;
       
   363 
       
   364     val phi = Proof_Context.export_morphism lthy_old lthy;
       
   365     val mor = fst (Term.dest_Const (Morphism.term phi mor_free));
       
   366     val mor_def = Morphism.thm phi mor_def_free;
       
   367 
       
   368     fun mk_mor Bs1 ss1 Bs2 ss2 fs =
       
   369       let
       
   370         val args = Bs1 @ ss1 @ Bs2 @ ss2 @ fs;
       
   371         val Ts = map fastype_of (Bs1 @ ss1 @ Bs2 @ ss2 @ fs);
       
   372         val morT = Library.foldr (op -->) (Ts, HOLogic.boolT);
       
   373       in
       
   374         Term.list_comb (Const (mor, morT), args)
       
   375       end;
       
   376 
       
   377     val (mor_image_thms, morE_thms) =
       
   378       let
       
   379         val prem = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs);
       
   380         fun mk_image_goal f B1 B2 = fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs)
       
   381           (Logic.mk_implies (prem, HOLogic.mk_Trueprop (mk_leq (mk_image f $ B1) B2)));
       
   382         val image_goals = map3 mk_image_goal fs Bs B's;
       
   383         fun mk_elim_prem sets x T = HOLogic.mk_Trueprop
       
   384           (HOLogic.mk_mem (x, mk_in (passive_UNIVs @ Bs) sets T));
       
   385         fun mk_elim_goal sets mapAsBs f s s' x T =
       
   386           fold_rev Logic.all (x :: Bs @ ss @ B's @ s's @ fs)
       
   387             (Logic.list_implies ([prem, mk_elim_prem sets x T],
       
   388               mk_Trueprop_eq (f $ (s $ x), s' $ Term.list_comb (mapAsBs, passive_ids @ fs @ [x]))));
       
   389         val elim_goals = map7 mk_elim_goal setssAs mapsAsBs fs ss s's xFs FTsAs;
       
   390         fun prove goal =
       
   391           Goal.prove_sorry lthy [] [] goal (K (mk_mor_elim_tac mor_def)) |> Thm.close_derivation;
       
   392       in
       
   393         (map prove image_goals, map prove elim_goals)
       
   394       end;
       
   395 
       
   396     val mor_incl_thm =
       
   397       let
       
   398         val prems = map2 (HOLogic.mk_Trueprop oo mk_leq) Bs Bs_copy;
       
   399         val concl = HOLogic.mk_Trueprop (mk_mor Bs ss Bs_copy ss active_ids);
       
   400       in
       
   401         Goal.prove_sorry lthy [] []
       
   402           (fold_rev Logic.all (Bs @ ss @ Bs_copy) (Logic.list_implies (prems, concl)))
       
   403           (K (mk_mor_incl_tac mor_def map_ids))
       
   404         |> Thm.close_derivation
       
   405       end;
       
   406 
       
   407     val mor_comp_thm =
       
   408       let
       
   409         val prems =
       
   410           [HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs),
       
   411            HOLogic.mk_Trueprop (mk_mor B's s's B''s s''s gs)];
       
   412         val concl =
       
   413           HOLogic.mk_Trueprop (mk_mor Bs ss B''s s''s (map2 (curry HOLogic.mk_comp) gs fs));
       
   414       in
       
   415         Goal.prove_sorry lthy [] []
       
   416           (fold_rev Logic.all (Bs @ ss @ B's @ s's @ B''s @ s''s @ fs @ gs)
       
   417              (Logic.list_implies (prems, concl)))
       
   418           (K (mk_mor_comp_tac mor_def set_mapss map_comp_id_thms))
       
   419         |> Thm.close_derivation
       
   420       end;
       
   421 
       
   422     val mor_inv_thm =
       
   423       let
       
   424         fun mk_inv_prem f inv_f B B' = HOLogic.mk_conj (mk_leq (mk_image inv_f $ B') B,
       
   425           HOLogic.mk_conj (mk_inver inv_f f B, mk_inver f inv_f B'));
       
   426         val prems = map HOLogic.mk_Trueprop
       
   427           ([mk_mor Bs ss B's s's fs,
       
   428           mk_alg passive_UNIVs Bs ss,
       
   429           mk_alg passive_UNIVs B's s's] @
       
   430           map4 mk_inv_prem fs inv_fs Bs B's);
       
   431         val concl = HOLogic.mk_Trueprop (mk_mor B's s's Bs ss inv_fs);
       
   432       in
       
   433         Goal.prove_sorry lthy [] []
       
   434           (fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs @ inv_fs)
       
   435             (Logic.list_implies (prems, concl)))
       
   436           (K (mk_mor_inv_tac alg_def mor_def set_mapss morE_thms map_comp_id_thms map_cong0L_thms))
       
   437         |> Thm.close_derivation
       
   438       end;
       
   439 
       
   440     val mor_cong_thm =
       
   441       let
       
   442         val prems = map HOLogic.mk_Trueprop
       
   443          (map2 (curry HOLogic.mk_eq) fs_copy fs @ [mk_mor Bs ss B's s's fs])
       
   444         val concl = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs_copy);
       
   445       in
       
   446         Goal.prove_sorry lthy [] []
       
   447           (fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs @ fs_copy)
       
   448              (Logic.list_implies (prems, concl)))
       
   449           (K ((hyp_subst_tac lthy THEN' atac) 1))
       
   450         |> Thm.close_derivation
       
   451       end;
       
   452 
       
   453     val mor_str_thm =
       
   454       let
       
   455         val maps = map2 (fn Ds => fn bnf => Term.list_comb
       
   456           (mk_map_of_bnf Ds (passiveAs @ FTsAs) allAs bnf, passive_ids @ ss)) Dss bnfs;
       
   457       in
       
   458         Goal.prove_sorry lthy [] []
       
   459           (fold_rev Logic.all ss (HOLogic.mk_Trueprop
       
   460             (mk_mor (map HOLogic.mk_UNIV FTsAs) maps active_UNIVs ss ss)))
       
   461           (K (mk_mor_str_tac ks mor_def))
       
   462         |> Thm.close_derivation
       
   463       end;
       
   464 
       
   465     val mor_convol_thm =
       
   466       let
       
   467         val maps = map3 (fn s => fn prod_s => fn mapx =>
       
   468           mk_convol (HOLogic.mk_comp (s, Term.list_comb (mapx, passive_ids @ fsts)), prod_s))
       
   469           s's prod_ss map_fsts;
       
   470       in
       
   471         Goal.prove_sorry lthy [] []
       
   472           (fold_rev Logic.all (s's @ prod_ss) (HOLogic.mk_Trueprop
       
   473             (mk_mor prod_UNIVs maps (map HOLogic.mk_UNIV activeBs) s's fsts)))
       
   474           (K (mk_mor_convol_tac ks mor_def))
       
   475         |> Thm.close_derivation
       
   476       end;
       
   477 
       
   478     val mor_UNIV_thm =
       
   479       let
       
   480         fun mk_conjunct mapAsBs f s s' = HOLogic.mk_eq
       
   481             (HOLogic.mk_comp (f, s),
       
   482             HOLogic.mk_comp (s', Term.list_comb (mapAsBs, passive_ids @ fs)));
       
   483         val lhs = mk_mor active_UNIVs ss (map HOLogic.mk_UNIV activeBs) s's fs;
       
   484         val rhs = Library.foldr1 HOLogic.mk_conj (map4 mk_conjunct mapsAsBs fs ss s's);
       
   485       in
       
   486         Goal.prove_sorry lthy [] [] (fold_rev Logic.all (ss @ s's @ fs) (mk_Trueprop_eq (lhs, rhs)))
       
   487           (K (mk_mor_UNIV_tac m morE_thms mor_def))
       
   488         |> Thm.close_derivation
       
   489       end;
       
   490 
       
   491     val timer = time (timer "Morphism definition & thms");
       
   492 
       
   493     (* isomorphism *)
       
   494 
       
   495     (*mor Bs1 ss1 Bs2 ss2 fs \<and> (\<exists>gs. mor Bs2 ss2 Bs1 ss1 fs \<and>
       
   496        forall i = 1 ... n. (inver gs[i] fs[i] Bs1[i] \<and> inver fs[i] gs[i] Bs2[i]))*)
       
   497     fun mk_iso Bs1 ss1 Bs2 ss2 fs gs =
       
   498       let
       
   499         val ex_inv_mor = list_exists_free gs
       
   500           (HOLogic.mk_conj (mk_mor Bs2 ss2 Bs1 ss1 gs,
       
   501             Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_conj)
       
   502               (map3 mk_inver gs fs Bs1) (map3 mk_inver fs gs Bs2))));
       
   503       in
       
   504         HOLogic.mk_conj (mk_mor Bs1 ss1 Bs2 ss2 fs, ex_inv_mor)
       
   505       end;
       
   506 
       
   507     val iso_alt_thm =
       
   508       let
       
   509         val prems = map HOLogic.mk_Trueprop
       
   510          [mk_alg passive_UNIVs Bs ss,
       
   511          mk_alg passive_UNIVs B's s's]
       
   512         val concl = mk_Trueprop_eq (mk_iso Bs ss B's s's fs inv_fs,
       
   513           HOLogic.mk_conj (mk_mor Bs ss B's s's fs,
       
   514             Library.foldr1 HOLogic.mk_conj (map3 mk_bij_betw fs Bs B's)));
       
   515       in
       
   516         Goal.prove_sorry lthy [] []
       
   517           (fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs) (Logic.list_implies (prems, concl)))
       
   518           (K (mk_iso_alt_tac mor_image_thms mor_inv_thm))
       
   519         |> Thm.close_derivation
       
   520       end;
       
   521 
       
   522     val timer = time (timer "Isomorphism definition & thms");
       
   523 
       
   524     (* algebra copies *)
       
   525 
       
   526     val (copy_alg_thm, ex_copy_alg_thm) =
       
   527       let
       
   528         val prems = map HOLogic.mk_Trueprop
       
   529          (mk_alg passive_UNIVs Bs ss :: map3 mk_bij_betw inv_fs B's Bs);
       
   530         val inver_prems = map HOLogic.mk_Trueprop
       
   531           (map3 mk_inver inv_fs fs Bs @ map3 mk_inver fs inv_fs B's);
       
   532         val all_prems = prems @ inver_prems;
       
   533         fun mk_s f s mapT y y' = Term.absfree y' (f $ (s $
       
   534           (Term.list_comb (mapT, passive_ids @ inv_fs) $ y)));
       
   535 
       
   536         val alg = HOLogic.mk_Trueprop
       
   537           (mk_alg passive_UNIVs B's (map5 mk_s fs ss mapsBsAs yFs yFs'));
       
   538         val copy_str_thm = Goal.prove_sorry lthy [] []
       
   539           (fold_rev Logic.all (Bs @ ss @ B's @ inv_fs @ fs)
       
   540             (Logic.list_implies (all_prems, alg)))
       
   541           (K (mk_copy_str_tac set_mapss alg_def alg_set_thms))
       
   542           |> Thm.close_derivation;
       
   543 
       
   544         val iso = HOLogic.mk_Trueprop
       
   545           (mk_iso B's (map5 mk_s fs ss mapsBsAs yFs yFs') Bs ss inv_fs fs_copy);
       
   546         val copy_alg_thm = Goal.prove_sorry lthy [] []
       
   547           (fold_rev Logic.all (Bs @ ss @ B's @ inv_fs @ fs)
       
   548             (Logic.list_implies (all_prems, iso)))
       
   549           (K (mk_copy_alg_tac set_mapss alg_set_thms mor_def iso_alt_thm copy_str_thm))
       
   550           |> Thm.close_derivation;
       
   551 
       
   552         val ex = HOLogic.mk_Trueprop
       
   553           (list_exists_free s's
       
   554             (HOLogic.mk_conj (mk_alg passive_UNIVs B's s's,
       
   555               mk_iso B's s's Bs ss inv_fs fs_copy)));
       
   556         val ex_copy_alg_thm = Goal.prove_sorry lthy [] []
       
   557           (fold_rev Logic.all (Bs @ ss @ B's @ inv_fs @ fs)
       
   558              (Logic.list_implies (prems, ex)))
       
   559           (K (mk_ex_copy_alg_tac n copy_str_thm copy_alg_thm))
       
   560           |> Thm.close_derivation;
       
   561       in
       
   562         (copy_alg_thm, ex_copy_alg_thm)
       
   563       end;
       
   564 
       
   565     val timer = time (timer "Copy thms");
       
   566 
       
   567 
       
   568     (* bounds *)
       
   569 
       
   570     val sum_Card_order = if n = 1 then bd_Card_order else @{thm Card_order_csum};
       
   571     val sum_Cnotzero = if n = 1 then bd_Cnotzero else bd_Cnotzero RS @{thm csum_Cnotzero1};
       
   572     val sum_Cinfinite = if n = 1 then bd_Cinfinite else bd_Cinfinite RS @{thm Cinfinite_csum1};
       
   573     fun mk_set_bd_sums i bd_Card_order bds =
       
   574       if n = 1 then bds
       
   575       else map (fn thm => bd_Card_order RS mk_ordLeq_csum n i thm) bds;
       
   576     val set_bd_sumss = map3 mk_set_bd_sums ks bd_Card_orders set_bdss;
       
   577 
       
   578     fun mk_in_bd_sum i Co Cnz bd =
       
   579       if n = 1 then bd
       
   580       else Cnz RS ((Co RS mk_ordLeq_csum n i (Co RS @{thm ordLeq_refl})) RS
       
   581         (bd RS @{thm ordLeq_transitive[OF _ cexp_mono2_Cnotzero[OF _ Card_order_csum]]}));
       
   582     val in_bd_sums = map4 mk_in_bd_sum ks bd_Card_orders bd_Cnotzeros in_bds;
       
   583 
       
   584     val sum_bd = Library.foldr1 (uncurry mk_csum) bds;
       
   585     val suc_bd = mk_cardSuc sum_bd;
       
   586     val field_suc_bd = mk_Field suc_bd;
       
   587     val suc_bdT = fst (dest_relT (fastype_of suc_bd));
       
   588     fun mk_Asuc_bd [] = mk_cexp ctwo suc_bd
       
   589       | mk_Asuc_bd As =
       
   590         mk_cexp (mk_csum (Library.foldr1 (uncurry mk_csum) (map mk_card_of As)) ctwo) suc_bd;
       
   591 
       
   592     val suc_bd_Card_order = if n = 1 then bd_Card_order RS @{thm cardSuc_Card_order}
       
   593       else @{thm cardSuc_Card_order[OF Card_order_csum]};
       
   594     val suc_bd_Cinfinite = if n = 1 then bd_Cinfinite RS @{thm Cinfinite_cardSuc}
       
   595       else bd_Cinfinite RS @{thm Cinfinite_cardSuc[OF Cinfinite_csum1]};
       
   596     val suc_bd_Cnotzero = suc_bd_Cinfinite RS @{thm Cinfinite_Cnotzero};
       
   597     val suc_bd_worel = suc_bd_Card_order RS @{thm Card_order_wo_rel}
       
   598     val basis_Asuc = if m = 0 then @{thm ordLeq_refl[OF Card_order_ctwo]}
       
   599         else @{thm ordLeq_csum2[OF Card_order_ctwo]};
       
   600     val Asuc_bd_Cinfinite = suc_bd_Cinfinite RS (basis_Asuc RS @{thm Cinfinite_cexp});
       
   601 
       
   602     val suc_bd_Asuc_bd = @{thm ordLess_ordLeq_trans[OF ordLess_ctwo_cexp cexp_mono1]} OF
       
   603       [suc_bd_Card_order, basis_Asuc, suc_bd_Card_order];
       
   604 
       
   605     val Asuc_bdT = fst (dest_relT (fastype_of (mk_Asuc_bd As)));
       
   606     val II_BTs = replicate n (HOLogic.mk_setT Asuc_bdT);
       
   607     val II_sTs = map2 (fn Ds => fn bnf =>
       
   608       mk_T_of_bnf Ds (passiveAs @ replicate n Asuc_bdT) bnf --> Asuc_bdT) Dss bnfs;
       
   609 
       
   610     val (((((((idxs, Asi_name), (idx, idx')), (jdx, jdx')), II_Bs), II_ss), Asuc_fs),
       
   611       names_lthy) = names_lthy
       
   612       |> mk_Frees "i" (replicate n suc_bdT)
       
   613       ||>> (fn ctxt => apfst the_single (mk_fresh_names ctxt 1 "Asi"))
       
   614       ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "i") suc_bdT
       
   615       ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "j") suc_bdT
       
   616       ||>> mk_Frees "IIB" II_BTs
       
   617       ||>> mk_Frees "IIs" II_sTs
       
   618       ||>> mk_Frees "f" (map (fn T => Asuc_bdT --> T) activeAs);
       
   619 
       
   620     val suc_bd_limit_thm =
       
   621       let
       
   622         val prem = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
       
   623           (map (fn idx => HOLogic.mk_mem (idx, field_suc_bd)) idxs));
       
   624         fun mk_conjunct idx = HOLogic.mk_conj (mk_not_eq idx jdx,
       
   625           HOLogic.mk_mem (HOLogic.mk_prod (idx, jdx), suc_bd));
       
   626         val concl = HOLogic.mk_Trueprop (mk_Bex field_suc_bd
       
   627           (Term.absfree jdx' (Library.foldr1 HOLogic.mk_conj (map mk_conjunct idxs))));
       
   628       in
       
   629         Goal.prove_sorry lthy [] []
       
   630           (fold_rev Logic.all idxs (Logic.list_implies ([prem], concl)))
       
   631           (K (mk_bd_limit_tac n suc_bd_Cinfinite))
       
   632         |> Thm.close_derivation
       
   633       end;
       
   634 
       
   635     val timer = time (timer "Bounds");
       
   636 
       
   637 
       
   638     (* minimal algebra *)
       
   639 
       
   640     fun mk_minG Asi i k = mk_UNION (mk_underS suc_bd $ i)
       
   641       (Term.absfree jdx' (mk_nthN n (Asi $ jdx) k));
       
   642 
       
   643     fun mk_minH_component As Asi i sets Ts s k =
       
   644       HOLogic.mk_binop @{const_name "sup"}
       
   645       (mk_minG Asi i k, mk_image s $ mk_in (As @ map (mk_minG Asi i) ks) sets Ts);
       
   646 
       
   647     fun mk_min_algs As ss =
       
   648       let
       
   649         val BTs = map (range_type o fastype_of) ss;
       
   650         val Ts = map (HOLogic.dest_setT o fastype_of) As @ BTs;
       
   651         val (Asi, Asi') = `Free (Asi_name, suc_bdT -->
       
   652           Library.foldr1 HOLogic.mk_prodT (map HOLogic.mk_setT BTs));
       
   653       in
       
   654          mk_worec suc_bd (Term.absfree Asi' (Term.absfree idx' (HOLogic.mk_tuple
       
   655            (map4 (mk_minH_component As Asi idx) (mk_setss Ts) (mk_FTs Ts) ss ks))))
       
   656       end;
       
   657 
       
   658     val (min_algs_thms, min_algs_mono_thms, card_of_min_algs_thm, least_min_algs_thm) =
       
   659       let
       
   660         val i_field = HOLogic.mk_mem (idx, field_suc_bd);
       
   661         val min_algs = mk_min_algs As ss;
       
   662         val min_algss = map (fn k => mk_nthN n (min_algs $ idx) k) ks;
       
   663 
       
   664         val concl = HOLogic.mk_Trueprop
       
   665           (HOLogic.mk_eq (min_algs $ idx, HOLogic.mk_tuple
       
   666             (map4 (mk_minH_component As min_algs idx) setssAs FTsAs ss ks)));
       
   667         val goal = fold_rev Logic.all (idx :: As @ ss)
       
   668           (Logic.mk_implies (HOLogic.mk_Trueprop i_field, concl));
       
   669 
       
   670         val min_algs_thm = Goal.prove_sorry lthy [] [] goal
       
   671           (K (mk_min_algs_tac suc_bd_worel in_cong'_thms))
       
   672           |> Thm.close_derivation;
       
   673 
       
   674         val min_algs_thms = map (fn k => min_algs_thm RS mk_nthI n k) ks;
       
   675 
       
   676         fun mk_mono_goal min_alg =
       
   677           fold_rev Logic.all (As @ ss) (HOLogic.mk_Trueprop (mk_relChain suc_bd
       
   678             (Term.absfree idx' min_alg)));
       
   679 
       
   680         val monos =
       
   681           map2 (fn goal => fn min_algs =>
       
   682             Goal.prove_sorry lthy [] [] goal (K (mk_min_algs_mono_tac lthy min_algs))
       
   683             |> Thm.close_derivation)
       
   684           (map mk_mono_goal min_algss) min_algs_thms;
       
   685 
       
   686         val Asuc_bd = mk_Asuc_bd As;
       
   687 
       
   688         fun mk_card_conjunct min_alg = mk_ordLeq (mk_card_of min_alg) Asuc_bd;
       
   689         val card_conjunction = Library.foldr1 HOLogic.mk_conj (map mk_card_conjunct min_algss);
       
   690         val card_cT = certifyT lthy suc_bdT;
       
   691         val card_ct = certify lthy (Term.absfree idx' card_conjunction);
       
   692 
       
   693         val card_of = singleton (Proof_Context.export names_lthy lthy)
       
   694           (Goal.prove_sorry lthy [] []
       
   695             (HOLogic.mk_Trueprop (HOLogic.mk_imp (i_field, card_conjunction)))
       
   696             (K (mk_min_algs_card_of_tac card_cT card_ct
       
   697               m suc_bd_worel min_algs_thms in_bd_sums
       
   698               sum_Card_order sum_Cnotzero suc_bd_Card_order suc_bd_Cinfinite suc_bd_Cnotzero
       
   699               suc_bd_Asuc_bd Asuc_bd_Cinfinite)))
       
   700           |> Thm.close_derivation;
       
   701 
       
   702         val least_prem = HOLogic.mk_Trueprop (mk_alg As Bs ss);
       
   703         val least_conjunction = Library.foldr1 HOLogic.mk_conj (map2 mk_leq min_algss Bs);
       
   704         val least_cT = certifyT lthy suc_bdT;
       
   705         val least_ct = certify lthy (Term.absfree idx' least_conjunction);
       
   706 
       
   707         val least = singleton (Proof_Context.export names_lthy lthy)
       
   708           (Goal.prove_sorry lthy [] []
       
   709             (Logic.mk_implies (least_prem,
       
   710               HOLogic.mk_Trueprop (HOLogic.mk_imp (i_field, least_conjunction))))
       
   711             (K (mk_min_algs_least_tac least_cT least_ct
       
   712               suc_bd_worel min_algs_thms alg_set_thms)))
       
   713           |> Thm.close_derivation;
       
   714       in
       
   715         (min_algs_thms, monos, card_of, least)
       
   716       end;
       
   717 
       
   718     val timer = time (timer "min_algs definition & thms");
       
   719 
       
   720     val min_alg_binds = mk_internal_bs min_algN;
       
   721     fun min_alg_bind i = nth min_alg_binds (i - 1);
       
   722     fun min_alg_name i = Binding.name_of (min_alg_bind i);
       
   723     val min_alg_def_bind = rpair [] o Thm.def_binding o min_alg_bind;
       
   724 
       
   725     fun min_alg_spec i =
       
   726       let
       
   727         val min_algT =
       
   728           Library.foldr (op -->) (ATs @ sTs, HOLogic.mk_setT (nth activeAs (i - 1)));
       
   729 
       
   730         val lhs = Term.list_comb (Free (min_alg_name i, min_algT), As @ ss);
       
   731         val rhs = mk_UNION (field_suc_bd)
       
   732           (Term.absfree idx' (mk_nthN n (mk_min_algs As ss $ idx) i));
       
   733       in
       
   734         mk_Trueprop_eq (lhs, rhs)
       
   735       end;
       
   736 
       
   737     val ((min_alg_frees, (_, min_alg_def_frees)), (lthy, lthy_old)) =
       
   738         lthy
       
   739         |> fold_map (fn i => Specification.definition
       
   740           (SOME (min_alg_bind i, NONE, NoSyn), (min_alg_def_bind i, min_alg_spec i))) ks
       
   741         |>> apsnd split_list o split_list
       
   742         ||> `Local_Theory.restore;
       
   743 
       
   744     val phi = Proof_Context.export_morphism lthy_old lthy;
       
   745     val min_algs = map (fst o Term.dest_Const o Morphism.term phi) min_alg_frees;
       
   746     val min_alg_defs = map (Morphism.thm phi) min_alg_def_frees;
       
   747 
       
   748     fun mk_min_alg As ss i =
       
   749       let
       
   750         val T = HOLogic.mk_setT (range_type (fastype_of (nth ss (i - 1))))
       
   751         val args = As @ ss;
       
   752         val Ts = map fastype_of args;
       
   753         val min_algT = Library.foldr (op -->) (Ts, T);
       
   754       in
       
   755         Term.list_comb (Const (nth min_algs (i - 1), min_algT), args)
       
   756       end;
       
   757 
       
   758     val (alg_min_alg_thm, card_of_min_alg_thms, least_min_alg_thms, mor_incl_min_alg_thm) =
       
   759       let
       
   760         val min_algs = map (mk_min_alg As ss) ks;
       
   761 
       
   762         val goal = fold_rev Logic.all (As @ ss) (HOLogic.mk_Trueprop (mk_alg As min_algs ss));
       
   763         val alg_min_alg = Goal.prove_sorry lthy [] [] goal
       
   764           (K (mk_alg_min_alg_tac m alg_def min_alg_defs suc_bd_limit_thm sum_Cinfinite
       
   765             set_bd_sumss min_algs_thms min_algs_mono_thms))
       
   766           |> Thm.close_derivation;
       
   767 
       
   768         val Asuc_bd = mk_Asuc_bd As;
       
   769         fun mk_card_of_thm min_alg def = Goal.prove_sorry lthy [] []
       
   770           (fold_rev Logic.all (As @ ss)
       
   771             (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of min_alg) Asuc_bd)))
       
   772           (K (mk_card_of_min_alg_tac def card_of_min_algs_thm
       
   773             suc_bd_Card_order suc_bd_Asuc_bd Asuc_bd_Cinfinite))
       
   774           |> Thm.close_derivation;
       
   775 
       
   776         val least_prem = HOLogic.mk_Trueprop (mk_alg As Bs ss);
       
   777         fun mk_least_thm min_alg B def = Goal.prove_sorry lthy [] []
       
   778           (fold_rev Logic.all (As @ Bs @ ss)
       
   779             (Logic.mk_implies (least_prem, HOLogic.mk_Trueprop (mk_leq min_alg B))))
       
   780           (K (mk_least_min_alg_tac def least_min_algs_thm))
       
   781           |> Thm.close_derivation;
       
   782 
       
   783         val leasts = map3 mk_least_thm min_algs Bs min_alg_defs;
       
   784 
       
   785         val incl_prem = HOLogic.mk_Trueprop (mk_alg passive_UNIVs Bs ss);
       
   786         val incl_min_algs = map (mk_min_alg passive_UNIVs ss) ks;
       
   787         val incl = Goal.prove_sorry lthy [] []
       
   788           (fold_rev Logic.all (Bs @ ss)
       
   789             (Logic.mk_implies (incl_prem,
       
   790               HOLogic.mk_Trueprop (mk_mor incl_min_algs ss Bs ss active_ids))))
       
   791           (K (EVERY' (rtac mor_incl_thm :: map etac leasts) 1))
       
   792           |> Thm.close_derivation;
       
   793       in
       
   794         (alg_min_alg, map2 mk_card_of_thm min_algs min_alg_defs, leasts, incl)
       
   795       end;
       
   796 
       
   797     val timer = time (timer "Minimal algebra definition & thms");
       
   798 
       
   799     val II_repT = HOLogic.mk_prodT (HOLogic.mk_tupleT II_BTs, HOLogic.mk_tupleT II_sTs);
       
   800     val IIT_bind = mk_internal_b IITN;
       
   801 
       
   802     val ((IIT_name, (IIT_glob_info, IIT_loc_info)), lthy) =
       
   803       typedef (IIT_bind, params, NoSyn)
       
   804         (HOLogic.mk_UNIV II_repT) NONE (EVERY' [rtac exI, rtac UNIV_I] 1) lthy;
       
   805 
       
   806     val IIT = Type (IIT_name, params');
       
   807     val Abs_IIT = Const (#Abs_name IIT_glob_info, II_repT --> IIT);
       
   808     val Rep_IIT = Const (#Rep_name IIT_glob_info, IIT --> II_repT);
       
   809     val Abs_IIT_inverse_thm = UNIV_I RS #Abs_inverse IIT_loc_info;
       
   810 
       
   811     val initT = IIT --> Asuc_bdT;
       
   812     val active_initTs = replicate n initT;
       
   813     val init_FTs = map2 (fn Ds => mk_T_of_bnf Ds (passiveAs @ active_initTs)) Dss bnfs;
       
   814     val init_fTs = map (fn T => initT --> T) activeAs;
       
   815 
       
   816     val (((((((iidx, iidx'), init_xs), (init_xFs, init_xFs')),
       
   817       init_fs), init_fs_copy), init_phis), names_lthy) = names_lthy
       
   818       |> yield_singleton (apfst (op ~~) oo mk_Frees' "i") IIT
       
   819       ||>> mk_Frees "ix" active_initTs
       
   820       ||>> mk_Frees' "x" init_FTs
       
   821       ||>> mk_Frees "f" init_fTs
       
   822       ||>> mk_Frees "f" init_fTs
       
   823       ||>> mk_Frees "P" (replicate n (mk_pred1T initT));
       
   824 
       
   825     val II = HOLogic.mk_Collect (fst iidx', IIT, list_exists_free (II_Bs @ II_ss)
       
   826       (HOLogic.mk_conj (HOLogic.mk_eq (iidx,
       
   827         Abs_IIT $ (HOLogic.mk_prod (HOLogic.mk_tuple II_Bs, HOLogic.mk_tuple II_ss))),
       
   828         mk_alg passive_UNIVs II_Bs II_ss)));
       
   829 
       
   830     val select_Bs = map (mk_nthN n (HOLogic.mk_fst (Rep_IIT $ iidx))) ks;
       
   831     val select_ss = map (mk_nthN n (HOLogic.mk_snd (Rep_IIT $ iidx))) ks;
       
   832 
       
   833     val str_init_binds = mk_internal_bs str_initN;
       
   834     fun str_init_bind i = nth str_init_binds (i - 1);
       
   835     val str_init_name = Binding.name_of o str_init_bind;
       
   836     val str_init_def_bind = rpair [] o Thm.def_binding o str_init_bind;
       
   837 
       
   838     fun str_init_spec i =
       
   839       let
       
   840         val T = nth init_FTs (i - 1);
       
   841         val init_xF = nth init_xFs (i - 1)
       
   842         val select_s = nth select_ss (i - 1);
       
   843         val map = mk_map_of_bnf (nth Dss (i - 1))
       
   844           (passiveAs @ active_initTs) (passiveAs @ replicate n Asuc_bdT)
       
   845           (nth bnfs (i - 1));
       
   846         val map_args = passive_ids @ replicate n (mk_rapp iidx Asuc_bdT);
       
   847         val str_initT = T --> IIT --> Asuc_bdT;
       
   848 
       
   849         val lhs = Term.list_comb (Free (str_init_name i, str_initT), [init_xF, iidx]);
       
   850         val rhs = select_s $ (Term.list_comb (map, map_args) $ init_xF);
       
   851       in
       
   852         mk_Trueprop_eq (lhs, rhs)
       
   853       end;
       
   854 
       
   855     val ((str_init_frees, (_, str_init_def_frees)), (lthy, lthy_old)) =
       
   856       lthy
       
   857       |> fold_map (fn i => Specification.definition
       
   858         (SOME (str_init_bind i, NONE, NoSyn), (str_init_def_bind i, str_init_spec i))) ks
       
   859       |>> apsnd split_list o split_list
       
   860       ||> `Local_Theory.restore;
       
   861 
       
   862     val phi = Proof_Context.export_morphism lthy_old lthy;
       
   863     val str_inits =
       
   864       map (Term.subst_atomic_types (map (`(Morphism.typ phi)) params') o Morphism.term phi)
       
   865         str_init_frees;
       
   866 
       
   867     val str_init_defs = map (Morphism.thm phi) str_init_def_frees;
       
   868 
       
   869     val car_inits = map (mk_min_alg passive_UNIVs str_inits) ks;
       
   870 
       
   871     (*TODO: replace with instantiate? (problem: figure out right type instantiation)*)
       
   872     val alg_init_thm = Goal.prove_sorry lthy [] []
       
   873       (HOLogic.mk_Trueprop (mk_alg passive_UNIVs car_inits str_inits))
       
   874       (K (rtac alg_min_alg_thm 1))
       
   875       |> Thm.close_derivation;
       
   876 
       
   877     val alg_select_thm = Goal.prove_sorry lthy [] []
       
   878       (HOLogic.mk_Trueprop (mk_Ball II
       
   879         (Term.absfree iidx' (mk_alg passive_UNIVs select_Bs select_ss))))
       
   880       (mk_alg_select_tac Abs_IIT_inverse_thm)
       
   881       |> Thm.close_derivation;
       
   882 
       
   883     val mor_select_thm =
       
   884       let
       
   885         val alg_prem = HOLogic.mk_Trueprop (mk_alg passive_UNIVs Bs ss);
       
   886         val i_prem = HOLogic.mk_Trueprop (HOLogic.mk_mem (iidx, II));
       
   887         val mor_prem = HOLogic.mk_Trueprop (mk_mor select_Bs select_ss Bs ss Asuc_fs);
       
   888         val prems = [alg_prem, i_prem, mor_prem];
       
   889         val concl = HOLogic.mk_Trueprop
       
   890           (mk_mor car_inits str_inits Bs ss
       
   891             (map (fn f => HOLogic.mk_comp (f, mk_rapp iidx Asuc_bdT)) Asuc_fs));
       
   892       in
       
   893         Goal.prove_sorry lthy [] []
       
   894           (fold_rev Logic.all (iidx :: Bs @ ss @ Asuc_fs) (Logic.list_implies (prems, concl)))
       
   895           (K (mk_mor_select_tac mor_def mor_cong_thm mor_comp_thm mor_incl_min_alg_thm alg_def
       
   896             alg_select_thm alg_set_thms set_mapss str_init_defs))
       
   897         |> Thm.close_derivation
       
   898       end;
       
   899 
       
   900     val (init_ex_mor_thm, init_unique_mor_thms) =
       
   901       let
       
   902         val prem = HOLogic.mk_Trueprop (mk_alg passive_UNIVs Bs ss);
       
   903         val concl = HOLogic.mk_Trueprop
       
   904           (list_exists_free init_fs (mk_mor car_inits str_inits Bs ss init_fs));
       
   905         val ex_mor = Goal.prove_sorry lthy [] []
       
   906           (fold_rev Logic.all (Bs @ ss) (Logic.mk_implies (prem, concl)))
       
   907           (mk_init_ex_mor_tac Abs_IIT_inverse_thm ex_copy_alg_thm alg_min_alg_thm
       
   908             card_of_min_alg_thms mor_comp_thm mor_select_thm mor_incl_min_alg_thm)
       
   909           |> Thm.close_derivation;
       
   910 
       
   911         val prems = map2 (HOLogic.mk_Trueprop oo curry HOLogic.mk_mem) init_xs car_inits
       
   912         val mor_prems = map HOLogic.mk_Trueprop
       
   913           [mk_mor car_inits str_inits Bs ss init_fs,
       
   914           mk_mor car_inits str_inits Bs ss init_fs_copy];
       
   915         fun mk_fun_eq f g x = HOLogic.mk_eq (f $ x, g $ x);
       
   916         val unique = HOLogic.mk_Trueprop
       
   917           (Library.foldr1 HOLogic.mk_conj (map3 mk_fun_eq init_fs init_fs_copy init_xs));
       
   918         val unique_mor = Goal.prove_sorry lthy [] []
       
   919           (fold_rev Logic.all (init_xs @ Bs @ ss @ init_fs @ init_fs_copy)
       
   920             (Logic.list_implies (prems @ mor_prems, unique)))
       
   921           (K (mk_init_unique_mor_tac m alg_def alg_init_thm least_min_alg_thms
       
   922             in_mono'_thms alg_set_thms morE_thms map_cong0s))
       
   923           |> Thm.close_derivation;
       
   924       in
       
   925         (ex_mor, split_conj_thm unique_mor)
       
   926       end;
       
   927 
       
   928     val init_setss = mk_setss (passiveAs @ active_initTs);
       
   929     val active_init_setss = map (drop m) init_setss;
       
   930     val init_ins = map2 (fn sets => mk_in (passive_UNIVs @ car_inits) sets) init_setss init_FTs;
       
   931 
       
   932     fun mk_closed phis =
       
   933       let
       
   934         fun mk_conjunct phi str_init init_sets init_in x x' =
       
   935           let
       
   936             val prem = Library.foldr1 HOLogic.mk_conj
       
   937               (map2 (fn set => mk_Ball (set $ x)) init_sets phis);
       
   938             val concl = phi $ (str_init $ x);
       
   939           in
       
   940             mk_Ball init_in (Term.absfree x' (HOLogic.mk_imp (prem, concl)))
       
   941           end;
       
   942       in
       
   943         Library.foldr1 HOLogic.mk_conj
       
   944           (map6 mk_conjunct phis str_inits active_init_setss init_ins init_xFs init_xFs')
       
   945       end;
       
   946 
       
   947     val init_induct_thm =
       
   948       let
       
   949         val prem = HOLogic.mk_Trueprop (mk_closed init_phis);
       
   950         val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
       
   951           (map2 mk_Ball car_inits init_phis));
       
   952       in
       
   953         Goal.prove_sorry lthy [] []
       
   954           (fold_rev Logic.all init_phis (Logic.mk_implies (prem, concl)))
       
   955           (K (mk_init_induct_tac m alg_def alg_init_thm least_min_alg_thms alg_set_thms))
       
   956         |> Thm.close_derivation
       
   957       end;
       
   958 
       
   959     val timer = time (timer "Initiality definition & thms");
       
   960 
       
   961     val ((T_names, (T_glob_infos, T_loc_infos)), lthy) =
       
   962       lthy
       
   963       |> fold_map3 (fn b => fn mx => fn car_init =>
       
   964         typedef (Binding.conceal b, params, mx) car_init NONE
       
   965           (EVERY' [rtac ssubst, rtac @{thm ex_in_conv}, resolve_tac alg_not_empty_thms,
       
   966             rtac alg_init_thm] 1)) bs mixfixes car_inits
       
   967       |>> apsnd split_list o split_list;
       
   968 
       
   969     val Ts = map (fn name => Type (name, params')) T_names;
       
   970     fun mk_Ts passive = map (Term.typ_subst_atomic (passiveAs ~~ passive)) Ts;
       
   971     val Ts' = mk_Ts passiveBs;
       
   972     val Rep_Ts = map2 (fn info => fn T => Const (#Rep_name info, T --> initT)) T_glob_infos Ts;
       
   973     val Abs_Ts = map2 (fn info => fn T => Const (#Abs_name info, initT --> T)) T_glob_infos Ts;
       
   974 
       
   975     val type_defs = map #type_definition T_loc_infos;
       
   976     val Reps = map #Rep T_loc_infos;
       
   977     val Rep_casess = map #Rep_cases T_loc_infos;
       
   978     val Rep_injects = map #Rep_inject T_loc_infos;
       
   979     val Rep_inverses = map #Rep_inverse T_loc_infos;
       
   980     val Abs_inverses = map #Abs_inverse T_loc_infos;
       
   981 
       
   982     fun mk_inver_thm mk_tac rep abs X thm =
       
   983       Goal.prove_sorry lthy [] []
       
   984         (HOLogic.mk_Trueprop (mk_inver rep abs X))
       
   985         (K (EVERY' [rtac ssubst, rtac @{thm inver_def}, rtac ballI, mk_tac thm] 1))
       
   986       |> Thm.close_derivation;
       
   987 
       
   988     val inver_Reps = map4 (mk_inver_thm rtac) Abs_Ts Rep_Ts (map HOLogic.mk_UNIV Ts) Rep_inverses;
       
   989     val inver_Abss = map4 (mk_inver_thm etac) Rep_Ts Abs_Ts car_inits Abs_inverses;
       
   990 
       
   991     val timer = time (timer "THE TYPEDEFs & Rep/Abs thms");
       
   992 
       
   993     val UNIVs = map HOLogic.mk_UNIV Ts;
       
   994     val FTs = mk_FTs (passiveAs @ Ts);
       
   995     val FTs' = mk_FTs (passiveBs @ Ts');
       
   996     fun mk_set_Ts T = passiveAs @ replicate n (HOLogic.mk_setT T);
       
   997     val setFTss = map (mk_FTs o mk_set_Ts) passiveAs;
       
   998     val FTs_setss = mk_setss (passiveAs @ Ts);
       
   999     val FTs'_setss = mk_setss (passiveBs @ Ts');
       
  1000     val map_FT_inits = map2 (fn Ds =>
       
  1001       mk_map_of_bnf Ds (passiveAs @ Ts) (passiveAs @ active_initTs)) Dss bnfs;
       
  1002     val fTs = map2 (curry op -->) Ts activeAs;
       
  1003     val foldT = Library.foldr1 HOLogic.mk_prodT (map2 (curry op -->) Ts activeAs);
       
  1004     val rec_sTs = map (Term.typ_subst_atomic (activeBs ~~ Ts)) prod_sTs;
       
  1005     val rec_maps = map (Term.subst_atomic_types (activeBs ~~ Ts)) map_fsts;
       
  1006     val rec_maps_rev = map (Term.subst_atomic_types (activeBs ~~ Ts)) map_fsts_rev;
       
  1007     val rec_fsts = map (Term.subst_atomic_types (activeBs ~~ Ts)) fsts;
       
  1008     val rec_UNIVs = map2 (HOLogic.mk_UNIV oo curry HOLogic.mk_prodT) Ts activeAs;
       
  1009 
       
  1010     val (((((((((Izs1, Izs1'), (Izs2, Izs2')), (xFs, xFs')), yFs), (AFss, AFss')),
       
  1011       (fold_f, fold_f')), fs), rec_ss), names_lthy) = names_lthy
       
  1012       |> mk_Frees' "z1" Ts
       
  1013       ||>> mk_Frees' "z2" Ts'
       
  1014       ||>> mk_Frees' "x" FTs
       
  1015       ||>> mk_Frees "y" FTs'
       
  1016       ||>> mk_Freess' "z" setFTss
       
  1017       ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "f") foldT
       
  1018       ||>> mk_Frees "f" fTs
       
  1019       ||>> mk_Frees "s" rec_sTs;
       
  1020 
       
  1021     val Izs = map2 retype_free Ts zs;
       
  1022     val phis = map2 retype_free (map mk_pred1T Ts) init_phis;
       
  1023     val phi2s = map2 retype_free (map2 mk_pred2T Ts Ts') init_phis;
       
  1024 
       
  1025     fun ctor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctorN ^ "_");
       
  1026     val ctor_name = Binding.name_of o ctor_bind;
       
  1027     val ctor_def_bind = rpair [] o Binding.conceal o Thm.def_binding o ctor_bind;
       
  1028 
       
  1029     fun ctor_spec i abs str map_FT_init x x' =
       
  1030       let
       
  1031         val ctorT = nth FTs (i - 1) --> nth Ts (i - 1);
       
  1032 
       
  1033         val lhs = Free (ctor_name i, ctorT);
       
  1034         val rhs = Term.absfree x' (abs $ (str $
       
  1035           (Term.list_comb (map_FT_init, map HOLogic.id_const passiveAs @ Rep_Ts) $ x)));
       
  1036       in
       
  1037         mk_Trueprop_eq (lhs, rhs)
       
  1038       end;
       
  1039 
       
  1040     val ((ctor_frees, (_, ctor_def_frees)), (lthy, lthy_old)) =
       
  1041       lthy
       
  1042       |> fold_map6 (fn i => fn abs => fn str => fn mapx => fn x => fn x' =>
       
  1043         Specification.definition
       
  1044           (SOME (ctor_bind i, NONE, NoSyn), (ctor_def_bind i, ctor_spec i abs str mapx x x')))
       
  1045           ks Abs_Ts str_inits map_FT_inits xFs xFs'
       
  1046       |>> apsnd split_list o split_list
       
  1047       ||> `Local_Theory.restore;
       
  1048 
       
  1049     val phi = Proof_Context.export_morphism lthy_old lthy;
       
  1050     fun mk_ctors passive =
       
  1051       map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ (mk_params passive)) o
       
  1052         Morphism.term phi) ctor_frees;
       
  1053     val ctors = mk_ctors passiveAs;
       
  1054     val ctor's = mk_ctors passiveBs;
       
  1055     val ctor_defs = map (Morphism.thm phi) ctor_def_frees;
       
  1056 
       
  1057     val (mor_Rep_thm, mor_Abs_thm) =
       
  1058       let
       
  1059         val copy = alg_init_thm RS copy_alg_thm;
       
  1060         fun mk_bij inj Rep cases = @{thm bij_betwI'} OF [inj, Rep, cases];
       
  1061         val bijs = map3 mk_bij Rep_injects Reps Rep_casess;
       
  1062         val mor_Rep =
       
  1063           Goal.prove_sorry lthy [] []
       
  1064             (HOLogic.mk_Trueprop (mk_mor UNIVs ctors car_inits str_inits Rep_Ts))
       
  1065             (mk_mor_Rep_tac ctor_defs copy bijs inver_Abss inver_Reps)
       
  1066           |> Thm.close_derivation;
       
  1067 
       
  1068         val inv = mor_inv_thm OF [mor_Rep, talg_thm, alg_init_thm];
       
  1069         val mor_Abs =
       
  1070           Goal.prove_sorry lthy [] []
       
  1071             (HOLogic.mk_Trueprop (mk_mor car_inits str_inits UNIVs ctors Abs_Ts))
       
  1072             (K (mk_mor_Abs_tac inv inver_Abss inver_Reps))
       
  1073           |> Thm.close_derivation;
       
  1074       in
       
  1075         (mor_Rep, mor_Abs)
       
  1076       end;
       
  1077 
       
  1078     val timer = time (timer "ctor definitions & thms");
       
  1079 
       
  1080     val fold_fun = Term.absfree fold_f'
       
  1081       (mk_mor UNIVs ctors active_UNIVs ss (map (mk_nthN n fold_f) ks));
       
  1082     val foldx = HOLogic.choice_const foldT $ fold_fun;
       
  1083 
       
  1084     fun fold_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctor_foldN ^ "_");
       
  1085     val fold_name = Binding.name_of o fold_bind;
       
  1086     val fold_def_bind = rpair [] o Binding.conceal o Thm.def_binding o fold_bind;
       
  1087 
       
  1088     fun fold_spec i T AT =
       
  1089       let
       
  1090         val foldT = Library.foldr (op -->) (sTs, T --> AT);
       
  1091 
       
  1092         val lhs = Term.list_comb (Free (fold_name i, foldT), ss);
       
  1093         val rhs = mk_nthN n foldx i;
       
  1094       in
       
  1095         mk_Trueprop_eq (lhs, rhs)
       
  1096       end;
       
  1097 
       
  1098     val ((fold_frees, (_, fold_def_frees)), (lthy, lthy_old)) =
       
  1099       lthy
       
  1100       |> fold_map3 (fn i => fn T => fn AT =>
       
  1101         Specification.definition
       
  1102           (SOME (fold_bind i, NONE, NoSyn), (fold_def_bind i, fold_spec i T AT)))
       
  1103           ks Ts activeAs
       
  1104       |>> apsnd split_list o split_list
       
  1105       ||> `Local_Theory.restore;
       
  1106 
       
  1107     val phi = Proof_Context.export_morphism lthy_old lthy;
       
  1108     val folds = map (Morphism.term phi) fold_frees;
       
  1109     val fold_names = map (fst o dest_Const) folds;
       
  1110     fun mk_folds passives actives =
       
  1111       map3 (fn name => fn T => fn active =>
       
  1112         Const (name, Library.foldr (op -->)
       
  1113           (map2 (curry op -->) (mk_FTs (passives @ actives)) actives, T --> active)))
       
  1114       fold_names (mk_Ts passives) actives;
       
  1115     fun mk_fold Ts ss i = Term.list_comb (Const (nth fold_names (i - 1), Library.foldr (op -->)
       
  1116       (map fastype_of ss, nth Ts (i - 1) --> range_type (fastype_of (nth ss (i - 1))))), ss);
       
  1117     val fold_defs = map (Morphism.thm phi) fold_def_frees;
       
  1118 
       
  1119     val mor_fold_thm =
       
  1120       let
       
  1121         val ex_mor = talg_thm RS init_ex_mor_thm;
       
  1122         val mor_cong = mor_cong_thm OF (map (mk_nth_conv n) ks);
       
  1123         val mor_comp = mor_Rep_thm RS mor_comp_thm;
       
  1124         val cT = certifyT lthy foldT;
       
  1125         val ct = certify lthy fold_fun
       
  1126       in
       
  1127         singleton (Proof_Context.export names_lthy lthy)
       
  1128           (Goal.prove_sorry lthy [] []
       
  1129             (HOLogic.mk_Trueprop (mk_mor UNIVs ctors active_UNIVs ss (map (mk_fold Ts ss) ks)))
       
  1130             (K (mk_mor_fold_tac cT ct fold_defs ex_mor (mor_comp RS mor_cong))))
       
  1131         |> Thm.close_derivation
       
  1132       end;
       
  1133 
       
  1134     val ctor_fold_thms = map (fn morE => rule_by_tactic lthy
       
  1135       ((rtac CollectI THEN' CONJ_WRAP' (K (rtac @{thm subset_UNIV})) (1 upto m + n)) 1)
       
  1136       (mor_fold_thm RS morE)) morE_thms;
       
  1137 
       
  1138     val (fold_unique_mor_thms, fold_unique_mor_thm) =
       
  1139       let
       
  1140         val prem = HOLogic.mk_Trueprop (mk_mor UNIVs ctors active_UNIVs ss fs);
       
  1141         fun mk_fun_eq f i = HOLogic.mk_eq (f, mk_fold Ts ss i);
       
  1142         val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_fun_eq fs ks));
       
  1143         val unique_mor = Goal.prove_sorry lthy [] []
       
  1144           (fold_rev Logic.all (ss @ fs) (Logic.mk_implies (prem, unique)))
       
  1145           (K (mk_fold_unique_mor_tac type_defs init_unique_mor_thms Reps
       
  1146             mor_comp_thm mor_Abs_thm mor_fold_thm))
       
  1147           |> Thm.close_derivation;
       
  1148       in
       
  1149         `split_conj_thm unique_mor
       
  1150       end;
       
  1151 
       
  1152     val (ctor_fold_unique_thms, ctor_fold_unique_thm) =
       
  1153       `split_conj_thm (mk_conjIN n RS
       
  1154         (mor_UNIV_thm RS iffD2 RS fold_unique_mor_thm))
       
  1155 
       
  1156     val fold_ctor_thms =
       
  1157       map (fn thm => (mor_incl_thm OF replicate n @{thm subset_UNIV}) RS thm RS sym)
       
  1158         fold_unique_mor_thms;
       
  1159 
       
  1160     val ctor_o_fold_thms =
       
  1161       let
       
  1162         val mor = mor_comp_thm OF [mor_fold_thm, mor_str_thm];
       
  1163       in
       
  1164         map2 (fn unique => fn fold_ctor =>
       
  1165           trans OF [mor RS unique, fold_ctor]) fold_unique_mor_thms fold_ctor_thms
       
  1166       end;
       
  1167 
       
  1168     val timer = time (timer "fold definitions & thms");
       
  1169 
       
  1170     val map_ctors = map2 (fn Ds => fn bnf =>
       
  1171       Term.list_comb (mk_map_of_bnf Ds (passiveAs @ FTs) (passiveAs @ Ts) bnf,
       
  1172         map HOLogic.id_const passiveAs @ ctors)) Dss bnfs;
       
  1173 
       
  1174     fun dtor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtorN ^ "_");
       
  1175     val dtor_name = Binding.name_of o dtor_bind;
       
  1176     val dtor_def_bind = rpair [] o Binding.conceal o Thm.def_binding o dtor_bind;
       
  1177 
       
  1178     fun dtor_spec i FT T =
       
  1179       let
       
  1180         val dtorT = T --> FT;
       
  1181 
       
  1182         val lhs = Free (dtor_name i, dtorT);
       
  1183         val rhs = mk_fold Ts map_ctors i;
       
  1184       in
       
  1185         mk_Trueprop_eq (lhs, rhs)
       
  1186       end;
       
  1187 
       
  1188     val ((dtor_frees, (_, dtor_def_frees)), (lthy, lthy_old)) =
       
  1189       lthy
       
  1190       |> fold_map3 (fn i => fn FT => fn T =>
       
  1191         Specification.definition
       
  1192           (SOME (dtor_bind i, NONE, NoSyn), (dtor_def_bind i, dtor_spec i FT T))) ks FTs Ts
       
  1193       |>> apsnd split_list o split_list
       
  1194       ||> `Local_Theory.restore;
       
  1195 
       
  1196     val phi = Proof_Context.export_morphism lthy_old lthy;
       
  1197     fun mk_dtors params =
       
  1198       map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ params) o Morphism.term phi)
       
  1199         dtor_frees;
       
  1200     val dtors = mk_dtors params';
       
  1201     val dtor_defs = map (Morphism.thm phi) dtor_def_frees;
       
  1202 
       
  1203     val ctor_o_dtor_thms = map2 (fold_thms lthy o single) dtor_defs ctor_o_fold_thms;
       
  1204 
       
  1205     val dtor_o_ctor_thms =
       
  1206       let
       
  1207         fun mk_goal dtor ctor FT =
       
  1208           mk_Trueprop_eq (HOLogic.mk_comp (dtor, ctor), HOLogic.id_const FT);
       
  1209         val goals = map3 mk_goal dtors ctors FTs;
       
  1210       in
       
  1211         map5 (fn goal => fn dtor_def => fn foldx => fn map_comp_id => fn map_cong0L =>
       
  1212           Goal.prove_sorry lthy [] [] goal
       
  1213             (K (mk_dtor_o_ctor_tac dtor_def foldx map_comp_id map_cong0L ctor_o_fold_thms))
       
  1214           |> Thm.close_derivation)
       
  1215         goals dtor_defs ctor_fold_thms map_comp_id_thms map_cong0L_thms
       
  1216       end;
       
  1217 
       
  1218     val dtor_ctor_thms = map (fn thm => thm RS @{thm pointfree_idE}) dtor_o_ctor_thms;
       
  1219     val ctor_dtor_thms = map (fn thm => thm RS @{thm pointfree_idE}) ctor_o_dtor_thms;
       
  1220 
       
  1221     val bij_dtor_thms =
       
  1222       map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) ctor_o_dtor_thms dtor_o_ctor_thms;
       
  1223     val inj_dtor_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_dtor_thms;
       
  1224     val surj_dtor_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_dtor_thms;
       
  1225     val dtor_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_dtor_thms;
       
  1226     val dtor_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_dtor_thms;
       
  1227     val dtor_exhaust_thms = map (fn thm => thm RS exE) dtor_nchotomy_thms;
       
  1228 
       
  1229     val bij_ctor_thms =
       
  1230       map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) dtor_o_ctor_thms ctor_o_dtor_thms;
       
  1231     val inj_ctor_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_ctor_thms;
       
  1232     val surj_ctor_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_ctor_thms;
       
  1233     val ctor_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_ctor_thms;
       
  1234     val ctor_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_ctor_thms;
       
  1235     val ctor_exhaust_thms = map (fn thm => thm RS exE) ctor_nchotomy_thms;
       
  1236 
       
  1237     val timer = time (timer "dtor definitions & thms");
       
  1238 
       
  1239     val fst_rec_pair_thms =
       
  1240       let
       
  1241         val mor = mor_comp_thm OF [mor_fold_thm, mor_convol_thm];
       
  1242       in
       
  1243         map2 (fn unique => fn fold_ctor =>
       
  1244           trans OF [mor RS unique, fold_ctor]) fold_unique_mor_thms fold_ctor_thms
       
  1245       end;
       
  1246 
       
  1247     fun rec_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctor_recN ^ "_");
       
  1248     val rec_name = Binding.name_of o rec_bind;
       
  1249     val rec_def_bind = rpair [] o Binding.conceal o Thm.def_binding o rec_bind;
       
  1250 
       
  1251     val rec_strs =
       
  1252       map3 (fn ctor => fn prod_s => fn mapx =>
       
  1253         mk_convol (HOLogic.mk_comp (ctor, Term.list_comb (mapx, passive_ids @ rec_fsts)), prod_s))
       
  1254       ctors rec_ss rec_maps;
       
  1255 
       
  1256     fun rec_spec i T AT =
       
  1257       let
       
  1258         val recT = Library.foldr (op -->) (rec_sTs, T --> AT);
       
  1259 
       
  1260         val lhs = Term.list_comb (Free (rec_name i, recT), rec_ss);
       
  1261         val rhs = HOLogic.mk_comp (snd_const (HOLogic.mk_prodT (T, AT)), mk_fold Ts rec_strs i);
       
  1262       in
       
  1263         mk_Trueprop_eq (lhs, rhs)
       
  1264       end;
       
  1265 
       
  1266     val ((rec_frees, (_, rec_def_frees)), (lthy, lthy_old)) =
       
  1267       lthy
       
  1268       |> fold_map3 (fn i => fn T => fn AT =>
       
  1269         Specification.definition
       
  1270           (SOME (rec_bind i, NONE, NoSyn), (rec_def_bind i, rec_spec i T AT)))
       
  1271           ks Ts activeAs
       
  1272       |>> apsnd split_list o split_list
       
  1273       ||> `Local_Theory.restore;
       
  1274 
       
  1275     val phi = Proof_Context.export_morphism lthy_old lthy;
       
  1276     val recs = map (Morphism.term phi) rec_frees;
       
  1277     val rec_names = map (fst o dest_Const) recs;
       
  1278     fun mk_rec ss i = Term.list_comb (Const (nth rec_names (i - 1), Library.foldr (op -->)
       
  1279       (map fastype_of ss, nth Ts (i - 1) --> range_type (fastype_of (nth ss (i - 1))))), ss);
       
  1280     val rec_defs = map (Morphism.thm phi) rec_def_frees;
       
  1281 
       
  1282     val convols = map2 (fn T => fn i => mk_convol (HOLogic.id_const T, mk_rec rec_ss i)) Ts ks;
       
  1283     val ctor_rec_thms =
       
  1284       let
       
  1285         fun mk_goal i rec_s rec_map ctor x =
       
  1286           let
       
  1287             val lhs = mk_rec rec_ss i $ (ctor $ x);
       
  1288             val rhs = rec_s $ (Term.list_comb (rec_map, passive_ids @ convols) $ x);
       
  1289           in
       
  1290             fold_rev Logic.all (x :: rec_ss) (mk_Trueprop_eq (lhs, rhs))
       
  1291           end;
       
  1292         val goals = map5 mk_goal ks rec_ss rec_maps_rev ctors xFs;
       
  1293       in
       
  1294         map2 (fn goal => fn foldx =>
       
  1295           Goal.prove_sorry lthy [] [] goal (mk_rec_tac rec_defs foldx fst_rec_pair_thms)
       
  1296           |> Thm.close_derivation)
       
  1297         goals ctor_fold_thms
       
  1298       end;
       
  1299 
       
  1300     val rec_unique_mor_thm =
       
  1301       let
       
  1302         val id_fs = map2 (fn T => fn f => mk_convol (HOLogic.id_const T, f)) Ts fs;
       
  1303         val prem = HOLogic.mk_Trueprop (mk_mor UNIVs ctors rec_UNIVs rec_strs id_fs);
       
  1304         fun mk_fun_eq f i = HOLogic.mk_eq (f, mk_rec rec_ss i);
       
  1305         val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_fun_eq fs ks));
       
  1306       in
       
  1307         Goal.prove_sorry lthy [] []
       
  1308           (fold_rev Logic.all (rec_ss @ fs) (Logic.mk_implies (prem, unique)))
       
  1309           (mk_rec_unique_mor_tac rec_defs fst_rec_pair_thms fold_unique_mor_thm)
       
  1310           |> Thm.close_derivation
       
  1311       end;
       
  1312 
       
  1313     val (ctor_rec_unique_thms, ctor_rec_unique_thm) =
       
  1314       `split_conj_thm (split_conj_prems n
       
  1315         (mor_UNIV_thm RS iffD2 RS rec_unique_mor_thm)
       
  1316         |> Local_Defs.unfold lthy (@{thms convol_o o_id id_o o_assoc[symmetric] fst_convol} @
       
  1317            map_id0s @ sym_map_comps) OF replicate n @{thm arg_cong2[of _ _ _ _ convol, OF refl]});
       
  1318 
       
  1319     val timer = time (timer "rec definitions & thms");
       
  1320 
       
  1321     val (ctor_induct_thm, induct_params) =
       
  1322       let
       
  1323         fun mk_prem phi ctor sets x =
       
  1324           let
       
  1325             fun mk_IH phi set z =
       
  1326               let
       
  1327                 val prem = HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set $ x));
       
  1328                 val concl = HOLogic.mk_Trueprop (phi $ z);
       
  1329               in
       
  1330                 Logic.all z (Logic.mk_implies (prem, concl))
       
  1331               end;
       
  1332 
       
  1333             val IHs = map3 mk_IH phis (drop m sets) Izs;
       
  1334             val concl = HOLogic.mk_Trueprop (phi $ (ctor $ x));
       
  1335           in
       
  1336             Logic.all x (Logic.list_implies (IHs, concl))
       
  1337           end;
       
  1338 
       
  1339         val prems = map4 mk_prem phis ctors FTs_setss xFs;
       
  1340 
       
  1341         fun mk_concl phi z = phi $ z;
       
  1342         val concl =
       
  1343           HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_concl phis Izs));
       
  1344 
       
  1345         val goal = Logic.list_implies (prems, concl);
       
  1346       in
       
  1347         (Goal.prove_sorry lthy [] []
       
  1348           (fold_rev Logic.all (phis @ Izs) goal)
       
  1349           (K (mk_ctor_induct_tac lthy m set_mapss init_induct_thm morE_thms mor_Abs_thm
       
  1350             Rep_inverses Abs_inverses Reps))
       
  1351         |> Thm.close_derivation,
       
  1352         rev (Term.add_tfrees goal []))
       
  1353       end;
       
  1354 
       
  1355     val cTs = map (SOME o certifyT lthy o TFree) induct_params;
       
  1356 
       
  1357     val weak_ctor_induct_thms =
       
  1358       let fun insts i = (replicate (i - 1) TrueI) @ (asm_rl :: replicate (n - i) TrueI);
       
  1359       in map (fn i => (ctor_induct_thm OF insts i) RS mk_conjunctN n i) ks end;
       
  1360 
       
  1361     val (ctor_induct2_thm, induct2_params) =
       
  1362       let
       
  1363         fun mk_prem phi ctor ctor' sets sets' x y =
       
  1364           let
       
  1365             fun mk_IH phi set set' z1 z2 =
       
  1366               let
       
  1367                 val prem1 = HOLogic.mk_Trueprop (HOLogic.mk_mem (z1, (set $ x)));
       
  1368                 val prem2 = HOLogic.mk_Trueprop (HOLogic.mk_mem (z2, (set' $ y)));
       
  1369                 val concl = HOLogic.mk_Trueprop (phi $ z1 $ z2);
       
  1370               in
       
  1371                 fold_rev Logic.all [z1, z2] (Logic.list_implies ([prem1, prem2], concl))
       
  1372               end;
       
  1373 
       
  1374             val IHs = map5 mk_IH phi2s (drop m sets) (drop m sets') Izs1 Izs2;
       
  1375             val concl = HOLogic.mk_Trueprop (phi $ (ctor $ x) $ (ctor' $ y));
       
  1376           in
       
  1377             fold_rev Logic.all [x, y] (Logic.list_implies (IHs, concl))
       
  1378           end;
       
  1379 
       
  1380         val prems = map7 mk_prem phi2s ctors ctor's FTs_setss FTs'_setss xFs yFs;
       
  1381 
       
  1382         fun mk_concl phi z1 z2 = phi $ z1 $ z2;
       
  1383         val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
       
  1384           (map3 mk_concl phi2s Izs1 Izs2));
       
  1385         fun mk_t phi (z1, z1') (z2, z2') =
       
  1386           Term.absfree z1' (HOLogic.mk_all (fst z2', snd z2', phi $ z1 $ z2));
       
  1387         val cts = map3 (SOME o certify lthy ooo mk_t) phi2s (Izs1 ~~ Izs1') (Izs2 ~~ Izs2');
       
  1388         val goal = Logic.list_implies (prems, concl);
       
  1389       in
       
  1390         (singleton (Proof_Context.export names_lthy lthy)
       
  1391           (Goal.prove_sorry lthy [] [] goal
       
  1392             (mk_ctor_induct2_tac cTs cts ctor_induct_thm weak_ctor_induct_thms))
       
  1393           |> Thm.close_derivation,
       
  1394         rev (Term.add_tfrees goal []))
       
  1395       end;
       
  1396 
       
  1397     val timer = time (timer "induction");
       
  1398 
       
  1399     fun mk_ctor_map_DEADID_thm ctor_inject map_id0 =
       
  1400       trans OF [id_apply, iffD2 OF [ctor_inject, map_id0 RS sym]];
       
  1401 
       
  1402     fun mk_ctor_Irel_DEADID_thm ctor_inject bnf =
       
  1403       trans OF [ctor_inject, rel_eq_of_bnf bnf RS @{thm predicate2_eqD} RS sym];
       
  1404 
       
  1405     val IphiTs = map2 mk_pred2T passiveAs passiveBs;
       
  1406     val Ipsi1Ts = map2 mk_pred2T passiveAs passiveCs;
       
  1407     val Ipsi2Ts = map2 mk_pred2T passiveCs passiveBs;
       
  1408     val activephiTs = map2 mk_pred2T activeAs activeBs;
       
  1409     val activeIphiTs = map2 mk_pred2T Ts Ts';
       
  1410     val (((((Iphis, Ipsi1s), Ipsi2s), activephis), activeIphis), names_lthy) = names_lthy
       
  1411       |> mk_Frees "R" IphiTs
       
  1412       ||>> mk_Frees "R" Ipsi1Ts
       
  1413       ||>> mk_Frees "Q" Ipsi2Ts
       
  1414       ||>> mk_Frees "S" activephiTs
       
  1415       ||>> mk_Frees "IR" activeIphiTs;
       
  1416     val rels = map2 (fn Ds => mk_rel_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
       
  1417 
       
  1418     (*register new datatypes as BNFs*)
       
  1419     val (timer, Ibnfs, (ctor_Imap_o_thms, ctor_Imap_thms), ctor_Iset_thmss',
       
  1420         ctor_Irel_thms, Ibnf_notes, lthy) =
       
  1421       if m = 0 then
       
  1422         (timer, replicate n DEADID_bnf,
       
  1423         map_split (`(mk_pointfree lthy)) (map2 mk_ctor_map_DEADID_thm ctor_inject_thms map_ids),
       
  1424         replicate n [], map2 mk_ctor_Irel_DEADID_thm ctor_inject_thms bnfs, [], lthy)
       
  1425       else let
       
  1426         val fTs = map2 (curry op -->) passiveAs passiveBs;
       
  1427         val uTs = map2 (curry op -->) Ts Ts';
       
  1428 
       
  1429         val (((((fs, fs'), fs_copy), us), (ys, ys')),
       
  1430           names_lthy) = names_lthy
       
  1431           |> mk_Frees' "f" fTs
       
  1432           ||>> mk_Frees "f" fTs
       
  1433           ||>> mk_Frees "u" uTs
       
  1434           ||>> mk_Frees' "y" passiveAs;
       
  1435 
       
  1436         val map_FTFT's = map2 (fn Ds =>
       
  1437           mk_map_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
       
  1438         fun mk_passive_maps ATs BTs Ts =
       
  1439           map2 (fn Ds => mk_map_of_bnf Ds (ATs @ Ts) (BTs @ Ts)) Dss bnfs;
       
  1440         fun mk_map_fold_arg fs Ts ctor fmap =
       
  1441           HOLogic.mk_comp (ctor, Term.list_comb (fmap, fs @ map HOLogic.id_const Ts));
       
  1442         fun mk_map Ts fs Ts' ctors mk_maps =
       
  1443           mk_fold Ts (map2 (mk_map_fold_arg fs Ts') ctors (mk_maps Ts'));
       
  1444         val pmapsABT' = mk_passive_maps passiveAs passiveBs;
       
  1445         val fs_maps = map (mk_map Ts fs Ts' ctor's pmapsABT') ks;
       
  1446 
       
  1447         val ls = 1 upto m;
       
  1448         val setsss = map (mk_setss o mk_set_Ts) passiveAs;
       
  1449 
       
  1450         fun mk_col l T z z' sets =
       
  1451           let
       
  1452             fun mk_UN set = mk_Union T $ (set $ z);
       
  1453           in
       
  1454             Term.absfree z'
       
  1455               (mk_union (nth sets (l - 1) $ z,
       
  1456                 Library.foldl1 mk_union (map mk_UN (drop m sets))))
       
  1457           end;
       
  1458 
       
  1459         val colss = map5 (fn l => fn T => map3 (mk_col l T)) ls passiveAs AFss AFss' setsss;
       
  1460         val setss_by_range = map (fn cols => map (mk_fold Ts cols) ks) colss;
       
  1461         val setss_by_bnf = transpose setss_by_range;
       
  1462 
       
  1463         val set_bss =
       
  1464           map (flat o map2 (fn B => fn b =>
       
  1465             if member (op =) resDs (TFree B) then [] else [b]) resBs) set_bss0;
       
  1466 
       
  1467         val ctor_witss =
       
  1468           let
       
  1469             val witss = map2 (fn Ds => fn bnf => mk_wits_of_bnf
       
  1470               (replicate (nwits_of_bnf bnf) Ds)
       
  1471               (replicate (nwits_of_bnf bnf) (passiveAs @ Ts)) bnf) Dss bnfs;
       
  1472             fun close_wit (I, wit) = fold_rev Term.absfree (map (nth ys') I) wit;
       
  1473             fun wit_apply (arg_I, arg_wit) (fun_I, fun_wit) =
       
  1474               (union (op =) arg_I fun_I, fun_wit $ arg_wit);
       
  1475 
       
  1476             fun gen_arg support i =
       
  1477               if i < m then [([i], nth ys i)]
       
  1478               else maps (mk_wit support (nth ctors (i - m)) (i - m)) (nth support (i - m))
       
  1479             and mk_wit support ctor i (I, wit) =
       
  1480               let val args = map (gen_arg (nth_map i (remove (op =) (I, wit)) support)) I;
       
  1481               in
       
  1482                 (args, [([], wit)])
       
  1483                 |-> fold (map_product wit_apply)
       
  1484                 |> map (apsnd (fn t => ctor $ t))
       
  1485                 |> minimize_wits
       
  1486               end;
       
  1487           in
       
  1488             map3 (fn ctor => fn i => map close_wit o minimize_wits o maps (mk_wit witss ctor i))
       
  1489               ctors (0 upto n - 1) witss
       
  1490           end;
       
  1491 
       
  1492         val (Ibnf_consts, lthy) =
       
  1493           fold_map8 (fn b => fn map_b => fn rel_b => fn set_bs => fn mapx => fn sets => fn wits =>
       
  1494               fn T => fn lthy =>
       
  1495             define_bnf_consts Dont_Inline (user_policy Note_Some lthy) (SOME deads)
       
  1496               map_b rel_b set_bs
       
  1497               ((((((b, T), fold_rev Term.absfree fs' mapx), sets), sum_bd), wits), NONE) lthy)
       
  1498           bs map_bs rel_bs set_bss fs_maps setss_by_bnf ctor_witss Ts lthy;
       
  1499 
       
  1500         val (_, Iconsts, Iconst_defs, mk_Iconsts) = split_list4 Ibnf_consts;
       
  1501         val (_, Isetss, Ibds_Ds, Iwitss_Ds, _) = split_list5 Iconsts;
       
  1502         val (Imap_defs, Iset_defss, Ibd_defs, Iwit_defss, Irel_defs) = split_list5 Iconst_defs;
       
  1503         val (mk_Imaps_Ds, mk_It_Ds, _, mk_Irels_Ds, _) = split_list5 mk_Iconsts;
       
  1504 
       
  1505         val Irel_unabs_defs = map (fn def => mk_unabs_def m (def RS meta_eq_to_obj_eq)) Irel_defs;
       
  1506         val Iset_defs = flat Iset_defss;
       
  1507 
       
  1508         fun mk_Imaps As Bs = map (fn mk => mk deads As Bs) mk_Imaps_Ds;
       
  1509         fun mk_Isetss As = map2 (fn mk => fn Isets => map (mk deads As) Isets) mk_It_Ds Isetss;
       
  1510         val Ibds = map2 (fn mk => mk deads passiveAs) mk_It_Ds Ibds_Ds;
       
  1511         val Iwitss =
       
  1512           map2 (fn mk => fn Iwits => map (mk deads passiveAs o snd) Iwits) mk_It_Ds Iwitss_Ds;
       
  1513         fun mk_Irels As Bs = map (fn mk => mk deads As Bs) mk_Irels_Ds;
       
  1514 
       
  1515         val Imaps = mk_Imaps passiveAs passiveBs;
       
  1516         val fs_Imaps = map (fn m => Term.list_comb (m, fs)) Imaps;
       
  1517         val fs_copy_Imaps = map (fn m => Term.list_comb (m, fs_copy)) Imaps;
       
  1518         val (Isetss_by_range, Isetss_by_bnf) = `transpose (mk_Isetss passiveAs);
       
  1519 
       
  1520         val map_setss = map (fn T => map2 (fn Ds =>
       
  1521           mk_map_of_bnf Ds (passiveAs @ Ts) (mk_set_Ts T)) Dss bnfs) passiveAs;
       
  1522 
       
  1523         val timer = time (timer "bnf constants for the new datatypes");
       
  1524 
       
  1525         val (ctor_Imap_thms, ctor_Imap_o_thms) =
       
  1526           let
       
  1527             fun mk_goal fs_map map ctor ctor' = fold_rev Logic.all fs
       
  1528               (mk_Trueprop_eq (HOLogic.mk_comp (fs_map, ctor),
       
  1529                 HOLogic.mk_comp (ctor', Term.list_comb (map, fs @ fs_Imaps))));
       
  1530             val goals = map4 mk_goal fs_Imaps map_FTFT's ctors ctor's;
       
  1531             val maps =
       
  1532               map4 (fn goal => fn foldx => fn map_comp_id => fn map_cong0 =>
       
  1533                 Goal.prove_sorry lthy [] [] goal
       
  1534                   (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Imap_defs THEN
       
  1535                     mk_map_tac m n foldx map_comp_id map_cong0)
       
  1536                 |> Thm.close_derivation)
       
  1537               goals ctor_fold_thms map_comp_id_thms map_cong0s;
       
  1538           in
       
  1539             `(map (fn thm => thm RS @{thm comp_eq_dest})) maps
       
  1540           end;
       
  1541 
       
  1542         val (ctor_Imap_unique_thms, ctor_Imap_unique_thm) =
       
  1543           let
       
  1544             fun mk_prem u map ctor ctor' =
       
  1545               mk_Trueprop_eq (HOLogic.mk_comp (u, ctor),
       
  1546                 HOLogic.mk_comp (ctor', Term.list_comb (map, fs @ us)));
       
  1547             val prems = map4 mk_prem us map_FTFT's ctors ctor's;
       
  1548             val goal =
       
  1549               HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
       
  1550                 (map2 (curry HOLogic.mk_eq) us fs_Imaps));
       
  1551             val unique = Goal.prove_sorry lthy [] []
       
  1552               (fold_rev Logic.all (us @ fs) (Logic.list_implies (prems, goal)))
       
  1553               (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Imap_defs THEN
       
  1554                 mk_ctor_map_unique_tac ctor_fold_unique_thm sym_map_comps ctxt)
       
  1555               |> Thm.close_derivation;
       
  1556           in
       
  1557             `split_conj_thm unique
       
  1558           end;
       
  1559 
       
  1560         val timer = time (timer "map functions for the new datatypes");
       
  1561 
       
  1562         val ctor_Iset_thmss =
       
  1563           let
       
  1564             fun mk_goal sets ctor set col map =
       
  1565               mk_Trueprop_eq (HOLogic.mk_comp (set, ctor),
       
  1566                 HOLogic.mk_comp (col, Term.list_comb (map, passive_ids @ sets)));
       
  1567             val goalss =
       
  1568               map3 (fn sets => map4 (mk_goal sets) ctors sets) Isetss_by_range colss map_setss;
       
  1569             val setss = map (map2 (fn foldx => fn goal =>
       
  1570                 Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} =>
       
  1571                   unfold_thms_tac ctxt Iset_defs THEN mk_set_tac foldx)
       
  1572                 |> Thm.close_derivation)
       
  1573               ctor_fold_thms) goalss;
       
  1574 
       
  1575             fun mk_simp_goal pas_set act_sets sets ctor z set =
       
  1576               Logic.all z (mk_Trueprop_eq (set $ (ctor $ z),
       
  1577                 mk_union (pas_set $ z,
       
  1578                   Library.foldl1 mk_union (map2 (fn X => mk_UNION (X $ z)) act_sets sets))));
       
  1579             val simp_goalss =
       
  1580               map2 (fn i => fn sets =>
       
  1581                 map4 (fn Fsets => mk_simp_goal (nth Fsets (i - 1)) (drop m Fsets) sets)
       
  1582                   FTs_setss ctors xFs sets)
       
  1583                 ls Isetss_by_range;
       
  1584 
       
  1585             val ctor_setss = map3 (fn i => map3 (fn set_nats => fn goal => fn set =>
       
  1586                 Goal.prove_sorry lthy [] [] goal
       
  1587                   (K (mk_ctor_set_tac set (nth set_nats (i - 1)) (drop m set_nats)))
       
  1588                 |> Thm.close_derivation)
       
  1589               set_mapss) ls simp_goalss setss;
       
  1590           in
       
  1591             ctor_setss
       
  1592           end;
       
  1593 
       
  1594         fun mk_set_thms ctor_set = (@{thm xt1(3)} OF [ctor_set, @{thm Un_upper1}]) ::
       
  1595           map (fn i => (@{thm xt1(3)} OF [ctor_set, @{thm Un_upper2}]) RS
       
  1596             (mk_Un_upper n i RS subset_trans) RSN
       
  1597             (2, @{thm UN_upper} RS subset_trans))
       
  1598             (1 upto n);
       
  1599         val set_Iset_thmsss = transpose (map (map mk_set_thms) ctor_Iset_thmss);
       
  1600 
       
  1601         val timer = time (timer "set functions for the new datatypes");
       
  1602 
       
  1603         val cxs = map (SOME o certify lthy) Izs;
       
  1604         val Isetss_by_range' =
       
  1605           map (map (Term.subst_atomic_types (passiveAs ~~ passiveBs))) Isetss_by_range;
       
  1606 
       
  1607         val Iset_Imap0_thmss =
       
  1608           let
       
  1609             fun mk_set_map0 f map z set set' =
       
  1610               HOLogic.mk_eq (mk_image f $ (set $ z), set' $ (map $ z));
       
  1611 
       
  1612             fun mk_cphi f map z set set' = certify lthy
       
  1613               (Term.absfree (dest_Free z) (mk_set_map0 f map z set set'));
       
  1614 
       
  1615             val csetss = map (map (certify lthy)) Isetss_by_range';
       
  1616 
       
  1617             val cphiss = map3 (fn f => fn sets => fn sets' =>
       
  1618               (map4 (mk_cphi f) fs_Imaps Izs sets sets')) fs Isetss_by_range Isetss_by_range';
       
  1619 
       
  1620             val inducts = map (fn cphis =>
       
  1621               Drule.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm) cphiss;
       
  1622 
       
  1623             val goals =
       
  1624               map3 (fn f => fn sets => fn sets' =>
       
  1625                 HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
       
  1626                   (map4 (mk_set_map0 f) fs_Imaps Izs sets sets')))
       
  1627                   fs Isetss_by_range Isetss_by_range';
       
  1628 
       
  1629             fun mk_tac induct = mk_set_nat_tac m (rtac induct) set_mapss ctor_Imap_thms;
       
  1630             val thms =
       
  1631               map5 (fn goal => fn csets => fn ctor_sets => fn induct => fn i =>
       
  1632                 singleton (Proof_Context.export names_lthy lthy)
       
  1633                   (Goal.prove_sorry lthy [] [] goal (mk_tac induct csets ctor_sets i))
       
  1634                 |> Thm.close_derivation)
       
  1635               goals csetss ctor_Iset_thmss inducts ls;
       
  1636           in
       
  1637             map split_conj_thm thms
       
  1638           end;
       
  1639 
       
  1640         val Iset_bd_thmss =
       
  1641           let
       
  1642             fun mk_set_bd z bd set = mk_ordLeq (mk_card_of (set $ z)) bd;
       
  1643 
       
  1644             fun mk_cphi z set = certify lthy (Term.absfree (dest_Free z) (mk_set_bd z sum_bd set));
       
  1645 
       
  1646             val cphiss = map (map2 mk_cphi Izs) Isetss_by_range;
       
  1647 
       
  1648             val inducts = map (fn cphis =>
       
  1649               Drule.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm) cphiss;
       
  1650 
       
  1651             val goals =
       
  1652               map (fn sets =>
       
  1653                 HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
       
  1654                   (map3 mk_set_bd Izs Ibds sets))) Isetss_by_range;
       
  1655 
       
  1656             fun mk_tac induct = mk_set_bd_tac m (rtac induct) sum_Cinfinite set_bd_sumss;
       
  1657             val thms =
       
  1658               map4 (fn goal => fn ctor_sets => fn induct => fn i =>
       
  1659                 singleton (Proof_Context.export names_lthy lthy)
       
  1660                   (Goal.prove_sorry lthy [] [] goal
       
  1661                     (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Ibd_defs THEN
       
  1662                       mk_tac induct ctor_sets i ctxt))
       
  1663                 |> Thm.close_derivation)
       
  1664               goals ctor_Iset_thmss inducts ls;
       
  1665           in
       
  1666             map split_conj_thm thms
       
  1667           end;
       
  1668 
       
  1669         val Imap_cong0_thms =
       
  1670           let
       
  1671             fun mk_prem z set f g y y' =
       
  1672               mk_Ball (set $ z) (Term.absfree y' (HOLogic.mk_eq (f $ y, g $ y)));
       
  1673 
       
  1674             fun mk_map_cong0 sets z fmap gmap =
       
  1675               HOLogic.mk_imp
       
  1676                 (Library.foldr1 HOLogic.mk_conj (map5 (mk_prem z) sets fs fs_copy ys ys'),
       
  1677                 HOLogic.mk_eq (fmap $ z, gmap $ z));
       
  1678 
       
  1679             fun mk_cphi sets z fmap gmap =
       
  1680               certify lthy (Term.absfree (dest_Free z) (mk_map_cong0 sets z fmap gmap));
       
  1681 
       
  1682             val cphis = map4 mk_cphi Isetss_by_bnf Izs fs_Imaps fs_copy_Imaps;
       
  1683 
       
  1684             val induct = Drule.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm;
       
  1685 
       
  1686             val goal =
       
  1687               HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
       
  1688                 (map4 mk_map_cong0 Isetss_by_bnf Izs fs_Imaps fs_copy_Imaps));
       
  1689 
       
  1690             val thm = singleton (Proof_Context.export names_lthy lthy)
       
  1691               (Goal.prove_sorry lthy [] [] goal
       
  1692               (mk_mcong_tac (rtac induct) set_Iset_thmsss map_cong0s ctor_Imap_thms))
       
  1693               |> Thm.close_derivation;
       
  1694           in
       
  1695             split_conj_thm thm
       
  1696           end;
       
  1697 
       
  1698         val in_rels = map in_rel_of_bnf bnfs;
       
  1699         val in_Irels = map (fn def => trans OF [def, @{thm OO_Grp_alt}] RS @{thm predicate2_eqD})
       
  1700             Irel_unabs_defs;
       
  1701 
       
  1702         val ctor_Iset_incl_thmss = map (map hd) set_Iset_thmsss;
       
  1703         val ctor_set_Iset_incl_thmsss = map (transpose o map tl) set_Iset_thmsss;
       
  1704         val ctor_Iset_thmss' = transpose ctor_Iset_thmss;
       
  1705 
       
  1706         val Irels = mk_Irels passiveAs passiveBs;
       
  1707         val Irelphis = map (fn rel => Term.list_comb (rel, Iphis)) Irels;
       
  1708         val relphis = map (fn rel => Term.list_comb (rel, Iphis @ Irelphis)) rels;
       
  1709         val Irelpsi1s = map (fn rel => Term.list_comb (rel, Ipsi1s)) (mk_Irels passiveAs passiveCs);
       
  1710         val Irelpsi2s = map (fn rel => Term.list_comb (rel, Ipsi2s)) (mk_Irels passiveCs passiveBs);
       
  1711         val Irelpsi12s = map (fn rel =>
       
  1712             Term.list_comb (rel, map2 (curry mk_rel_compp) Ipsi1s Ipsi2s)) Irels;
       
  1713 
       
  1714         val ctor_Irel_thms =
       
  1715           let
       
  1716             fun mk_goal xF yF ctor ctor' Irelphi relphi = fold_rev Logic.all (xF :: yF :: Iphis)
       
  1717               (mk_Trueprop_eq (Irelphi $ (ctor $ xF) $ (ctor' $ yF), relphi $ xF $ yF));
       
  1718             val goals = map6 mk_goal xFs yFs ctors ctor's Irelphis relphis;
       
  1719           in
       
  1720             map12 (fn i => fn goal => fn in_rel => fn map_comp0 => fn map_cong0 =>
       
  1721               fn ctor_map => fn ctor_sets => fn ctor_inject => fn ctor_dtor =>
       
  1722               fn set_map0s => fn ctor_set_incls => fn ctor_set_set_inclss =>
       
  1723               Goal.prove_sorry lthy [] [] goal
       
  1724                (K (mk_ctor_rel_tac lthy in_Irels i in_rel map_comp0 map_cong0 ctor_map ctor_sets
       
  1725                  ctor_inject ctor_dtor set_map0s ctor_set_incls ctor_set_set_inclss))
       
  1726               |> Thm.close_derivation)
       
  1727             ks goals in_rels map_comps map_cong0s ctor_Imap_thms ctor_Iset_thmss'
       
  1728               ctor_inject_thms ctor_dtor_thms set_mapss ctor_Iset_incl_thmss
       
  1729               ctor_set_Iset_incl_thmsss
       
  1730           end;
       
  1731 
       
  1732         val le_Irel_OO_thm =
       
  1733           let
       
  1734             fun mk_le_Irel_OO Irelpsi1 Irelpsi2 Irelpsi12 Iz1 Iz2 =
       
  1735               HOLogic.mk_imp (mk_rel_compp (Irelpsi1, Irelpsi2) $ Iz1 $ Iz2,
       
  1736                 Irelpsi12 $ Iz1 $ Iz2);
       
  1737             val goals = map5 mk_le_Irel_OO Irelpsi1s Irelpsi2s Irelpsi12s Izs1 Izs2;
       
  1738 
       
  1739             val cTs = map (SOME o certifyT lthy o TFree) induct2_params;
       
  1740             val cxs = map (SOME o certify lthy) (splice Izs1 Izs2);
       
  1741             fun mk_cphi z1 z2 goal = SOME (certify lthy (Term.absfree z1 (Term.absfree z2 goal)));
       
  1742             val cphis = map3 mk_cphi Izs1' Izs2' goals;
       
  1743             val induct = Drule.instantiate' cTs (cphis @ cxs) ctor_induct2_thm;
       
  1744 
       
  1745             val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj goals);
       
  1746           in
       
  1747             singleton (Proof_Context.export names_lthy lthy)
       
  1748               (Goal.prove_sorry lthy [] [] goal
       
  1749                 (mk_le_rel_OO_tac m induct ctor_nchotomy_thms ctor_Irel_thms rel_mono_strongs
       
  1750                   rel_OOs))
       
  1751               |> Thm.close_derivation
       
  1752           end;
       
  1753 
       
  1754         val timer = time (timer "helpers for BNF properties");
       
  1755 
       
  1756         val map_id0_tacs = map (K o mk_map_id0_tac map_id0s) ctor_Imap_unique_thms;
       
  1757         val map_comp0_tacs =
       
  1758           map2 (K oo mk_map_comp0_tac map_comps ctor_Imap_thms) ctor_Imap_unique_thms ks;
       
  1759         val map_cong0_tacs = map (mk_map_cong0_tac m) Imap_cong0_thms;
       
  1760         val set_map0_tacss = map (map (K o mk_set_map0_tac)) (transpose Iset_Imap0_thmss);
       
  1761         val bd_co_tacs = replicate n (fn {context = ctxt, prems = _} =>
       
  1762           unfold_thms_tac ctxt Ibd_defs THEN mk_bd_card_order_tac bd_card_orders);
       
  1763         val bd_cinf_tacs = replicate n (fn {context = ctxt, prems = _} =>
       
  1764           unfold_thms_tac ctxt Ibd_defs THEN rtac (sum_Cinfinite RS conjunct1) 1);
       
  1765         val set_bd_tacss = map (map (fn thm => K (rtac thm 1))) (transpose Iset_bd_thmss);
       
  1766         val le_rel_OO_tacs = map (fn i =>
       
  1767           K ((rtac @{thm predicate2I} THEN' etac (le_Irel_OO_thm RS mk_conjunctN n i RS mp)) 1)) ks;
       
  1768 
       
  1769         val rel_OO_Grp_tacs = map (fn def => K (rtac def 1)) Irel_unabs_defs;
       
  1770 
       
  1771         val tacss = map9 zip_axioms map_id0_tacs map_comp0_tacs map_cong0_tacs set_map0_tacss
       
  1772           bd_co_tacs bd_cinf_tacs set_bd_tacss le_rel_OO_tacs rel_OO_Grp_tacs;
       
  1773 
       
  1774         fun wit_tac {context = ctxt, prems = _} = unfold_thms_tac ctxt (flat Iwit_defss) THEN
       
  1775           mk_wit_tac ctxt n (flat ctor_Iset_thmss) (maps wit_thms_of_bnf bnfs);
       
  1776 
       
  1777         val (Ibnfs, lthy) =
       
  1778           fold_map5 (fn tacs => fn map_b => fn rel_b => fn set_bs => fn consts => fn lthy =>
       
  1779             bnf_def Do_Inline (user_policy Note_Some) I tacs wit_tac (SOME deads)
       
  1780               map_b rel_b set_bs consts lthy
       
  1781             |> register_bnf (Local_Theory.full_name lthy b))
       
  1782           tacss map_bs rel_bs set_bss
       
  1783             ((((((bs ~~ Ts) ~~ Imaps) ~~ Isetss_by_bnf) ~~ Ibds) ~~ Iwitss) ~~ map SOME Irels)
       
  1784             lthy;
       
  1785 
       
  1786         val timer = time (timer "registered new datatypes as BNFs");
       
  1787 
       
  1788         val ls' = if m = 1 then [0] else ls
       
  1789 
       
  1790         val Ibnf_common_notes =
       
  1791           [(ctor_map_uniqueN, [ctor_Imap_unique_thm])]
       
  1792           |> map (fn (thmN, thms) =>
       
  1793             ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]));
       
  1794 
       
  1795         val Ibnf_notes =
       
  1796           [(ctor_mapN, map single ctor_Imap_thms),
       
  1797           (ctor_relN, map single ctor_Irel_thms),
       
  1798           (ctor_set_inclN, ctor_Iset_incl_thmss),
       
  1799           (ctor_set_set_inclN, map flat ctor_set_Iset_incl_thmsss)] @
       
  1800           map2 (fn i => fn thms => (mk_ctor_setN i, map single thms)) ls' ctor_Iset_thmss
       
  1801           |> maps (fn (thmN, thmss) =>
       
  1802             map2 (fn b => fn thms =>
       
  1803               ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]))
       
  1804             bs thmss)
       
  1805       in
       
  1806         (timer, Ibnfs, (ctor_Imap_o_thms, ctor_Imap_thms), ctor_Iset_thmss',
       
  1807           ctor_Irel_thms, Ibnf_common_notes @ Ibnf_notes, lthy)
       
  1808       end;
       
  1809 
       
  1810       val ctor_fold_o_Imap_thms = mk_xtor_un_fold_o_map_thms Least_FP false m ctor_fold_unique_thm
       
  1811         ctor_Imap_o_thms (map (mk_pointfree lthy) ctor_fold_thms) sym_map_comps map_cong0s;
       
  1812       val ctor_rec_o_Imap_thms = mk_xtor_un_fold_o_map_thms Least_FP true m ctor_rec_unique_thm
       
  1813         ctor_Imap_o_thms (map (mk_pointfree lthy) ctor_rec_thms) sym_map_comps map_cong0s;
       
  1814 
       
  1815       val Irels = if m = 0 then map HOLogic.eq_const Ts
       
  1816         else map (mk_rel_of_bnf deads passiveAs passiveBs) Ibnfs;
       
  1817       val Irel_induct_thm =
       
  1818         mk_rel_xtor_co_induct_thm Least_FP rels activeIphis Irels Iphis xFs yFs ctors ctor's
       
  1819           (mk_rel_induct_tac m ctor_induct2_thm ks ctor_Irel_thms rel_mono_strongs) lthy;
       
  1820 
       
  1821       val rels = map2 (fn Ds => mk_rel_of_bnf Ds allAs allBs') Dss bnfs;
       
  1822       val ctor_fold_transfer_thms =
       
  1823         mk_un_fold_transfer_thms Least_FP rels activephis Irels Iphis
       
  1824           (mk_folds passiveAs activeAs) (mk_folds passiveBs activeBs)
       
  1825           (mk_fold_transfer_tac m Irel_induct_thm (map map_transfer_of_bnf bnfs) ctor_fold_thms)
       
  1826           lthy;
       
  1827 
       
  1828       val timer = time (timer "relator induction");
       
  1829 
       
  1830       val common_notes =
       
  1831         [(ctor_inductN, [ctor_induct_thm]),
       
  1832         (ctor_induct2N, [ctor_induct2_thm]),
       
  1833         (rel_inductN, [Irel_induct_thm]),
       
  1834         (ctor_fold_transferN, ctor_fold_transfer_thms)]
       
  1835         |> map (fn (thmN, thms) =>
       
  1836           ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]));
       
  1837 
       
  1838       val notes =
       
  1839         [(ctor_dtorN, ctor_dtor_thms),
       
  1840         (ctor_exhaustN, ctor_exhaust_thms),
       
  1841         (ctor_foldN, ctor_fold_thms),
       
  1842         (ctor_fold_uniqueN, ctor_fold_unique_thms),
       
  1843         (ctor_rec_uniqueN, ctor_rec_unique_thms),
       
  1844         (ctor_fold_o_mapN, ctor_fold_o_Imap_thms),
       
  1845         (ctor_rec_o_mapN, ctor_rec_o_Imap_thms),
       
  1846         (ctor_injectN, ctor_inject_thms),
       
  1847         (ctor_recN, ctor_rec_thms),
       
  1848         (dtor_ctorN, dtor_ctor_thms),
       
  1849         (dtor_exhaustN, dtor_exhaust_thms),
       
  1850         (dtor_injectN, dtor_inject_thms)]
       
  1851         |> map (apsnd (map single))
       
  1852         |> maps (fn (thmN, thmss) =>
       
  1853           map2 (fn b => fn thms =>
       
  1854             ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]))
       
  1855           bs thmss);
       
  1856 
       
  1857     (*FIXME: once the package exports all the necessary high-level characteristic theorems,
       
  1858        those should not only be concealed but rather not noted at all*)
       
  1859     val maybe_conceal_notes = note_all = false ? map (apfst (apfst Binding.conceal));
       
  1860   in
       
  1861     timer;
       
  1862     ({Ts = Ts, bnfs = Ibnfs, ctors = ctors, dtors = dtors, xtor_co_iterss = transpose [folds, recs],
       
  1863       xtor_co_induct = ctor_induct_thm, dtor_ctors = dtor_ctor_thms, ctor_dtors = ctor_dtor_thms,
       
  1864       ctor_injects = ctor_inject_thms, dtor_injects = dtor_inject_thms,
       
  1865       xtor_map_thms = ctor_Imap_thms, xtor_set_thmss = ctor_Iset_thmss',
       
  1866       xtor_rel_thms = ctor_Irel_thms,
       
  1867       xtor_co_iter_thmss = transpose [ctor_fold_thms, ctor_rec_thms],
       
  1868       xtor_co_iter_o_map_thmss = transpose [ctor_fold_o_Imap_thms, ctor_rec_o_Imap_thms],
       
  1869       rel_xtor_co_induct_thm = Irel_induct_thm},
       
  1870      lthy |> Local_Theory.notes (maybe_conceal_notes (common_notes @ notes @ Ibnf_notes)) |> snd)
       
  1871   end;
       
  1872 
       
  1873 val _ =
       
  1874   Outer_Syntax.local_theory @{command_spec "datatype_new"} "define new-style inductive datatypes"
       
  1875     (parse_co_datatype_cmd Least_FP construct_lfp);
       
  1876 
       
  1877 val _ = Outer_Syntax.local_theory @{command_spec "primrec_new"}
       
  1878   "define primitive recursive functions"
       
  1879   (Parse.fixes -- Parse_Spec.where_alt_specs >> (snd oo uncurry add_primrec_cmd));
       
  1880 
       
  1881 end;