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