avoid nested 'Tools' directories
authorblanchet
Mon, 20 Jan 2014 18:24:56 +0100
changeset 55060 3105434fb02f
parent 55059 ef2e0fb783c6
child 55061 a0adf838e2d1
avoid nested 'Tools' directories
src/HOL/Tools/BNF/Tools/bnf_comp.ML
src/HOL/Tools/BNF/Tools/bnf_comp_tactics.ML
src/HOL/Tools/BNF/Tools/bnf_decl.ML
src/HOL/Tools/BNF/Tools/bnf_def.ML
src/HOL/Tools/BNF/Tools/bnf_def_tactics.ML
src/HOL/Tools/BNF/Tools/bnf_fp_def_sugar.ML
src/HOL/Tools/BNF/Tools/bnf_fp_def_sugar_tactics.ML
src/HOL/Tools/BNF/Tools/bnf_fp_n2m.ML
src/HOL/Tools/BNF/Tools/bnf_fp_n2m_sugar.ML
src/HOL/Tools/BNF/Tools/bnf_fp_n2m_tactics.ML
src/HOL/Tools/BNF/Tools/bnf_fp_rec_sugar_util.ML
src/HOL/Tools/BNF/Tools/bnf_fp_util.ML
src/HOL/Tools/BNF/Tools/bnf_gfp.ML
src/HOL/Tools/BNF/Tools/bnf_gfp_rec_sugar.ML
src/HOL/Tools/BNF/Tools/bnf_gfp_rec_sugar_tactics.ML
src/HOL/Tools/BNF/Tools/bnf_gfp_tactics.ML
src/HOL/Tools/BNF/Tools/bnf_gfp_util.ML
src/HOL/Tools/BNF/Tools/bnf_lfp.ML
src/HOL/Tools/BNF/Tools/bnf_lfp_compat.ML
src/HOL/Tools/BNF/Tools/bnf_lfp_rec_sugar.ML
src/HOL/Tools/BNF/Tools/bnf_lfp_tactics.ML
src/HOL/Tools/BNF/Tools/bnf_lfp_util.ML
src/HOL/Tools/BNF/Tools/bnf_tactics.ML
src/HOL/Tools/BNF/Tools/bnf_util.ML
src/HOL/Tools/BNF/bnf_comp.ML
src/HOL/Tools/BNF/bnf_comp_tactics.ML
src/HOL/Tools/BNF/bnf_def.ML
src/HOL/Tools/BNF/bnf_def_tactics.ML
src/HOL/Tools/BNF/bnf_fp_def_sugar.ML
src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML
src/HOL/Tools/BNF/bnf_fp_n2m.ML
src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML
src/HOL/Tools/BNF/bnf_fp_n2m_tactics.ML
src/HOL/Tools/BNF/bnf_fp_rec_sugar_util.ML
src/HOL/Tools/BNF/bnf_fp_util.ML
src/HOL/Tools/BNF/bnf_gfp.ML
src/HOL/Tools/BNF/bnf_gfp_rec_sugar.ML
src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML
src/HOL/Tools/BNF/bnf_gfp_tactics.ML
src/HOL/Tools/BNF/bnf_gfp_util.ML
src/HOL/Tools/BNF/bnf_lfp.ML
src/HOL/Tools/BNF/bnf_lfp_compat.ML
src/HOL/Tools/BNF/bnf_lfp_rec_sugar.ML
src/HOL/Tools/BNF/bnf_lfp_tactics.ML
src/HOL/Tools/BNF/bnf_lfp_util.ML
src/HOL/Tools/BNF/bnf_tactics.ML
src/HOL/Tools/BNF/bnf_util.ML
--- a/src/HOL/Tools/BNF/Tools/bnf_comp.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,704 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_comp.ML
-    Author:     Dmitriy Traytel, TU Muenchen
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2012
-
-Composition of bounded natural functors.
-*)
-
-signature BNF_COMP =
-sig
-  val ID_bnf: BNF_Def.bnf
-  val DEADID_bnf: BNF_Def.bnf
-
-  type unfold_set
-  val empty_unfolds: unfold_set
-
-  exception BAD_DEAD of typ * typ
-
-  val bnf_of_typ: BNF_Def.const_policy -> (binding -> binding) ->
-    ((string * sort) list list -> (string * sort) list) -> (string * sort) list -> typ ->
-    unfold_set * Proof.context ->
-    (BNF_Def.bnf * (typ list * typ list)) * (unfold_set * Proof.context)
-  val default_comp_sort: (string * sort) list list -> (string * sort) list
-  val normalize_bnfs: (int -> binding -> binding) -> ''a list list -> ''a list ->
-    (''a list list -> ''a list) -> BNF_Def.bnf list -> unfold_set -> Proof.context ->
-    (int list list * ''a list) * (BNF_Def.bnf list * (unfold_set * Proof.context))
-  val seal_bnf: (binding -> binding) -> unfold_set -> binding -> typ list -> BNF_Def.bnf ->
-    Proof.context -> (BNF_Def.bnf * typ list) * local_theory
-end;
-
-structure BNF_Comp : BNF_COMP =
-struct
-
-open BNF_Def
-open BNF_Util
-open BNF_Tactics
-open BNF_Comp_Tactics
-
-val ID_bnf = the (bnf_of @{context} "Basic_BNFs.ID");
-val DEADID_bnf = the (bnf_of @{context} "Basic_BNFs.DEADID");
-
-(* TODO: Replace by "BNF_Defs.defs list" *)
-type unfold_set = {
-  map_unfolds: thm list,
-  set_unfoldss: thm list list,
-  rel_unfolds: thm list
-};
-
-val empty_unfolds = {map_unfolds = [], set_unfoldss = [], rel_unfolds = []};
-
-fun add_to_thms thms new = thms |> not (Thm.is_reflexive new) ? insert Thm.eq_thm new;
-fun adds_to_thms thms news = insert (eq_set Thm.eq_thm) (no_reflexive news) thms;
-
-fun add_to_unfolds map sets rel
-  {map_unfolds, set_unfoldss, rel_unfolds} =
-  {map_unfolds = add_to_thms map_unfolds map,
-    set_unfoldss = adds_to_thms set_unfoldss sets,
-    rel_unfolds = add_to_thms rel_unfolds rel};
-
-fun add_bnf_to_unfolds bnf =
-  add_to_unfolds (map_def_of_bnf bnf) (set_defs_of_bnf bnf) (rel_def_of_bnf bnf);
-
-val bdTN = "bdT";
-
-fun mk_killN n = "_kill" ^ string_of_int n;
-fun mk_liftN n = "_lift" ^ string_of_int n;
-fun mk_permuteN src dest =
-  "_permute_" ^ implode (map string_of_int src) ^ "_" ^ implode (map string_of_int dest);
-
-(*copied from Envir.expand_term_free*)
-fun expand_term_const defs =
-  let
-    val eqs = map ((fn ((x, U), u) => (x, (U, u))) o apfst dest_Const) defs;
-    val get = fn Const (x, _) => AList.lookup (op =) eqs x | _ => NONE;
-  in Envir.expand_term get end;
-
-fun clean_compose_bnf const_policy qualify b outer inners (unfold_set, lthy) =
-  let
-    val olive = live_of_bnf outer;
-    val onwits = nwits_of_bnf outer;
-    val odead = dead_of_bnf outer;
-    val inner = hd inners;
-    val ilive = live_of_bnf inner;
-    val ideads = map dead_of_bnf inners;
-    val inwitss = map nwits_of_bnf inners;
-
-    (* TODO: check olive = length inners > 0,
-                   forall inner from inners. ilive = live,
-                   forall inner from inners. idead = dead  *)
-
-    val (oDs, lthy1) = apfst (map TFree)
-      (Variable.invent_types (replicate odead HOLogic.typeS) lthy);
-    val (Dss, lthy2) = apfst (map (map TFree))
-        (fold_map Variable.invent_types (map (fn n => replicate n HOLogic.typeS) ideads) lthy1);
-    val (Ass, lthy3) = apfst (replicate ilive o map TFree)
-      (Variable.invent_types (replicate ilive HOLogic.typeS) lthy2);
-    val As = if ilive > 0 then hd Ass else [];
-    val Ass_repl = replicate olive As;
-    val (Bs, _(*lthy4*)) = apfst (map TFree)
-      (Variable.invent_types (replicate ilive HOLogic.typeS) lthy3);
-    val Bss_repl = replicate olive Bs;
-
-    val ((((fs', Qs'), Asets), xs), _(*names_lthy*)) = lthy
-      |> apfst snd o mk_Frees' "f" (map2 (curry op -->) As Bs)
-      ||>> apfst snd o mk_Frees' "Q" (map2 mk_pred2T As Bs)
-      ||>> mk_Frees "A" (map HOLogic.mk_setT As)
-      ||>> mk_Frees "x" As;
-
-    val CAs = map3 mk_T_of_bnf Dss Ass_repl inners;
-    val CCA = mk_T_of_bnf oDs CAs outer;
-    val CBs = map3 mk_T_of_bnf Dss Bss_repl inners;
-    val outer_sets = mk_sets_of_bnf (replicate olive oDs) (replicate olive CAs) outer;
-    val inner_setss = map3 mk_sets_of_bnf (map (replicate ilive) Dss) (replicate olive Ass) inners;
-    val inner_bds = map3 mk_bd_of_bnf Dss Ass_repl inners;
-    val outer_bd = mk_bd_of_bnf oDs CAs outer;
-
-    (*%f1 ... fn. outer.map (inner_1.map f1 ... fn) ... (inner_m.map f1 ... fn)*)
-    val mapx = fold_rev Term.abs fs'
-      (Term.list_comb (mk_map_of_bnf oDs CAs CBs outer,
-        map2 (fn Ds => (fn f => Term.list_comb (f, map Bound (ilive - 1 downto 0))) o
-          mk_map_of_bnf Ds As Bs) Dss inners));
-    (*%Q1 ... Qn. outer.rel (inner_1.rel Q1 ... Qn) ... (inner_m.rel Q1 ... Qn)*)
-    val rel = fold_rev Term.abs Qs'
-      (Term.list_comb (mk_rel_of_bnf oDs CAs CBs outer,
-        map2 (fn Ds => (fn f => Term.list_comb (f, map Bound (ilive - 1 downto 0))) o
-          mk_rel_of_bnf Ds As Bs) Dss inners));
-
-    (*Union o collect {outer.set_1 ... outer.set_m} o outer.map inner_1.set_i ... inner_m.set_i*)
-    (*Union o collect {image inner_1.set_i o outer.set_1 ... image inner_m.set_i o outer.set_m}*)
-    fun mk_set i =
-      let
-        val (setTs, T) = `(replicate olive o HOLogic.mk_setT) (nth As i);
-        val outer_set = mk_collect
-          (mk_sets_of_bnf (replicate olive oDs) (replicate olive setTs) outer)
-          (mk_T_of_bnf oDs setTs outer --> HOLogic.mk_setT T);
-        val inner_sets = map (fn sets => nth sets i) inner_setss;
-        val outer_map = mk_map_of_bnf oDs CAs setTs outer;
-        val map_inner_sets = Term.list_comb (outer_map, inner_sets);
-        val collect_image = mk_collect
-          (map2 (fn f => fn set => HOLogic.mk_comp (mk_image f, set)) inner_sets outer_sets)
-          (CCA --> HOLogic.mk_setT T);
-      in
-        (Library.foldl1 HOLogic.mk_comp [mk_Union T, outer_set, map_inner_sets],
-        HOLogic.mk_comp (mk_Union T, collect_image))
-      end;
-
-    val (sets, sets_alt) = map_split mk_set (0 upto ilive - 1);
-
-    (*(inner_1.bd +c ... +c inner_m.bd) *c outer.bd*)
-    val bd = mk_cprod (Library.foldr1 (uncurry mk_csum) inner_bds) outer_bd;
-
-    fun map_id0_tac _ =
-      mk_comp_map_id0_tac (map_id0_of_bnf outer) (map_cong0_of_bnf outer)
-        (map map_id0_of_bnf inners);
-
-    fun map_comp0_tac _ =
-      mk_comp_map_comp0_tac (map_comp0_of_bnf outer) (map_cong0_of_bnf outer)
-        (map map_comp0_of_bnf inners);
-
-    fun mk_single_set_map0_tac i _ =
-      mk_comp_set_map0_tac (map_comp0_of_bnf outer) (map_cong0_of_bnf outer)
-        (collect_set_map_of_bnf outer)
-        (map ((fn thms => nth thms i) o set_map0_of_bnf) inners);
-
-    val set_map0_tacs = map mk_single_set_map0_tac (0 upto ilive - 1);
-
-    fun bd_card_order_tac _ =
-      mk_comp_bd_card_order_tac (map bd_card_order_of_bnf inners) (bd_card_order_of_bnf outer);
-
-    fun bd_cinfinite_tac _ =
-      mk_comp_bd_cinfinite_tac (bd_cinfinite_of_bnf inner) (bd_cinfinite_of_bnf outer);
-
-    val set_alt_thms =
-      if Config.get lthy quick_and_dirty then
-        []
-      else
-        map (fn goal =>
-          Goal.prove_sorry lthy [] [] goal
-            (fn {context = ctxt, prems = _} =>
-              mk_comp_set_alt_tac ctxt (collect_set_map_of_bnf outer))
-          |> Thm.close_derivation)
-        (map2 (curry (HOLogic.mk_Trueprop o HOLogic.mk_eq)) sets sets_alt);
-
-    fun map_cong0_tac _ =
-      mk_comp_map_cong0_tac set_alt_thms (map_cong0_of_bnf outer) (map map_cong0_of_bnf inners);
-
-    val set_bd_tacs =
-      if Config.get lthy quick_and_dirty then
-        replicate ilive (K all_tac)
-      else
-        let
-          val outer_set_bds = set_bd_of_bnf outer;
-          val inner_set_bdss = map set_bd_of_bnf inners;
-          val inner_bd_Card_orders = map bd_Card_order_of_bnf inners;
-          fun single_set_bd_thm i j =
-            @{thm comp_single_set_bd} OF [nth inner_bd_Card_orders j, nth (nth inner_set_bdss j) i,
-              nth outer_set_bds j]
-          val single_set_bd_thmss =
-            map ((fn f => map f (0 upto olive - 1)) o single_set_bd_thm) (0 upto ilive - 1);
-        in
-          map2 (fn set_alt => fn single_set_bds => fn {context = ctxt, prems = _} =>
-            mk_comp_set_bd_tac ctxt set_alt single_set_bds)
-          set_alt_thms single_set_bd_thmss
-        end;
-
-    val in_alt_thm =
-      let
-        val inx = mk_in Asets sets CCA;
-        val in_alt = mk_in (map2 (mk_in Asets) inner_setss CAs) outer_sets CCA;
-        val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));
-      in
-        Goal.prove_sorry lthy [] [] goal
-          (fn {context = ctxt, prems = _} => mk_comp_in_alt_tac ctxt set_alt_thms)
-        |> Thm.close_derivation
-      end;
-
-    fun le_rel_OO_tac _ = mk_le_rel_OO_tac (le_rel_OO_of_bnf outer) (rel_mono_of_bnf outer)
-      (map le_rel_OO_of_bnf inners);
-
-    fun rel_OO_Grp_tac _ =
-      let
-        val outer_rel_Grp = rel_Grp_of_bnf outer RS sym;
-        val outer_rel_cong = rel_cong_of_bnf outer;
-        val thm =
-          (trans OF [in_alt_thm RS @{thm OO_Grp_cong},
-             trans OF [@{thm arg_cong2[of _ _ _ _ relcompp]} OF
-               [trans OF [outer_rel_Grp RS @{thm arg_cong[of _ _ conversep]},
-                 rel_conversep_of_bnf outer RS sym], outer_rel_Grp],
-               trans OF [rel_OO_of_bnf outer RS sym, outer_rel_cong OF
-                 (map (fn bnf => rel_OO_Grp_of_bnf bnf RS sym) inners)]]] RS sym)
-          (*|> unfold_thms lthy (rel_def_of_bnf outer :: map rel_def_of_bnf inners)*);
-      in
-        rtac thm 1
-      end;
-
-    val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac
-      bd_cinfinite_tac set_bd_tacs le_rel_OO_tac rel_OO_Grp_tac;
-
-    val outer_wits = mk_wits_of_bnf (replicate onwits oDs) (replicate onwits CAs) outer;
-
-    val inner_witss = map (map (fn (I, wit) => Term.list_comb (wit, map (nth xs) I)))
-      (map3 (fn Ds => fn n => mk_wits_of_bnf (replicate n Ds) (replicate n As))
-        Dss inwitss inners);
-
-    val inner_witsss = map (map (nth inner_witss) o fst) outer_wits;
-
-    val wits = (inner_witsss, (map (single o snd) outer_wits))
-      |-> map2 (fold (map_product (fn iwit => fn owit => owit $ iwit)))
-      |> flat
-      |> map (`(fn t => Term.add_frees t []))
-      |> minimize_wits
-      |> map (fn (frees, t) => fold absfree frees t);
-
-    fun wit_tac {context = ctxt, prems = _} =
-      mk_comp_wit_tac ctxt (wit_thms_of_bnf outer) (collect_set_map_of_bnf outer)
-        (maps wit_thms_of_bnf inners);
-
-    val (bnf', lthy') =
-      bnf_def const_policy (K Dont_Note) qualify tacs wit_tac (SOME (oDs @ flat Dss)) Binding.empty
-        Binding.empty [] ((((((b, CCA), mapx), sets), bd), wits), SOME rel) lthy;
-  in
-    (bnf', (add_bnf_to_unfolds bnf' unfold_set, lthy'))
-  end;
-
-(* Killing live variables *)
-
-fun kill_bnf qualify n bnf (unfold_set, lthy) = if n = 0 then (bnf, (unfold_set, lthy)) else
-  let
-    val b = Binding.suffix_name (mk_killN n) (name_of_bnf bnf);
-    val live = live_of_bnf bnf;
-    val dead = dead_of_bnf bnf;
-    val nwits = nwits_of_bnf bnf;
-
-    (* TODO: check 0 < n <= live *)
-
-    val (Ds, lthy1) = apfst (map TFree)
-      (Variable.invent_types (replicate dead HOLogic.typeS) lthy);
-    val ((killedAs, As), lthy2) = apfst (`(take n) o map TFree)
-      (Variable.invent_types (replicate live HOLogic.typeS) lthy1);
-    val (Bs, _(*lthy3*)) = apfst (append killedAs o map TFree)
-      (Variable.invent_types (replicate (live - n) HOLogic.typeS) lthy2);
-
-    val ((Asets, lives), _(*names_lthy*)) = lthy
-      |> mk_Frees "A" (map HOLogic.mk_setT (drop n As))
-      ||>> mk_Frees "x" (drop n As);
-    val xs = map (fn T => HOLogic.choice_const T $ absdummy T @{term True}) killedAs @ lives;
-
-    val T = mk_T_of_bnf Ds As bnf;
-
-    (*bnf.map id ... id*)
-    val mapx = Term.list_comb (mk_map_of_bnf Ds As Bs bnf, map HOLogic.id_const killedAs);
-    (*bnf.rel (op =) ... (op =)*)
-    val rel = Term.list_comb (mk_rel_of_bnf Ds As Bs bnf, map HOLogic.eq_const killedAs);
-
-    val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf;
-    val sets = drop n bnf_sets;
-
-    (*(|UNIV :: A1 set| +c ... +c |UNIV :: An set|) *c bnf.bd*)
-    val bnf_bd = mk_bd_of_bnf Ds As bnf;
-    val bd = mk_cprod
-      (Library.foldr1 (uncurry mk_csum) (map (mk_card_of o HOLogic.mk_UNIV) killedAs)) bnf_bd;
-
-    fun map_id0_tac _ = rtac (map_id0_of_bnf bnf) 1;
-    fun map_comp0_tac {context = ctxt, prems = _} =
-      unfold_thms_tac ctxt ((map_comp0_of_bnf bnf RS sym) :: @{thms o_assoc id_o o_id}) THEN
-      rtac refl 1;
-    fun map_cong0_tac {context = ctxt, prems = _} =
-      mk_kill_map_cong0_tac ctxt n (live - n) (map_cong0_of_bnf bnf);
-    val set_map0_tacs = map (fn thm => fn _ => rtac thm 1) (drop n (set_map0_of_bnf bnf));
-    fun bd_card_order_tac _ = mk_kill_bd_card_order_tac n (bd_card_order_of_bnf bnf);
-    fun bd_cinfinite_tac _ = mk_kill_bd_cinfinite_tac (bd_Cinfinite_of_bnf bnf);
-    val set_bd_tacs =
-      map (fn thm => fn _ => mk_kill_set_bd_tac (bd_Card_order_of_bnf bnf) thm)
-        (drop n (set_bd_of_bnf bnf));
-
-    val in_alt_thm =
-      let
-        val inx = mk_in Asets sets T;
-        val in_alt = mk_in (map HOLogic.mk_UNIV killedAs @ Asets) bnf_sets T;
-        val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));
-      in
-        Goal.prove_sorry lthy [] [] goal (K kill_in_alt_tac) |> Thm.close_derivation
-      end;
-
-    fun le_rel_OO_tac {context = ctxt, prems = _} =
-      EVERY' [rtac @{thm ord_le_eq_trans}, rtac (le_rel_OO_of_bnf bnf)] 1 THEN
-      unfold_thms_tac ctxt @{thms eq_OO} THEN rtac refl 1;
-
-    fun rel_OO_Grp_tac _ =
-      let
-        val rel_Grp = rel_Grp_of_bnf bnf RS sym
-        val thm =
-          (trans OF [in_alt_thm RS @{thm OO_Grp_cong},
-            trans OF [@{thm arg_cong2[of _ _ _ _ relcompp]} OF
-              [trans OF [rel_Grp RS @{thm arg_cong[of _ _ conversep]},
-                rel_conversep_of_bnf bnf RS sym], rel_Grp],
-              trans OF [rel_OO_of_bnf bnf RS sym, rel_cong_of_bnf bnf OF
-                (replicate n @{thm trans[OF Grp_UNIV_id[OF refl] eq_alt[symmetric]]} @
-                 replicate (live - n) @{thm Grp_fst_snd})]]] RS sym);
-      in
-        rtac thm 1
-      end;
-
-    val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac
-      bd_cinfinite_tac set_bd_tacs le_rel_OO_tac rel_OO_Grp_tac;
-
-    val bnf_wits = mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf;
-
-    val wits = map (fn t => fold absfree (Term.add_frees t []) t)
-      (map (fn (I, wit) => Term.list_comb (wit, map (nth xs) I)) bnf_wits);
-
-    fun wit_tac _ = mk_simple_wit_tac (wit_thms_of_bnf bnf);
-
-    val (bnf', lthy') =
-      bnf_def Smart_Inline (K Dont_Note) qualify tacs wit_tac (SOME (killedAs @ Ds)) Binding.empty
-        Binding.empty [] ((((((b, T), mapx), sets), bd), wits), SOME rel) lthy;
-  in
-    (bnf', (add_bnf_to_unfolds bnf' unfold_set, lthy'))
-  end;
-
-(* Adding dummy live variables *)
-
-fun lift_bnf qualify n bnf (unfold_set, lthy) = if n = 0 then (bnf, (unfold_set, lthy)) else
-  let
-    val b = Binding.suffix_name (mk_liftN n) (name_of_bnf bnf);
-    val live = live_of_bnf bnf;
-    val dead = dead_of_bnf bnf;
-    val nwits = nwits_of_bnf bnf;
-
-    (* TODO: check 0 < n *)
-
-    val (Ds, lthy1) = apfst (map TFree)
-      (Variable.invent_types (replicate dead HOLogic.typeS) lthy);
-    val ((newAs, As), lthy2) = apfst (chop n o map TFree)
-      (Variable.invent_types (replicate (n + live) HOLogic.typeS) lthy1);
-    val ((newBs, Bs), _(*lthy3*)) = apfst (chop n o map TFree)
-      (Variable.invent_types (replicate (n + live) HOLogic.typeS) lthy2);
-
-    val (Asets, _(*names_lthy*)) = lthy
-      |> mk_Frees "A" (map HOLogic.mk_setT (newAs @ As));
-
-    val T = mk_T_of_bnf Ds As bnf;
-
-    (*%f1 ... fn. bnf.map*)
-    val mapx =
-      fold_rev Term.absdummy (map2 (curry op -->) newAs newBs) (mk_map_of_bnf Ds As Bs bnf);
-    (*%Q1 ... Qn. bnf.rel*)
-    val rel = fold_rev Term.absdummy (map2 mk_pred2T newAs newBs) (mk_rel_of_bnf Ds As Bs bnf);
-
-    val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf;
-    val sets = map (fn A => absdummy T (HOLogic.mk_set A [])) newAs @ bnf_sets;
-
-    val bd = mk_bd_of_bnf Ds As bnf;
-
-    fun map_id0_tac _ = rtac (map_id0_of_bnf bnf) 1;
-    fun map_comp0_tac {context = ctxt, prems = _} =
-      unfold_thms_tac ctxt ((map_comp0_of_bnf bnf RS sym) :: @{thms o_assoc id_o o_id}) THEN
-      rtac refl 1;
-    fun map_cong0_tac {context = ctxt, prems = _} =
-      rtac (map_cong0_of_bnf bnf) 1 THEN REPEAT_DETERM_N live (Goal.assume_rule_tac ctxt 1);
-    val set_map0_tacs =
-      if Config.get lthy quick_and_dirty then
-        replicate (n + live) (K all_tac)
-      else
-        replicate n (K empty_natural_tac) @
-        map (fn thm => fn _ => rtac thm 1) (set_map0_of_bnf bnf);
-    fun bd_card_order_tac _ = rtac (bd_card_order_of_bnf bnf) 1;
-    fun bd_cinfinite_tac _ = rtac (bd_cinfinite_of_bnf bnf) 1;
-    val set_bd_tacs =
-      if Config.get lthy quick_and_dirty then
-        replicate (n + live) (K all_tac)
-      else
-        replicate n (K (mk_lift_set_bd_tac (bd_Card_order_of_bnf bnf))) @
-        (map (fn thm => fn _ => rtac thm 1) (set_bd_of_bnf bnf));
-
-    val in_alt_thm =
-      let
-        val inx = mk_in Asets sets T;
-        val in_alt = mk_in (drop n Asets) bnf_sets T;
-        val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));
-      in
-        Goal.prove_sorry lthy [] [] goal (K lift_in_alt_tac) |> Thm.close_derivation
-      end;
-
-    fun le_rel_OO_tac _ = rtac (le_rel_OO_of_bnf bnf) 1;
-
-    fun rel_OO_Grp_tac _ = mk_simple_rel_OO_Grp_tac (rel_OO_Grp_of_bnf bnf) in_alt_thm;
-
-    val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac
-      bd_cinfinite_tac set_bd_tacs le_rel_OO_tac rel_OO_Grp_tac;
-
-    val wits = map snd (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf);
-
-    fun wit_tac _ = mk_simple_wit_tac (wit_thms_of_bnf bnf);
-
-    val (bnf', lthy') =
-      bnf_def Smart_Inline (K Dont_Note) qualify tacs wit_tac (SOME Ds) Binding.empty Binding.empty
-        [] ((((((b, T), mapx), sets), bd), wits), SOME rel) lthy;
-  in
-    (bnf', (add_bnf_to_unfolds bnf' unfold_set, lthy'))
-  end;
-
-(* Changing the order of live variables *)
-
-fun permute_bnf qualify src dest bnf (unfold_set, lthy) =
-  if src = dest then (bnf, (unfold_set, lthy)) else
-  let
-    val b = Binding.suffix_name (mk_permuteN src dest) (name_of_bnf bnf);
-    val live = live_of_bnf bnf;
-    val dead = dead_of_bnf bnf;
-    val nwits = nwits_of_bnf bnf;
-    fun permute xs = permute_like (op =) src dest xs;
-    fun unpermute xs = permute_like (op =) dest src xs;
-
-    val (Ds, lthy1) = apfst (map TFree)
-      (Variable.invent_types (replicate dead HOLogic.typeS) lthy);
-    val (As, lthy2) = apfst (map TFree)
-      (Variable.invent_types (replicate live HOLogic.typeS) lthy1);
-    val (Bs, _(*lthy3*)) = apfst (map TFree)
-      (Variable.invent_types (replicate live HOLogic.typeS) lthy2);
-
-    val (Asets, _(*names_lthy*)) = lthy
-      |> mk_Frees "A" (map HOLogic.mk_setT (permute As));
-
-    val T = mk_T_of_bnf Ds As bnf;
-
-    (*%f(1) ... f(n). bnf.map f\<sigma>(1) ... f\<sigma>(n)*)
-    val mapx = fold_rev Term.absdummy (permute (map2 (curry op -->) As Bs))
-      (Term.list_comb (mk_map_of_bnf Ds As Bs bnf, unpermute (map Bound (live - 1 downto 0))));
-    (*%Q(1) ... Q(n). bnf.rel Q\<sigma>(1) ... Q\<sigma>(n)*)
-    val rel = fold_rev Term.absdummy (permute (map2 mk_pred2T As Bs))
-      (Term.list_comb (mk_rel_of_bnf Ds As Bs bnf, unpermute (map Bound (live - 1 downto 0))));
-
-    val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf;
-    val sets = permute bnf_sets;
-
-    val bd = mk_bd_of_bnf Ds As bnf;
-
-    fun map_id0_tac _ = rtac (map_id0_of_bnf bnf) 1;
-    fun map_comp0_tac _ = rtac (map_comp0_of_bnf bnf) 1;
-    fun map_cong0_tac {context = ctxt, prems = _} =
-      rtac (map_cong0_of_bnf bnf) 1 THEN REPEAT_DETERM_N live (Goal.assume_rule_tac ctxt 1);
-    val set_map0_tacs = permute (map (fn thm => fn _ => rtac thm 1) (set_map0_of_bnf bnf));
-    fun bd_card_order_tac _ = rtac (bd_card_order_of_bnf bnf) 1;
-    fun bd_cinfinite_tac _ = rtac (bd_cinfinite_of_bnf bnf) 1;
-    val set_bd_tacs = permute (map (fn thm => fn _ => rtac thm 1) (set_bd_of_bnf bnf));
-
-    val in_alt_thm =
-      let
-        val inx = mk_in Asets sets T;
-        val in_alt = mk_in (unpermute Asets) bnf_sets T;
-        val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));
-      in
-        Goal.prove_sorry lthy [] [] goal (K (mk_permute_in_alt_tac src dest))
-        |> Thm.close_derivation
-      end;
-
-    fun le_rel_OO_tac _ = rtac (le_rel_OO_of_bnf bnf) 1;
-
-    fun rel_OO_Grp_tac _ = mk_simple_rel_OO_Grp_tac (rel_OO_Grp_of_bnf bnf) in_alt_thm;
-
-    val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac
-      bd_cinfinite_tac set_bd_tacs le_rel_OO_tac rel_OO_Grp_tac;
-
-    val wits = map snd (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf);
-
-    fun wit_tac _ = mk_simple_wit_tac (wit_thms_of_bnf bnf);
-
-    val (bnf', lthy') =
-      bnf_def Smart_Inline (K Dont_Note) qualify tacs wit_tac (SOME Ds) Binding.empty Binding.empty
-        [] ((((((b, T), mapx), sets), bd), wits), SOME rel) lthy;
-  in
-    (bnf', (add_bnf_to_unfolds bnf' unfold_set, lthy'))
-  end;
-
-(* Composition pipeline *)
-
-fun permute_and_kill qualify n src dest bnf =
-  bnf
-  |> permute_bnf qualify src dest
-  #> uncurry (kill_bnf qualify n);
-
-fun lift_and_permute qualify n src dest bnf =
-  bnf
-  |> lift_bnf qualify n
-  #> uncurry (permute_bnf qualify src dest);
-
-fun normalize_bnfs qualify Ass Ds sort bnfs unfold_set lthy =
-  let
-    val before_kill_src = map (fn As => 0 upto (length As - 1)) Ass;
-    val kill_poss = map (find_indices op = Ds) Ass;
-    val live_poss = map2 (subtract op =) kill_poss before_kill_src;
-    val before_kill_dest = map2 append kill_poss live_poss;
-    val kill_ns = map length kill_poss;
-    val (inners', (unfold_set', lthy')) =
-      fold_map5 (fn i => permute_and_kill (qualify i))
-        (if length bnfs = 1 then [0] else (1 upto length bnfs))
-        kill_ns before_kill_src before_kill_dest bnfs (unfold_set, lthy);
-
-    val Ass' = map2 (map o nth) Ass live_poss;
-    val As = sort Ass';
-    val after_lift_dest = replicate (length Ass') (0 upto (length As - 1));
-    val old_poss = map (map (fn x => find_index (fn y => x = y) As)) Ass';
-    val new_poss = map2 (subtract op =) old_poss after_lift_dest;
-    val after_lift_src = map2 append new_poss old_poss;
-    val lift_ns = map (fn xs => length As - length xs) Ass';
-  in
-    ((kill_poss, As), fold_map5 (fn i => lift_and_permute (qualify i))
-      (if length bnfs = 1 then [0] else (1 upto length bnfs))
-      lift_ns after_lift_src after_lift_dest inners' (unfold_set', lthy'))
-  end;
-
-fun default_comp_sort Ass =
-  Library.sort (Term_Ord.typ_ord o pairself TFree) (fold (fold (insert (op =))) Ass []);
-
-fun compose_bnf const_policy qualify sort outer inners oDs Dss tfreess (unfold_set, lthy) =
-  let
-    val b = name_of_bnf outer;
-
-    val Ass = map (map Term.dest_TFree) tfreess;
-    val Ds = fold (fold Term.add_tfreesT) (oDs :: Dss) [];
-
-    val ((kill_poss, As), (inners', (unfold_set', lthy'))) =
-      normalize_bnfs qualify Ass Ds sort inners unfold_set lthy;
-
-    val Ds = oDs @ flat (map3 (append oo map o nth) tfreess kill_poss Dss);
-    val As = map TFree As;
-  in
-    apfst (rpair (Ds, As))
-      (clean_compose_bnf const_policy (qualify 0) b outer inners' (unfold_set', lthy'))
-  end;
-
-(* Hide the type of the bound (optimization) and unfold the definitions (nicer to the user) *)
-
-fun seal_bnf qualify (unfold_set : unfold_set) b Ds bnf lthy =
-  let
-    val live = live_of_bnf bnf;
-    val nwits = nwits_of_bnf bnf;
-
-    val (As, lthy1) = apfst (map TFree)
-      (Variable.invent_types (replicate live HOLogic.typeS) (fold Variable.declare_typ Ds lthy));
-    val (Bs, _) = apfst (map TFree)
-      (Variable.invent_types (replicate live HOLogic.typeS) lthy1);
-
-    val map_unfolds = #map_unfolds unfold_set;
-    val set_unfoldss = #set_unfoldss unfold_set;
-    val rel_unfolds = #rel_unfolds unfold_set;
-
-    val expand_maps =
-      fold expand_term_const (map (single o Logic.dest_equals o Thm.prop_of) map_unfolds);
-    val expand_sets =
-      fold expand_term_const (map (map (Logic.dest_equals o Thm.prop_of)) set_unfoldss);
-    val expand_rels =
-      fold expand_term_const (map (single o Logic.dest_equals o Thm.prop_of) rel_unfolds);
-    fun unfold_maps ctxt = fold (unfold_thms ctxt o single) map_unfolds;
-    fun unfold_sets ctxt = fold (unfold_thms ctxt) set_unfoldss;
-    fun unfold_rels ctxt = unfold_thms ctxt rel_unfolds;
-    fun unfold_all ctxt = unfold_sets ctxt o unfold_maps ctxt o unfold_rels ctxt;
-    val bnf_map = expand_maps (mk_map_of_bnf Ds As Bs bnf);
-    val bnf_sets = map (expand_maps o expand_sets)
-      (mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf);
-    val bnf_bd = mk_bd_of_bnf Ds As bnf;
-    val bnf_rel = expand_rels (mk_rel_of_bnf Ds As Bs bnf);
-    val T = mk_T_of_bnf Ds As bnf;
-
-    (*bd should only depend on dead type variables!*)
-    val bd_repT = fst (dest_relT (fastype_of bnf_bd));
-    val bdT_bind = qualify (Binding.suffix_name ("_" ^ bdTN) b);
-    val params = fold Term.add_tfreesT Ds [];
-    val deads = map TFree params;
-
-    val ((bdT_name, (bdT_glob_info, bdT_loc_info)), lthy) =
-      typedef (bdT_bind, params, NoSyn)
-        (HOLogic.mk_UNIV bd_repT) NONE (EVERY' [rtac exI, rtac UNIV_I] 1) lthy;
-
-    val bnf_bd' = mk_dir_image bnf_bd
-      (Const (#Abs_name bdT_glob_info, bd_repT --> Type (bdT_name, deads)))
-
-    val Abs_bdT_inj = mk_Abs_inj_thm (#Abs_inject bdT_loc_info);
-    val Abs_bdT_bij = mk_Abs_bij_thm lthy Abs_bdT_inj (#Abs_cases bdT_loc_info);
-
-    val bd_ordIso = @{thm dir_image} OF [Abs_bdT_inj, bd_Card_order_of_bnf bnf];
-    val bd_card_order =
-      @{thm card_order_dir_image} OF [Abs_bdT_bij, bd_card_order_of_bnf bnf];
-    val bd_cinfinite =
-      (@{thm Cinfinite_cong} OF [bd_ordIso, bd_Cinfinite_of_bnf bnf]) RS conjunct1;
-
-    val set_bds =
-      map (fn thm => @{thm ordLeq_ordIso_trans} OF [thm, bd_ordIso]) (set_bd_of_bnf bnf);
-
-    fun mk_tac thm {context = ctxt, prems = _} =
-      (rtac (unfold_all ctxt thm) THEN'
-      SOLVE o REPEAT_DETERM o (atac ORELSE' Goal.assume_rule_tac ctxt)) 1;
-
-    val tacs = zip_axioms (mk_tac (map_id0_of_bnf bnf)) (mk_tac (map_comp0_of_bnf bnf))
-      (mk_tac (map_cong0_of_bnf bnf)) (map mk_tac (set_map0_of_bnf bnf))
-      (K (rtac bd_card_order 1)) (K (rtac bd_cinfinite 1)) (map mk_tac set_bds)
-      (mk_tac (le_rel_OO_of_bnf bnf))
-      (mk_tac (rel_OO_Grp_of_bnf bnf));
-
-    val bnf_wits = map snd (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf);
-
-    fun wit_tac {context = ctxt, prems = _} =
-      mk_simple_wit_tac (map (unfold_all ctxt) (wit_thms_of_bnf bnf));
-
-    val (bnf', lthy') =
-      bnf_def Hardly_Inline (user_policy Dont_Note) qualify tacs wit_tac (SOME deads)
-        Binding.empty Binding.empty []
-        ((((((b, T), bnf_map), bnf_sets), bnf_bd'), bnf_wits), SOME bnf_rel) lthy;
-  in
-    ((bnf', deads), lthy')
-  end;
-
-exception BAD_DEAD of typ * typ;
-
-fun bnf_of_typ _ _ _ _ (T as TFree _) accum = ((ID_bnf, ([], [T])), accum)
-  | bnf_of_typ _ _ _ _ (TVar _) _ = error "Unexpected schematic variable"
-  | bnf_of_typ const_policy qualify' sort Xs (T as Type (C, Ts)) (unfold_set, lthy) =
-    let
-      fun check_bad_dead ((_, (deads, _)), _) =
-        let val Ds = fold Term.add_tfreesT deads [] in
-          (case Library.inter (op =) Ds Xs of [] => ()
-           | X :: _ => raise BAD_DEAD (TFree X, T))
-        end;
-
-      val tfrees = Term.add_tfreesT T [];
-      val bnf_opt = if null tfrees then NONE else bnf_of lthy C;
-    in
-      (case bnf_opt of
-        NONE => ((DEADID_bnf, ([T], [])), (unfold_set, lthy))
-      | SOME bnf =>
-        if forall (can Term.dest_TFree) Ts andalso length Ts = length tfrees then
-          let
-            val T' = T_of_bnf bnf;
-            val deads = deads_of_bnf bnf;
-            val lives = lives_of_bnf bnf;
-            val tvars' = Term.add_tvarsT T' [];
-            val deads_lives =
-              pairself (map (Term.typ_subst_TVars (map fst tvars' ~~ map TFree tfrees)))
-                (deads, lives);
-          in ((bnf, deads_lives), (unfold_set, lthy)) end
-        else
-          let
-            val name = Long_Name.base_name C;
-            fun qualify i =
-              let val namei = name ^ nonzero_string_of_int i;
-              in qualify' o Binding.qualify true namei end;
-            val odead = dead_of_bnf bnf;
-            val olive = live_of_bnf bnf;
-            val oDs_pos = find_indices op = [TFree ("dead", [])] (snd (Term.dest_Type
-              (mk_T_of_bnf (replicate odead (TFree ("dead", []))) (replicate olive dummyT) bnf)));
-            val oDs = map (nth Ts) oDs_pos;
-            val Ts' = map (nth Ts) (subtract (op =) oDs_pos (0 upto length Ts - 1));
-            val ((inners, (Dss, Ass)), (unfold_set', lthy')) =
-              apfst (apsnd split_list o split_list)
-                (fold_map2 (fn i => bnf_of_typ Smart_Inline (qualify i) sort Xs)
-                (if length Ts' = 1 then [0] else (1 upto length Ts')) Ts' (unfold_set, lthy));
-          in
-            compose_bnf const_policy qualify sort bnf inners oDs Dss Ass (unfold_set', lthy')
-          end)
-      |> tap check_bad_dead
-    end;
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_comp_tactics.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,252 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_comp_tactics.ML
-    Author:     Dmitriy Traytel, TU Muenchen
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2012
-
-Tactics for composition of bounded natural functors.
-*)
-
-signature BNF_COMP_TACTICS =
-sig
-  val mk_comp_bd_card_order_tac: thm list -> thm -> tactic
-  val mk_comp_bd_cinfinite_tac: thm -> thm -> tactic
-  val mk_comp_in_alt_tac: Proof.context -> thm list -> tactic
-  val mk_comp_map_comp0_tac: thm -> thm -> thm list -> tactic
-  val mk_comp_map_cong0_tac: thm list -> thm -> thm list -> tactic
-  val mk_comp_map_id0_tac: thm -> thm -> thm list -> tactic
-  val mk_comp_set_alt_tac: Proof.context -> thm -> tactic
-  val mk_comp_set_bd_tac: Proof.context -> thm -> thm list -> tactic
-  val mk_comp_set_map0_tac: thm -> thm -> thm -> thm list -> tactic
-  val mk_comp_wit_tac: Proof.context -> thm list -> thm -> thm list -> tactic
-
-  val mk_kill_bd_card_order_tac: int -> thm -> tactic
-  val mk_kill_bd_cinfinite_tac: thm -> tactic
-  val kill_in_alt_tac: tactic
-  val mk_kill_map_cong0_tac: Proof.context -> int -> int -> thm -> tactic
-  val mk_kill_set_bd_tac: thm -> thm -> tactic
-
-  val empty_natural_tac: tactic
-  val lift_in_alt_tac: tactic
-  val mk_lift_set_bd_tac: thm -> tactic
-
-  val mk_permute_in_alt_tac: ''a list -> ''a list -> tactic
-
-  val mk_le_rel_OO_tac: thm -> thm -> thm list -> tactic
-  val mk_simple_rel_OO_Grp_tac: thm -> thm -> tactic
-  val mk_simple_wit_tac: thm list -> tactic
-end;
-
-structure BNF_Comp_Tactics : BNF_COMP_TACTICS =
-struct
-
-open BNF_Util
-open BNF_Tactics
-
-val Cnotzero_UNIV = @{thm Cnotzero_UNIV};
-val arg_cong_Union = @{thm arg_cong[of _ _ Union]};
-val csum_Cnotzero1 = @{thm csum_Cnotzero1};
-val o_eq_dest_lhs = @{thm o_eq_dest_lhs};
-val trans_image_cong_o_apply = @{thm trans[OF image_cong[OF o_apply refl]]};
-val trans_o_apply = @{thm trans[OF o_apply]};
-
-
-
-(* Composition *)
-
-fun mk_comp_set_alt_tac ctxt collect_set_map =
-  unfold_thms_tac ctxt @{thms sym[OF o_assoc]} THEN
-  unfold_thms_tac ctxt [collect_set_map RS sym] THEN
-  rtac refl 1;
-
-fun mk_comp_map_id0_tac Gmap_id0 Gmap_cong0 map_id0s =
-  EVERY' ([rtac ext, rtac (Gmap_cong0 RS trans)] @
-    map (fn thm => rtac (thm RS fun_cong)) map_id0s @ [rtac (Gmap_id0 RS fun_cong)]) 1;
-
-fun mk_comp_map_comp0_tac Gmap_comp0 Gmap_cong0 map_comp0s =
-  EVERY' ([rtac ext, rtac sym, rtac trans_o_apply,
-    rtac (Gmap_comp0 RS sym RS o_eq_dest_lhs RS trans), rtac Gmap_cong0] @
-    map (fn thm => rtac (thm RS sym RS fun_cong)) map_comp0s) 1;
-
-fun mk_comp_set_map0_tac Gmap_comp0 Gmap_cong0 Gset_map0 set_map0s =
-  EVERY' ([rtac ext] @
-    replicate 3 (rtac trans_o_apply) @
-    [rtac (arg_cong_Union RS trans),
-     rtac (@{thm arg_cong2[of _ _ _ _ collect, OF refl]} RS trans),
-     rtac (Gmap_comp0 RS sym RS o_eq_dest_lhs RS trans),
-     rtac Gmap_cong0] @
-     map (fn thm => rtac (thm RS fun_cong)) set_map0s @
-     [rtac (Gset_map0 RS o_eq_dest_lhs), rtac sym, rtac trans_o_apply,
-     rtac trans_image_cong_o_apply, rtac trans_image_cong_o_apply,
-     rtac (@{thm image_cong} OF [Gset_map0 RS o_eq_dest_lhs RS arg_cong_Union, refl] RS trans),
-     rtac @{thm trans[OF comp_eq_dest[OF Union_natural[symmetric]]]}, rtac arg_cong_Union,
-     rtac @{thm trans[OF o_eq_dest_lhs[OF image_o_collect[symmetric]]]},
-     rtac @{thm fun_cong[OF arg_cong[of _ _ collect]]}] @
-     [REPEAT_DETERM_N (length set_map0s) o EVERY' [rtac @{thm trans[OF image_insert]},
-        rtac @{thm arg_cong2[of _ _ _ _ insert]}, rtac ext, rtac trans_o_apply,
-        rtac trans_image_cong_o_apply, rtac @{thm trans[OF image_image]},
-        rtac @{thm sym[OF trans[OF o_apply]]}, rtac @{thm image_cong[OF refl o_apply]}],
-     rtac @{thm image_empty}]) 1;
-
-fun mk_comp_map_cong0_tac comp_set_alts map_cong0 map_cong0s =
-  let
-     val n = length comp_set_alts;
-  in
-    (if n = 0 then rtac refl 1
-    else rtac map_cong0 1 THEN
-      EVERY' (map_index (fn (i, map_cong0) =>
-        rtac map_cong0 THEN' EVERY' (map_index (fn (k, set_alt) =>
-          EVERY' [select_prem_tac n (dtac @{thm meta_spec}) (k + 1), etac meta_mp,
-            rtac (equalityD2 RS set_mp), rtac (set_alt RS fun_cong RS trans),
-            rtac trans_o_apply, rtac (@{thm collect_def} RS arg_cong_Union),
-            rtac @{thm UnionI}, rtac @{thm UN_I}, REPEAT_DETERM_N i o rtac @{thm insertI2},
-            rtac @{thm insertI1}, rtac (o_apply RS equalityD2 RS set_mp),
-            etac @{thm imageI}, atac])
-          comp_set_alts))
-      map_cong0s) 1)
-  end;
-
-fun mk_comp_bd_card_order_tac Fbd_card_orders Gbd_card_order =
-  let
-    val (card_orders, last_card_order) = split_last Fbd_card_orders;
-    fun gen_before thm = rtac @{thm card_order_csum} THEN' rtac thm;
-  in
-    (rtac @{thm card_order_cprod} THEN'
-    WRAP' gen_before (K (K all_tac)) card_orders (rtac last_card_order) THEN'
-    rtac Gbd_card_order) 1
-  end;
-
-fun mk_comp_bd_cinfinite_tac Fbd_cinfinite Gbd_cinfinite =
-  (rtac @{thm cinfinite_cprod} THEN'
-   ((K (TRY ((rtac @{thm cinfinite_csum} THEN' rtac disjI1) 1)) THEN'
-     ((rtac @{thm cinfinite_csum} THEN' rtac disjI1 THEN' rtac Fbd_cinfinite) ORELSE'
-      rtac Fbd_cinfinite)) ORELSE'
-    rtac Fbd_cinfinite) THEN'
-   rtac Gbd_cinfinite) 1;
-
-fun mk_comp_set_bd_tac ctxt comp_set_alt Gset_Fset_bds =
-  let
-    val (bds, last_bd) = split_last Gset_Fset_bds;
-    fun gen_before bd =
-      rtac ctrans THEN' rtac @{thm Un_csum} THEN'
-      rtac ctrans THEN' rtac @{thm csum_mono} THEN'
-      rtac bd;
-    fun gen_after _ = rtac @{thm ordIso_imp_ordLeq} THEN' rtac @{thm cprod_csum_distrib1};
-  in
-    unfold_thms_tac ctxt [comp_set_alt] THEN
-    rtac @{thm comp_set_bd_Union_o_collect} 1 THEN
-    unfold_thms_tac ctxt @{thms Union_image_insert Union_image_empty Union_Un_distrib o_apply} THEN
-    (rtac ctrans THEN'
-     WRAP' gen_before gen_after bds (rtac last_bd) THEN'
-     rtac @{thm ordIso_imp_ordLeq} THEN'
-     rtac @{thm cprod_com}) 1
-  end;
-
-val comp_in_alt_thms = @{thms o_apply collect_def SUP_def image_insert image_empty Union_insert
-  Union_empty Un_empty_right Union_Un_distrib Un_subset_iff conj_subset_def UN_image_subset
-  conj_assoc};
-
-fun mk_comp_in_alt_tac ctxt comp_set_alts =
-  unfold_thms_tac ctxt (comp_set_alts @ comp_in_alt_thms) THEN
-  unfold_thms_tac ctxt @{thms set_eq_subset} THEN
-  rtac conjI 1 THEN
-  REPEAT_DETERM (
-    rtac @{thm subsetI} 1 THEN
-    unfold_thms_tac ctxt @{thms mem_Collect_eq Ball_def} THEN
-    (REPEAT_DETERM (CHANGED (etac conjE 1)) THEN
-     REPEAT_DETERM (CHANGED ((
-       (rtac conjI THEN' (atac ORELSE' rtac subset_UNIV)) ORELSE'
-       atac ORELSE'
-       (rtac subset_UNIV)) 1)) ORELSE rtac subset_UNIV 1));
-
-val comp_wit_thms = @{thms Union_empty_conv o_apply collect_def SUP_def
-  Union_image_insert Union_image_empty};
-
-fun mk_comp_wit_tac ctxt Gwit_thms collect_set_map Fwit_thms =
-  ALLGOALS (dtac @{thm in_Union_o_assoc}) THEN
-  unfold_thms_tac ctxt (collect_set_map :: comp_wit_thms) THEN
-  REPEAT_DETERM ((atac ORELSE'
-    REPEAT_DETERM o eresolve_tac @{thms UnionE UnE} THEN'
-    etac imageE THEN' TRY o dresolve_tac Gwit_thms THEN'
-    (etac FalseE ORELSE'
-    hyp_subst_tac ctxt THEN'
-    dresolve_tac Fwit_thms THEN'
-    (etac FalseE ORELSE' atac))) 1);
-
-
-
-(* Kill operation *)
-
-fun mk_kill_map_cong0_tac ctxt n m map_cong0 =
-  (rtac map_cong0 THEN' EVERY' (replicate n (rtac refl)) THEN'
-    EVERY' (replicate m (Goal.assume_rule_tac ctxt))) 1;
-
-fun mk_kill_bd_card_order_tac n bd_card_order =
-  (rtac @{thm card_order_cprod} THEN'
-  K (REPEAT_DETERM_N (n - 1)
-    ((rtac @{thm card_order_csum} THEN'
-    rtac @{thm card_of_card_order_on}) 1)) THEN'
-  rtac @{thm card_of_card_order_on} THEN'
-  rtac bd_card_order) 1;
-
-fun mk_kill_bd_cinfinite_tac bd_Cinfinite =
-  (rtac @{thm cinfinite_cprod2} THEN'
-  TRY o rtac csum_Cnotzero1 THEN'
-  rtac Cnotzero_UNIV THEN'
-  rtac bd_Cinfinite) 1;
-
-fun mk_kill_set_bd_tac bd_Card_order set_bd =
-  (rtac ctrans THEN'
-  rtac set_bd THEN'
-  rtac @{thm ordLeq_cprod2} THEN'
-  TRY o rtac csum_Cnotzero1 THEN'
-  rtac Cnotzero_UNIV THEN'
-  rtac bd_Card_order) 1
-
-val kill_in_alt_tac =
-  ((rtac @{thm Collect_cong} THEN' rtac iffI) 1 THEN
-  REPEAT_DETERM (CHANGED (etac conjE 1)) THEN
-  REPEAT_DETERM (CHANGED ((etac conjI ORELSE'
-    rtac conjI THEN' rtac subset_UNIV) 1)) THEN
-  (rtac subset_UNIV ORELSE' atac) 1 THEN
-  REPEAT_DETERM (CHANGED (etac conjE 1)) THEN
-  REPEAT_DETERM (CHANGED ((etac conjI ORELSE' atac) 1))) ORELSE
-  ((rtac @{thm UNIV_eq_I} THEN' rtac CollectI) 1 THEN
-    REPEAT_DETERM (TRY (rtac conjI 1) THEN rtac subset_UNIV 1));
-
-
-
-(* Lift operation *)
-
-val empty_natural_tac = rtac @{thm empty_natural} 1;
-
-fun mk_lift_set_bd_tac bd_Card_order = (rtac @{thm Card_order_empty} THEN' rtac bd_Card_order) 1;
-
-val lift_in_alt_tac =
-  ((rtac @{thm Collect_cong} THEN' rtac iffI) 1 THEN
-  REPEAT_DETERM (CHANGED (etac conjE 1)) THEN
-  REPEAT_DETERM (CHANGED ((etac conjI ORELSE' atac) 1)) THEN
-  REPEAT_DETERM (CHANGED (etac conjE 1)) THEN
-  REPEAT_DETERM (CHANGED ((etac conjI ORELSE'
-    rtac conjI THEN' rtac @{thm empty_subsetI}) 1)) THEN
-  (rtac @{thm empty_subsetI} ORELSE' atac) 1) ORELSE
-  ((rtac sym THEN' rtac @{thm UNIV_eq_I} THEN' rtac CollectI) 1 THEN
-    REPEAT_DETERM (TRY (rtac conjI 1) THEN rtac @{thm empty_subsetI} 1));
-
-
-
-(* Permute operation *)
-
-fun mk_permute_in_alt_tac src dest =
-  (rtac @{thm Collect_cong} THEN'
-  mk_rotate_eq_tac (rtac refl) trans @{thm conj_assoc} @{thm conj_commute} @{thm conj_cong}
-    dest src) 1;
-
-fun mk_le_rel_OO_tac outer_le_rel_OO outer_rel_mono inner_le_rel_OOs =
-  EVERY' (map rtac (@{thm order_trans} :: outer_le_rel_OO :: outer_rel_mono :: inner_le_rel_OOs)) 1;
-
-fun mk_simple_rel_OO_Grp_tac rel_OO_Grp in_alt_thm =
-  rtac (trans OF [rel_OO_Grp, in_alt_thm RS @{thm OO_Grp_cong} RS sym]) 1;
-
-fun mk_simple_wit_tac wit_thms = ALLGOALS (atac ORELSE' eresolve_tac (@{thm emptyE} :: wit_thms));
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_decl.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,117 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_decl.ML
-    Author:     Dmitriy Traytel, TU Muenchen
-    Copyright   2013
-
-Axiomatic declaration of bounded natural functors.
-*)
-
-signature BNF_DECL =
-sig
-  val bnf_decl: (binding option * (typ * sort)) list -> binding -> mixfix -> binding -> binding ->
-    typ list -> local_theory -> BNF_Def.bnf * local_theory
-end
-
-structure BNF_Decl : BNF_DECL =
-struct
-
-open BNF_Util
-open BNF_Def
-
-fun prepare_decl prepare_constraint prepare_typ raw_vars b mx user_mapb user_relb user_witTs lthy =
-  let
-   fun prepare_type_arg (set_opt, (ty, c)) =
-      let val s = fst (dest_TFree (prepare_typ lthy ty)) in
-        (set_opt, (s, prepare_constraint lthy c))
-      end;
-    val ((user_setbs, vars), raw_vars') =
-      map prepare_type_arg raw_vars
-      |> `split_list
-      |>> apfst (map_filter I);
-    val deads = map_filter (fn (NONE, x) => SOME x | _ => NONE) raw_vars';
-
-    fun mk_b name user_b =
-      (if Binding.is_empty user_b then Binding.prefix_name (name ^ "_") b else user_b)
-      |> Binding.qualify false (Binding.name_of b);
-    val (Tname, lthy) = Typedecl.basic_typedecl (b, length vars, mx) lthy;
-    val (bd_type_Tname, lthy) =
-      Typedecl.basic_typedecl (mk_b "bd_type" Binding.empty, length deads, NoSyn) lthy;
-    val T = Type (Tname, map TFree vars);
-    val bd_type_T = Type (bd_type_Tname, map TFree deads);
-    val lives = map TFree (filter_out (member (op =) deads) vars);
-    val live = length lives;
-    val _ = "Trying to declare a BNF with no live variables" |> null lives ? error;
-    val (lives', _) = BNF_Util.mk_TFrees (length lives)
-      (fold Variable.declare_typ (map TFree vars) lthy);
-    val T' = Term.typ_subst_atomic (lives ~~ lives') T;
-    val mapT = map2 (curry op -->) lives lives' ---> T --> T';
-    val setTs = map (fn U => T --> HOLogic.mk_setT U) lives;
-    val bdT = BNF_Util.mk_relT (bd_type_T, bd_type_T);
-    val mapb = mk_b BNF_Def.mapN user_mapb;
-    val bdb = mk_b "bd" Binding.empty;
-    val setbs = map2 (fn b => fn i => mk_b (BNF_Def.mk_setN i) b) user_setbs
-      (if live = 1 then [0] else 1 upto live);
-
-    val witTs = map (prepare_typ lthy) user_witTs;
-    val nwits = length witTs;
-    val witbs = map (fn i => mk_b (BNF_Def.mk_witN i) Binding.empty)
-      (if nwits = 1 then [0] else 1 upto nwits);
-
-    val lthy = Local_Theory.background_theory
-      (Sign.add_consts_i ((mapb, mapT, NoSyn) :: (bdb, bdT, NoSyn) ::
-        map2 (fn b => fn T => (b, T, NoSyn)) setbs setTs @
-        map2 (fn b => fn T => (b, T, NoSyn)) witbs witTs))
-      lthy;
-    val Fmap = Const (Local_Theory.full_name lthy mapb, mapT);
-    val Fsets = map2 (fn setb => fn setT =>
-      Const (Local_Theory.full_name lthy setb, setT)) setbs setTs;
-    val Fbd = Const (Local_Theory.full_name lthy bdb, bdT);
-    val Fwits = map2 (fn witb => fn witT =>
-      Const (Local_Theory.full_name lthy witb, witT)) witbs witTs;
-    val (key, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, _) =
-      prepare_def Do_Inline (user_policy Note_Some) I (K I) (K I) (SOME (map TFree deads))
-      user_mapb user_relb user_setbs ((((((Binding.empty, T), Fmap), Fsets), Fbd), Fwits), NONE)
-      lthy;
-
-    fun mk_wits_tac set_maps = K (TRYALL Goal.conjunction_tac) THEN' the triv_tac_opt set_maps;
-    val wit_goals = map Logic.mk_conjunction_balanced wit_goalss;
-    val all_goalss = map single goals @ (if nwits > 0 then wit_goalss else []);
-
-    val (((_, [raw_thms])), (lthy_old, lthy)) = Local_Theory.background_theory_result
-      (Specification.axiomatization [] [((mk_b "axioms" Binding.empty, []), flat all_goalss)]) lthy
-      ||> `Local_Theory.restore;
-
-    fun mk_wit_thms set_maps =
-      Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals) (mk_wits_tac set_maps)
-      |> Conjunction.elim_balanced (length wit_goals)
-      |> map2 (Conjunction.elim_balanced o length) wit_goalss
-      |> map (map (Thm.close_derivation o Thm.forall_elim_vars 0));
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-    val thms = unflat all_goalss (Morphism.fact phi raw_thms);
-  in
-    BNF_Def.register_bnf key (after_qed mk_wit_thms thms lthy)
-  end;
-
-val bnf_decl = prepare_decl (K I) (K I);
-
-fun read_constraint _ NONE = HOLogic.typeS
-  | read_constraint ctxt (SOME s) = Syntax.read_sort ctxt s;
-
-val bnf_decl_cmd = prepare_decl read_constraint Syntax.read_typ;
-
-val parse_witTs =
-  @{keyword "["} |-- (Parse.short_ident --| @{keyword ":"} -- Scan.repeat Parse.typ
-    >> (fn ("wits", Ts) => Ts
-         | (s, _) => error ("Unknown label " ^ quote s ^ " (expected \"wits\")"))) --|
-  @{keyword "]"} || Scan.succeed [];
-
-val parse_bnf_decl =
-  parse_type_args_named_constrained -- parse_binding -- parse_map_rel_bindings --
-    parse_witTs -- Parse.opt_mixfix;
-
-val _ =
-  Outer_Syntax.local_theory @{command_spec "bnf_decl"} "bnf declaration"
-    (parse_bnf_decl >> 
-      (fn ((((bsTs, b), (mapb, relb)), witTs), mx) =>
-         bnf_decl_cmd bsTs b mx mapb relb witTs #> snd));
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_def.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1393 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_def.ML
-    Author:     Dmitriy Traytel, TU Muenchen
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2012
-
-Definition of bounded natural functors.
-*)
-
-signature BNF_DEF =
-sig
-  type bnf
-  type nonemptiness_witness = {I: int list, wit: term, prop: thm list}
-
-  val morph_bnf: morphism -> bnf -> bnf
-  val eq_bnf: bnf * bnf -> bool
-  val bnf_of: Proof.context -> string -> bnf option
-  val register_bnf: string -> (bnf * local_theory) -> (bnf * local_theory)
-
-  val name_of_bnf: bnf -> binding
-  val T_of_bnf: bnf -> typ
-  val live_of_bnf: bnf -> int
-  val lives_of_bnf: bnf -> typ list
-  val dead_of_bnf: bnf -> int
-  val deads_of_bnf: bnf -> typ list
-  val nwits_of_bnf: bnf -> int
-
-  val mapN: string
-  val relN: string
-  val setN: string
-  val mk_setN: int -> string
-  val mk_witN: int -> string
-
-  val map_of_bnf: bnf -> term
-  val sets_of_bnf: bnf -> term list
-  val rel_of_bnf: bnf -> term
-
-  val mk_T_of_bnf: typ list -> typ list -> bnf -> typ
-  val mk_bd_of_bnf: typ list -> typ list -> bnf -> term
-  val mk_map_of_bnf: typ list -> typ list -> typ list -> bnf -> term
-  val mk_rel_of_bnf: typ list -> typ list -> typ list -> bnf -> term
-  val mk_sets_of_bnf: typ list list -> typ list list -> bnf -> term list
-  val mk_wits_of_bnf: typ list list -> typ list list -> bnf -> (int list * term) list
-
-  val bd_Card_order_of_bnf: bnf -> thm
-  val bd_Cinfinite_of_bnf: bnf -> thm
-  val bd_Cnotzero_of_bnf: bnf -> thm
-  val bd_card_order_of_bnf: bnf -> thm
-  val bd_cinfinite_of_bnf: bnf -> thm
-  val collect_set_map_of_bnf: bnf -> thm
-  val in_bd_of_bnf: bnf -> thm
-  val in_cong_of_bnf: bnf -> thm
-  val in_mono_of_bnf: bnf -> thm
-  val in_rel_of_bnf: bnf -> thm
-  val map_comp0_of_bnf: bnf -> thm
-  val map_comp_of_bnf: bnf -> thm
-  val map_cong0_of_bnf: bnf -> thm
-  val map_cong_of_bnf: bnf -> thm
-  val map_def_of_bnf: bnf -> thm
-  val map_id0_of_bnf: bnf -> thm
-  val map_id_of_bnf: bnf -> thm
-  val map_transfer_of_bnf: bnf -> thm
-  val le_rel_OO_of_bnf: bnf -> thm
-  val rel_def_of_bnf: bnf -> thm
-  val rel_Grp_of_bnf: bnf -> thm
-  val rel_OO_of_bnf: bnf -> thm
-  val rel_OO_Grp_of_bnf: bnf -> thm
-  val rel_cong_of_bnf: bnf -> thm
-  val rel_conversep_of_bnf: bnf -> thm
-  val rel_mono_of_bnf: bnf -> thm
-  val rel_mono_strong_of_bnf: bnf -> thm
-  val rel_eq_of_bnf: bnf -> thm
-  val rel_flip_of_bnf: bnf -> thm
-  val set_bd_of_bnf: bnf -> thm list
-  val set_defs_of_bnf: bnf -> thm list
-  val set_map0_of_bnf: bnf -> thm list
-  val set_map_of_bnf: bnf -> thm list
-  val wit_thms_of_bnf: bnf -> thm list
-  val wit_thmss_of_bnf: bnf -> thm list list
-
-  val mk_map: int -> typ list -> typ list -> term -> term
-  val mk_rel: int -> typ list -> typ list -> term -> term
-  val build_map: Proof.context -> (typ * typ -> term) -> typ * typ -> term
-  val build_rel: Proof.context -> (typ * typ -> term) -> typ * typ -> term
-  val flatten_type_args_of_bnf: bnf -> 'a -> 'a list -> 'a list
-  val map_flattened_map_args: Proof.context -> string -> (term list -> 'a list) -> term list ->
-    'a list
-
-  val mk_witness: int list * term -> thm list -> nonemptiness_witness
-  val minimize_wits: (''a list * 'b) list -> (''a list * 'b) list
-  val wits_of_bnf: bnf -> nonemptiness_witness list
-
-  val zip_axioms: 'a -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list
-
-  datatype const_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline
-  datatype fact_policy = Dont_Note | Note_Some | Note_All
-
-  val bnf_note_all: bool Config.T
-  val bnf_timing: bool Config.T
-  val user_policy: fact_policy -> Proof.context -> fact_policy
-  val note_bnf_thms: fact_policy -> (binding -> binding) -> binding -> bnf -> Proof.context ->
-    Proof.context
-
-  val print_bnfs: Proof.context -> unit
-  val prepare_def: const_policy -> (Proof.context -> fact_policy) -> (binding -> binding) ->
-    (Proof.context -> 'a -> typ) -> (Proof.context -> 'b -> term) -> typ list option ->
-    binding -> binding -> binding list ->
-    (((((binding * 'a) * 'b) * 'b list) * 'b) * 'b list) * 'b option -> Proof.context ->
-    string * term list *
-    ((thm list -> {context: Proof.context, prems: thm list} -> tactic) option * term list list) *
-    ((thm list -> thm list list) -> thm list list -> Proof.context -> bnf * local_theory) *
-    local_theory * thm list
-
-  val define_bnf_consts: const_policy -> fact_policy -> typ list option ->
-    binding -> binding -> binding list ->
-    (((((binding * typ) * term) * term list) * term) * term list) * term option -> local_theory ->
-      ((typ list * typ list * typ list * typ) *
-       (term * term list * term * (int list * term) list * term) *
-       (thm * thm list * thm * thm list * thm) *
-       ((typ list -> typ list -> typ list -> term) *
-        (typ list -> typ list -> term -> term) *
-        (typ list -> typ list -> typ -> typ) *
-        (typ list -> typ list -> typ list -> term) *
-        (typ list -> typ list -> typ list -> term))) * local_theory
-
-  val bnf_def: const_policy -> (Proof.context -> fact_policy) -> (binding -> binding) ->
-    ({prems: thm list, context: Proof.context} -> tactic) list ->
-    ({prems: thm list, context: Proof.context} -> tactic) -> typ list option -> binding ->
-    binding -> binding list ->
-    (((((binding * typ) * term) * term list) * term) * term list) * term option ->
-    local_theory -> bnf * local_theory
-end;
-
-structure BNF_Def : BNF_DEF =
-struct
-
-open BNF_Util
-open BNF_Tactics
-open BNF_Def_Tactics
-
-val fundefcong_attrs = @{attributes [fundef_cong]};
-
-type axioms = {
-  map_id0: thm,
-  map_comp0: thm,
-  map_cong0: thm,
-  set_map0: thm list,
-  bd_card_order: thm,
-  bd_cinfinite: thm,
-  set_bd: thm list,
-  le_rel_OO: thm,
-  rel_OO_Grp: thm
-};
-
-fun mk_axioms' ((((((((id, comp), cong), map), c_o), cinf), set_bd), le_rel_OO), rel) =
-  {map_id0 = id, map_comp0 = comp, map_cong0 = cong, set_map0 = map, bd_card_order = c_o,
-   bd_cinfinite = cinf, set_bd = set_bd, le_rel_OO = le_rel_OO, rel_OO_Grp = rel};
-
-fun dest_cons [] = raise List.Empty
-  | dest_cons (x :: xs) = (x, xs);
-
-fun mk_axioms n thms = thms
-  |> map the_single
-  |> dest_cons
-  ||>> dest_cons
-  ||>> dest_cons
-  ||>> chop n
-  ||>> dest_cons
-  ||>> dest_cons
-  ||>> chop n
-  ||>> dest_cons
-  ||> the_single
-  |> mk_axioms';
-
-fun zip_axioms mid mcomp mcong smap bdco bdinf sbd le_rel_OO rel =
-  [mid, mcomp, mcong] @ smap @ [bdco, bdinf] @ sbd @ [le_rel_OO, rel];
-
-fun dest_axioms {map_id0, map_comp0, map_cong0, set_map0, bd_card_order, bd_cinfinite, set_bd,
-  le_rel_OO, rel_OO_Grp} =
-  zip_axioms map_id0 map_comp0 map_cong0 set_map0 bd_card_order bd_cinfinite set_bd le_rel_OO
-    rel_OO_Grp;
-
-fun map_axioms f {map_id0, map_comp0, map_cong0, set_map0, bd_card_order, bd_cinfinite, set_bd,
-  le_rel_OO, rel_OO_Grp} =
-  {map_id0 = f map_id0,
-    map_comp0 = f map_comp0,
-    map_cong0 = f map_cong0,
-    set_map0 = map f set_map0,
-    bd_card_order = f bd_card_order,
-    bd_cinfinite = f bd_cinfinite,
-    set_bd = map f set_bd,
-    le_rel_OO = f le_rel_OO,
-    rel_OO_Grp = f rel_OO_Grp};
-
-val morph_axioms = map_axioms o Morphism.thm;
-
-type defs = {
-  map_def: thm,
-  set_defs: thm list,
-  rel_def: thm
-}
-
-fun mk_defs map sets rel = {map_def = map, set_defs = sets, rel_def = rel};
-
-fun map_defs f {map_def, set_defs, rel_def} =
-  {map_def = f map_def, set_defs = map f set_defs, rel_def = f rel_def};
-
-val morph_defs = map_defs o Morphism.thm;
-
-type facts = {
-  bd_Card_order: thm,
-  bd_Cinfinite: thm,
-  bd_Cnotzero: thm,
-  collect_set_map: thm lazy,
-  in_bd: thm lazy,
-  in_cong: thm lazy,
-  in_mono: thm lazy,
-  in_rel: thm lazy,
-  map_comp: thm lazy,
-  map_cong: thm lazy,
-  map_id: thm lazy,
-  map_transfer: thm lazy,
-  rel_eq: thm lazy,
-  rel_flip: thm lazy,
-  set_map: thm lazy list,
-  rel_cong: thm lazy,
-  rel_mono: thm lazy,
-  rel_mono_strong: thm lazy,
-  rel_Grp: thm lazy,
-  rel_conversep: thm lazy,
-  rel_OO: thm lazy
-};
-
-fun mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_map in_bd in_cong in_mono in_rel
-    map_comp map_cong map_id map_transfer rel_eq rel_flip set_map rel_cong rel_mono
-    rel_mono_strong rel_Grp rel_conversep rel_OO = {
-  bd_Card_order = bd_Card_order,
-  bd_Cinfinite = bd_Cinfinite,
-  bd_Cnotzero = bd_Cnotzero,
-  collect_set_map = collect_set_map,
-  in_bd = in_bd,
-  in_cong = in_cong,
-  in_mono = in_mono,
-  in_rel = in_rel,
-  map_comp = map_comp,
-  map_cong = map_cong,
-  map_id = map_id,
-  map_transfer = map_transfer,
-  rel_eq = rel_eq,
-  rel_flip = rel_flip,
-  set_map = set_map,
-  rel_cong = rel_cong,
-  rel_mono = rel_mono,
-  rel_mono_strong = rel_mono_strong,
-  rel_Grp = rel_Grp,
-  rel_conversep = rel_conversep,
-  rel_OO = rel_OO};
-
-fun map_facts f {
-  bd_Card_order,
-  bd_Cinfinite,
-  bd_Cnotzero,
-  collect_set_map,
-  in_bd,
-  in_cong,
-  in_mono,
-  in_rel,
-  map_comp,
-  map_cong,
-  map_id,
-  map_transfer,
-  rel_eq,
-  rel_flip,
-  set_map,
-  rel_cong,
-  rel_mono,
-  rel_mono_strong,
-  rel_Grp,
-  rel_conversep,
-  rel_OO} =
-  {bd_Card_order = f bd_Card_order,
-    bd_Cinfinite = f bd_Cinfinite,
-    bd_Cnotzero = f bd_Cnotzero,
-    collect_set_map = Lazy.map f collect_set_map,
-    in_bd = Lazy.map f in_bd,
-    in_cong = Lazy.map f in_cong,
-    in_mono = Lazy.map f in_mono,
-    in_rel = Lazy.map f in_rel,
-    map_comp = Lazy.map f map_comp,
-    map_cong = Lazy.map f map_cong,
-    map_id = Lazy.map f map_id,
-    map_transfer = Lazy.map f map_transfer,
-    rel_eq = Lazy.map f rel_eq,
-    rel_flip = Lazy.map f rel_flip,
-    set_map = map (Lazy.map f) set_map,
-    rel_cong = Lazy.map f rel_cong,
-    rel_mono = Lazy.map f rel_mono,
-    rel_mono_strong = Lazy.map f rel_mono_strong,
-    rel_Grp = Lazy.map f rel_Grp,
-    rel_conversep = Lazy.map f rel_conversep,
-    rel_OO = Lazy.map f rel_OO};
-
-val morph_facts = map_facts o Morphism.thm;
-
-type nonemptiness_witness = {
-  I: int list,
-  wit: term,
-  prop: thm list
-};
-
-fun mk_witness (I, wit) prop = {I = I, wit = wit, prop = prop};
-fun map_witness f g {I, wit, prop} = {I = I, wit = f wit, prop = map g prop};
-fun morph_witness phi = map_witness (Morphism.term phi) (Morphism.thm phi);
-
-datatype bnf = BNF of {
-  name: binding,
-  T: typ,
-  live: int,
-  lives: typ list, (*source type variables of map*)
-  lives': typ list, (*target type variables of map*)
-  dead: int,
-  deads: typ list,
-  map: term,
-  sets: term list,
-  bd: term,
-  axioms: axioms,
-  defs: defs,
-  facts: facts,
-  nwits: int,
-  wits: nonemptiness_witness list,
-  rel: term
-};
-
-(* getters *)
-
-fun rep_bnf (BNF bnf) = bnf;
-val name_of_bnf = #name o rep_bnf;
-val T_of_bnf = #T o rep_bnf;
-fun mk_T_of_bnf Ds Ts bnf =
-  let val bnf_rep = rep_bnf bnf
-  in Term.typ_subst_atomic ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#T bnf_rep) end;
-val live_of_bnf = #live o rep_bnf;
-val lives_of_bnf = #lives o rep_bnf;
-val dead_of_bnf = #dead o rep_bnf;
-val deads_of_bnf = #deads o rep_bnf;
-val axioms_of_bnf = #axioms o rep_bnf;
-val facts_of_bnf = #facts o rep_bnf;
-val nwits_of_bnf = #nwits o rep_bnf;
-val wits_of_bnf = #wits o rep_bnf;
-
-fun flatten_type_args_of_bnf bnf dead_x xs =
-  let
-    val Type (_, Ts) = T_of_bnf bnf;
-    val lives = lives_of_bnf bnf;
-    val deads = deads_of_bnf bnf;
-  in
-    permute_like (op =) (deads @ lives) Ts (replicate (length deads) dead_x @ xs)
-  end;
-
-(*terms*)
-val map_of_bnf = #map o rep_bnf;
-val sets_of_bnf = #sets o rep_bnf;
-fun mk_map_of_bnf Ds Ts Us bnf =
-  let val bnf_rep = rep_bnf bnf;
-  in
-    Term.subst_atomic_types
-      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#map bnf_rep)
-  end;
-fun mk_sets_of_bnf Dss Tss bnf =
-  let val bnf_rep = rep_bnf bnf;
-  in
-    map2 (fn (Ds, Ts) => Term.subst_atomic_types
-      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts))) (Dss ~~ Tss) (#sets bnf_rep)
-  end;
-val bd_of_bnf = #bd o rep_bnf;
-fun mk_bd_of_bnf Ds Ts bnf =
-  let val bnf_rep = rep_bnf bnf;
-  in Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#bd bnf_rep) end;
-fun mk_wits_of_bnf Dss Tss bnf =
-  let
-    val bnf_rep = rep_bnf bnf;
-    val wits = map (fn x => (#I x, #wit x)) (#wits bnf_rep);
-  in
-    map2 (fn (Ds, Ts) => apsnd (Term.subst_atomic_types
-      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)))) (Dss ~~ Tss) wits
-  end;
-val rel_of_bnf = #rel o rep_bnf;
-fun mk_rel_of_bnf Ds Ts Us bnf =
-  let val bnf_rep = rep_bnf bnf;
-  in
-    Term.subst_atomic_types
-      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#rel bnf_rep)
-  end;
-
-(*thms*)
-val bd_card_order_of_bnf = #bd_card_order o #axioms o rep_bnf;
-val bd_cinfinite_of_bnf = #bd_cinfinite o #axioms o rep_bnf;
-val bd_Card_order_of_bnf = #bd_Card_order o #facts o rep_bnf;
-val bd_Cinfinite_of_bnf = #bd_Cinfinite o #facts o rep_bnf;
-val bd_Cnotzero_of_bnf = #bd_Cnotzero o #facts o rep_bnf;
-val collect_set_map_of_bnf = Lazy.force o #collect_set_map o #facts o rep_bnf;
-val in_bd_of_bnf = Lazy.force o #in_bd o #facts o rep_bnf;
-val in_cong_of_bnf = Lazy.force o #in_cong o #facts o rep_bnf;
-val in_mono_of_bnf = Lazy.force o #in_mono o #facts o rep_bnf;
-val in_rel_of_bnf = Lazy.force o #in_rel o #facts o rep_bnf;
-val map_def_of_bnf = #map_def o #defs o rep_bnf;
-val map_id0_of_bnf = #map_id0 o #axioms o rep_bnf;
-val map_id_of_bnf = Lazy.force o #map_id o #facts o rep_bnf;
-val map_comp0_of_bnf = #map_comp0 o #axioms o rep_bnf;
-val map_comp_of_bnf = Lazy.force o #map_comp o #facts o rep_bnf;
-val map_cong0_of_bnf = #map_cong0 o #axioms o rep_bnf;
-val map_cong_of_bnf = Lazy.force o #map_cong o #facts o rep_bnf;
-val map_transfer_of_bnf = Lazy.force o #map_transfer o #facts o rep_bnf;
-val le_rel_OO_of_bnf = #le_rel_OO o #axioms o rep_bnf;
-val rel_def_of_bnf = #rel_def o #defs o rep_bnf;
-val rel_eq_of_bnf = Lazy.force o #rel_eq o #facts o rep_bnf;
-val rel_flip_of_bnf = Lazy.force o #rel_flip o #facts o rep_bnf;
-val set_bd_of_bnf = #set_bd o #axioms o rep_bnf;
-val set_defs_of_bnf = #set_defs o #defs o rep_bnf;
-val set_map0_of_bnf = #set_map0 o #axioms o rep_bnf;
-val set_map_of_bnf = map Lazy.force o #set_map o #facts o rep_bnf;
-val rel_cong_of_bnf = Lazy.force o #rel_cong o #facts o rep_bnf;
-val rel_mono_of_bnf = Lazy.force o #rel_mono o #facts o rep_bnf;
-val rel_mono_strong_of_bnf = Lazy.force o #rel_mono_strong o #facts o rep_bnf;
-val rel_Grp_of_bnf = Lazy.force o #rel_Grp o #facts o rep_bnf;
-val rel_conversep_of_bnf = Lazy.force o #rel_conversep o #facts o rep_bnf;
-val rel_OO_of_bnf = Lazy.force o #rel_OO o #facts o rep_bnf;
-val rel_OO_Grp_of_bnf = #rel_OO_Grp o #axioms o rep_bnf;
-val wit_thms_of_bnf = maps #prop o wits_of_bnf;
-val wit_thmss_of_bnf = map #prop o wits_of_bnf;
-
-fun mk_bnf name T live lives lives' dead deads map sets bd axioms defs facts wits rel =
-  BNF {name = name, T = T,
-       live = live, lives = lives, lives' = lives', dead = dead, deads = deads,
-       map = map, sets = sets, bd = bd,
-       axioms = axioms, defs = defs, facts = facts,
-       nwits = length wits, wits = wits, rel = rel};
-
-fun morph_bnf phi (BNF {name = name, T = T, live = live, lives = lives, lives' = lives',
-  dead = dead, deads = deads, map = map, sets = sets, bd = bd,
-  axioms = axioms, defs = defs, facts = facts,
-  nwits = nwits, wits = wits, rel = rel}) =
-  BNF {name = Morphism.binding phi name, T = Morphism.typ phi T,
-    live = live, lives = List.map (Morphism.typ phi) lives,
-    lives' = List.map (Morphism.typ phi) lives',
-    dead = dead, deads = List.map (Morphism.typ phi) deads,
-    map = Morphism.term phi map, sets = List.map (Morphism.term phi) sets,
-    bd = Morphism.term phi bd,
-    axioms = morph_axioms phi axioms,
-    defs = morph_defs phi defs,
-    facts = morph_facts phi facts,
-    nwits = nwits,
-    wits = List.map (morph_witness phi) wits,
-    rel = Morphism.term phi rel};
-
-fun eq_bnf (BNF {T = T1, live = live1, dead = dead1, ...},
-  BNF {T = T2, live = live2, dead = dead2, ...}) =
-  Type.could_unify (T1, T2) andalso live1 = live2 andalso dead1 = dead2;
-
-structure Data = Generic_Data
-(
-  type T = bnf Symtab.table;
-  val empty = Symtab.empty;
-  val extend = I;
-  val merge = Symtab.merge eq_bnf;
-);
-
-fun bnf_of ctxt =
-  Symtab.lookup (Data.get (Context.Proof ctxt))
-  #> Option.map (morph_bnf (Morphism.transfer_morphism (Proof_Context.theory_of ctxt)));
-
-
-(* Utilities *)
-
-fun normalize_set insts instA set =
-  let
-    val (T, T') = dest_funT (fastype_of set);
-    val A = fst (Term.dest_TVar (HOLogic.dest_setT T'));
-    val params = Term.add_tvar_namesT T [];
-  in Term.subst_TVars ((A :: params) ~~ (instA :: insts)) set end;
-
-fun normalize_rel ctxt instTs instA instB rel =
-  let
-    val thy = Proof_Context.theory_of ctxt;
-    val tyenv =
-      Sign.typ_match thy (fastype_of rel, Library.foldr (op -->) (instTs, mk_pred2T instA instB))
-        Vartab.empty;
-  in Envir.subst_term (tyenv, Vartab.empty) rel end
-  handle Type.TYPE_MATCH => error "Bad relator";
-
-fun normalize_wit insts CA As wit =
-  let
-    fun strip_param (Ts, T as Type (@{type_name fun}, [T1, T2])) =
-        if Type.raw_instance (CA, T) then (Ts, T) else strip_param (T1 :: Ts, T2)
-      | strip_param x = x;
-    val (Ts, T) = strip_param ([], fastype_of wit);
-    val subst = Term.add_tvar_namesT T [] ~~ insts;
-    fun find y = find_index (fn x => x = y) As;
-  in
-    (map (find o Term.typ_subst_TVars subst) (rev Ts), Term.subst_TVars subst wit)
-  end;
-
-fun minimize_wits wits =
- let
-   fun minimize done [] = done
-     | minimize done ((I, wit) :: todo) =
-       if exists (fn (J, _) => subset (op =) (J, I)) (done @ todo)
-       then minimize done todo
-       else minimize ((I, wit) :: done) todo;
- in minimize [] wits end;
-
-fun mk_map live Ts Us t =
-  let val (Type (_, Ts0), Type (_, Us0)) = strip_typeN (live + 1) (fastype_of t) |>> List.last in
-    Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
-  end;
-
-fun mk_rel live Ts Us t =
-  let val [Type (_, Ts0), Type (_, Us0)] = binder_types (snd (strip_typeN live (fastype_of t))) in
-    Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
-  end;
-
-fun build_map_or_rel mk const of_bnf dest ctxt build_simple =
-  let
-    fun build (TU as (T, U)) =
-      if T = U then
-        const T
-      else
-        (case TU of
-          (Type (s, Ts), Type (s', Us)) =>
-          if s = s' then
-            let
-              val bnf = the (bnf_of ctxt s);
-              val live = live_of_bnf bnf;
-              val mapx = mk live Ts Us (of_bnf bnf);
-              val TUs' = map dest (fst (strip_typeN live (fastype_of mapx)));
-            in Term.list_comb (mapx, map build TUs') end
-          else
-            build_simple TU
-        | _ => build_simple TU);
-  in build end;
-
-val build_map = build_map_or_rel mk_map HOLogic.id_const map_of_bnf dest_funT;
-val build_rel = build_map_or_rel mk_rel HOLogic.eq_const rel_of_bnf dest_pred2T;
-
-fun map_flattened_map_args ctxt s map_args fs =
-  let
-    val flat_fs = flatten_type_args_of_bnf (the (bnf_of ctxt s)) Term.dummy fs;
-    val flat_fs' = map_args flat_fs;
-  in
-    permute_like (op aconv) flat_fs fs flat_fs'
-  end;
-
-
-(* Names *)
-
-val mapN = "map";
-val setN = "set";
-fun mk_setN i = setN ^ nonzero_string_of_int i;
-val bdN = "bd";
-val witN = "wit";
-fun mk_witN i = witN ^ nonzero_string_of_int i;
-val relN = "rel";
-
-val bd_card_orderN = "bd_card_order";
-val bd_cinfiniteN = "bd_cinfinite";
-val bd_Card_orderN = "bd_Card_order";
-val bd_CinfiniteN = "bd_Cinfinite";
-val bd_CnotzeroN = "bd_Cnotzero";
-val collect_set_mapN = "collect_set_map";
-val in_bdN = "in_bd";
-val in_monoN = "in_mono";
-val in_relN = "in_rel";
-val map_id0N = "map_id0";
-val map_idN = "map_id";
-val map_comp0N = "map_comp0";
-val map_compN = "map_comp";
-val map_cong0N = "map_cong0";
-val map_congN = "map_cong";
-val map_transferN = "map_transfer";
-val rel_eqN = "rel_eq";
-val rel_flipN = "rel_flip";
-val set_map0N = "set_map0";
-val set_mapN = "set_map";
-val set_bdN = "set_bd";
-val rel_GrpN = "rel_Grp";
-val rel_conversepN = "rel_conversep";
-val rel_monoN = "rel_mono"
-val rel_mono_strongN = "rel_mono_strong"
-val rel_comppN = "rel_compp";
-val rel_compp_GrpN = "rel_compp_Grp";
-
-datatype const_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline;
-
-datatype fact_policy = Dont_Note | Note_Some | Note_All;
-
-val bnf_note_all = Attrib.setup_config_bool @{binding bnf_note_all} (K false);
-val bnf_timing = Attrib.setup_config_bool @{binding bnf_timing} (K false);
-
-fun user_policy policy ctxt = if Config.get ctxt bnf_note_all then Note_All else policy;
-
-val smart_max_inline_size = 25; (*FUDGE*)
-
-fun note_bnf_thms fact_policy qualify' bnf_b bnf =
-  let
-    val axioms = axioms_of_bnf bnf;
-    val facts = facts_of_bnf bnf;
-    val wits = wits_of_bnf bnf;
-    val qualify =
-      let val (_, qs, _) = Binding.dest bnf_b;
-      in fold_rev (fn (s, mand) => Binding.qualify mand s) qs #> qualify' end;
-  in
-    (if fact_policy = Note_All then
-      let
-        val witNs = if length wits = 1 then [witN] else map mk_witN (1 upto length wits);
-        val notes =
-          [(bd_card_orderN, [#bd_card_order axioms]),
-            (bd_cinfiniteN, [#bd_cinfinite axioms]),
-            (bd_Card_orderN, [#bd_Card_order facts]),
-            (bd_CinfiniteN, [#bd_Cinfinite facts]),
-            (bd_CnotzeroN, [#bd_Cnotzero facts]),
-            (collect_set_mapN, [Lazy.force (#collect_set_map facts)]),
-            (in_bdN, [Lazy.force (#in_bd facts)]),
-            (in_monoN, [Lazy.force (#in_mono facts)]),
-            (in_relN, [Lazy.force (#in_rel facts)]),
-            (map_comp0N, [#map_comp0 axioms]),
-            (map_id0N, [#map_id0 axioms]),
-            (map_transferN, [Lazy.force (#map_transfer facts)]),
-            (rel_mono_strongN, [Lazy.force (#rel_mono_strong facts)]),
-            (set_map0N, #set_map0 axioms),
-            (set_bdN, #set_bd axioms)] @
-            (witNs ~~ wit_thmss_of_bnf bnf)
-            |> map (fn (thmN, thms) =>
-              ((qualify (Binding.qualify true (Binding.name_of bnf_b) (Binding.name thmN)), []),
-              [(thms, [])]));
-        in
-          Local_Theory.notes notes #> snd
-        end
-      else
-        I)
-    #> (if fact_policy <> Dont_Note then
-        let
-          val notes =
-            [(map_compN, [Lazy.force (#map_comp facts)], []),
-            (map_cong0N, [#map_cong0 axioms], []),
-            (map_congN, [Lazy.force (#map_cong facts)], fundefcong_attrs),
-            (map_idN, [Lazy.force (#map_id facts)], []),
-            (rel_comppN, [Lazy.force (#rel_OO facts)], []),
-            (rel_compp_GrpN, no_refl [#rel_OO_Grp axioms], []),
-            (rel_conversepN, [Lazy.force (#rel_conversep facts)], []),
-            (rel_eqN, [Lazy.force (#rel_eq facts)], []),
-            (rel_flipN, [Lazy.force (#rel_flip facts)], []),
-            (rel_GrpN, [Lazy.force (#rel_Grp facts)], []),
-            (rel_monoN, [Lazy.force (#rel_mono facts)], []),
-            (set_mapN, map Lazy.force (#set_map facts), [])]
-            |> filter_out (null o #2)
-            |> map (fn (thmN, thms, attrs) =>
-              ((qualify (Binding.qualify true (Binding.name_of bnf_b) (Binding.name thmN)),
-                attrs), [(thms, [])]));
-        in
-          Local_Theory.notes notes #> snd
-        end
-      else
-        I)
-  end;
-
-
-(* Define new BNFs *)
-
-fun define_bnf_consts const_policy fact_policy Ds_opt map_b rel_b set_bs
-  ((((((bnf_b, T_rhs), map_rhs), set_rhss), bd_rhs), wit_rhss), rel_rhs_opt) no_defs_lthy =
-  let
-    val live = length set_rhss;
-
-    val def_qualify = Binding.conceal o Binding.qualify false (Binding.name_of bnf_b);
-
-    fun mk_prefix_binding pre = Binding.prefix_name (pre ^ "_") bnf_b;
-
-    fun maybe_define user_specified (b, rhs) lthy =
-      let
-        val inline =
-          (user_specified orelse fact_policy = Dont_Note) andalso
-          (case const_policy of
-            Dont_Inline => false
-          | Hardly_Inline => Term.is_Free rhs orelse Term.is_Const rhs
-          | Smart_Inline => Term.size_of_term rhs <= smart_max_inline_size
-          | Do_Inline => true)
-      in
-        if inline then
-          ((rhs, Drule.reflexive_thm), lthy)
-        else
-          let val b = b () in
-            apfst (apsnd snd) (Local_Theory.define ((b, NoSyn), ((Thm.def_binding b, []), rhs))
-              lthy)
-          end
-      end;
-
-    fun maybe_restore lthy_old lthy =
-      lthy |> not (pointer_eq (lthy_old, lthy)) ? Local_Theory.restore;
-
-    val map_bind_def =
-      (fn () => def_qualify (if Binding.is_empty map_b then mk_prefix_binding mapN else map_b),
-         map_rhs);
-    val set_binds_defs =
-      let
-        fun set_name i get_b =
-          (case try (nth set_bs) (i - 1) of
-            SOME b => if Binding.is_empty b then get_b else K b
-          | NONE => get_b) #> def_qualify;
-        val bs = if live = 1 then [set_name 1 (fn () => mk_prefix_binding setN)]
-          else map (fn i => set_name i (fn () => mk_prefix_binding (mk_setN i))) (1 upto live);
-      in bs ~~ set_rhss end;
-    val bd_bind_def = (fn () => def_qualify (mk_prefix_binding bdN), bd_rhs);
-
-    val ((((bnf_map_term, raw_map_def),
-      (bnf_set_terms, raw_set_defs)),
-      (bnf_bd_term, raw_bd_def)), (lthy, lthy_old)) =
-        no_defs_lthy
-        |> maybe_define true map_bind_def
-        ||>> apfst split_list o fold_map (maybe_define true) set_binds_defs
-        ||>> maybe_define true bd_bind_def
-        ||> `(maybe_restore no_defs_lthy);
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-
-
-    val bnf_map_def = Morphism.thm phi raw_map_def;
-    val bnf_set_defs = map (Morphism.thm phi) raw_set_defs;
-    val bnf_bd_def = Morphism.thm phi raw_bd_def;
-
-    val bnf_map = Morphism.term phi bnf_map_term;
-
-    (*TODO: handle errors*)
-    (*simple shape analysis of a map function*)
-    val ((alphas, betas), (Calpha, _)) =
-      fastype_of bnf_map
-      |> strip_typeN live
-      |>> map_split dest_funT
-      ||> dest_funT
-      handle TYPE _ => error "Bad map function";
-
-    val Calpha_params = map TVar (Term.add_tvarsT Calpha []);
-
-    val bnf_T = Morphism.typ phi T_rhs;
-    val bad_args = Term.add_tfreesT bnf_T [];
-    val _ = if null bad_args then () else error ("Locally fixed type arguments " ^
-      commas_quote (map (Syntax.string_of_typ no_defs_lthy o TFree) bad_args));
-
-    val bnf_sets =
-      map2 (normalize_set Calpha_params) alphas (map (Morphism.term phi) bnf_set_terms);
-    val bnf_bd =
-      Term.subst_TVars (Term.add_tvar_namesT bnf_T [] ~~ Calpha_params)
-        (Morphism.term phi bnf_bd_term);
-
-    (*TODO: assert Ds = (TVars of bnf_map) \ (alphas @ betas) as sets*)
-    val deads = (case Ds_opt of
-      NONE => subtract (op =) (alphas @ betas) (map TVar (Term.add_tvars bnf_map []))
-    | SOME Ds => map (Morphism.typ phi) Ds);
-
-    (*TODO: further checks of type of bnf_map*)
-    (*TODO: check types of bnf_sets*)
-    (*TODO: check type of bnf_bd*)
-    (*TODO: check type of bnf_rel*)
-
-    fun mk_bnf_map Ds As' Bs' =
-      Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As') @ (betas ~~ Bs')) bnf_map;
-    fun mk_bnf_t Ds As' = Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As'));
-    fun mk_bnf_T Ds As' = Term.typ_subst_atomic ((deads ~~ Ds) @ (alphas ~~ As'));
-
-    val (((As, Bs), Ds), names_lthy) = lthy
-      |> mk_TFrees live
-      ||>> mk_TFrees live
-      ||>> mk_TFrees (length deads);
-    val RTs = map2 (curry HOLogic.mk_prodT) As Bs;
-    val pred2RTs = map2 mk_pred2T As Bs;
-    val (Rs, Rs') = names_lthy |> mk_Frees' "R" pred2RTs |> fst
-    val CA = mk_bnf_T Ds As Calpha;
-    val CR = mk_bnf_T Ds RTs Calpha;
-    val setRs =
-      map3 (fn R => fn T => fn U =>
-          HOLogic.Collect_const (HOLogic.mk_prodT (T, U)) $ HOLogic.mk_split R) Rs As Bs;
-
-    (*Grp (in (Collect (split R1) .. Collect (split Rn))) (map fst .. fst)^--1 OO
-      Grp (in (Collect (split R1) .. Collect (split Rn))) (map snd .. snd)*)
-    val OO_Grp =
-      let
-        val map1 = Term.list_comb (mk_bnf_map Ds RTs As, map fst_const RTs);
-        val map2 = Term.list_comb (mk_bnf_map Ds RTs Bs, map snd_const RTs);
-        val bnf_in = mk_in setRs (map (mk_bnf_t Ds RTs) bnf_sets) CR;
-      in
-        mk_rel_compp (mk_conversep (mk_Grp bnf_in map1), mk_Grp bnf_in map2)
-        |> fold_rev Term.absfree Rs'
-      end;
-
-    val rel_rhs = the_default OO_Grp rel_rhs_opt;
-
-    val rel_bind_def =
-      (fn () => def_qualify (if Binding.is_empty rel_b then mk_prefix_binding relN else rel_b),
-         rel_rhs);
-
-    val wit_rhss =
-      if null wit_rhss then
-        [fold_rev Term.absdummy As (Term.list_comb (mk_bnf_map Ds As As,
-          map2 (fn T => fn i => Term.absdummy T (Bound i)) As (live downto 1)) $
-          Const (@{const_name undefined}, CA))]
-      else wit_rhss;
-    val nwits = length wit_rhss;
-    val wit_binds_defs =
-      let
-        val bs = if nwits = 1 then [fn () => def_qualify (mk_prefix_binding witN)]
-          else map (fn i => fn () => def_qualify (mk_prefix_binding (mk_witN i))) (1 upto nwits);
-      in bs ~~ wit_rhss end;
-
-    val (((bnf_rel_term, raw_rel_def), (bnf_wit_terms, raw_wit_defs)), (lthy, lthy_old)) =
-      lthy
-      |> maybe_define (is_some rel_rhs_opt) rel_bind_def
-      ||>> apfst split_list o fold_map (maybe_define (not (null wit_rhss))) wit_binds_defs
-      ||> `(maybe_restore lthy);
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-    val bnf_rel_def = Morphism.thm phi raw_rel_def;
-    val bnf_rel = Morphism.term phi bnf_rel_term;
-    fun mk_bnf_rel Ds As' Bs' =
-      normalize_rel lthy (map2 mk_pred2T As' Bs') (mk_bnf_T Ds As' Calpha) (mk_bnf_T Ds Bs' Calpha)
-        bnf_rel;
-
-    val bnf_wit_defs = map (Morphism.thm phi) raw_wit_defs;
-    val bnf_wits =
-      map (normalize_wit Calpha_params Calpha alphas o Morphism.term phi) bnf_wit_terms;
-
-    fun mk_OO_Grp Ds' As' Bs' =
-      Term.subst_atomic_types ((Ds ~~ Ds') @ (As ~~ As') @ (Bs ~~ Bs')) OO_Grp;
-  in
-    (((alphas, betas, deads, Calpha),
-     (bnf_map, bnf_sets, bnf_bd, bnf_wits, bnf_rel),
-     (bnf_map_def, bnf_set_defs, bnf_bd_def, bnf_wit_defs, bnf_rel_def),
-     (mk_bnf_map, mk_bnf_t, mk_bnf_T, mk_bnf_rel, mk_OO_Grp)), lthy)
-  end;
-
-fun prepare_def const_policy mk_fact_policy qualify prep_typ prep_term Ds_opt map_b rel_b set_bs
-  ((((((raw_bnf_b, raw_bnf_T), raw_map), raw_sets), raw_bd), raw_wits), raw_rel_opt)
-  no_defs_lthy =
-  let
-    val fact_policy = mk_fact_policy no_defs_lthy;
-    val bnf_b = qualify raw_bnf_b;
-    val live = length raw_sets;
-
-    val T_rhs = prep_typ no_defs_lthy raw_bnf_T;
-    val map_rhs = prep_term no_defs_lthy raw_map;
-    val set_rhss = map (prep_term no_defs_lthy) raw_sets;
-    val bd_rhs = prep_term no_defs_lthy raw_bd;
-    val wit_rhss = map (prep_term no_defs_lthy) raw_wits;
-    val rel_rhs_opt = Option.map (prep_term no_defs_lthy) raw_rel_opt;
-
-    fun err T =
-      error ("Trying to register the type " ^ quote (Syntax.string_of_typ no_defs_lthy T) ^
-        " as unnamed BNF");
-
-    val (bnf_b, key) =
-      if Binding.eq_name (bnf_b, Binding.empty) then
-        (case T_rhs of
-          Type (C, Ts) => if forall (can dest_TFree) Ts
-            then (Binding.qualified_name C, C) else err T_rhs
-        | T => err T)
-      else (bnf_b, Local_Theory.full_name no_defs_lthy bnf_b);
-
-    val (((alphas, betas, deads, Calpha),
-     (bnf_map, bnf_sets, bnf_bd, bnf_wits, bnf_rel),
-     (bnf_map_def, bnf_set_defs, bnf_bd_def, bnf_wit_defs, bnf_rel_def),
-     (mk_bnf_map_Ds, mk_bnf_t_Ds, mk_bnf_T_Ds, _, mk_OO_Grp)), lthy) =
-       define_bnf_consts const_policy fact_policy Ds_opt map_b rel_b set_bs
-         ((((((bnf_b, T_rhs), map_rhs), set_rhss), bd_rhs), wit_rhss), rel_rhs_opt) no_defs_lthy;
-
-    val dead = length deads;
-
-    val ((((((As', Bs'), Cs), Ds), B1Ts), B2Ts), (Ts, T)) = lthy
-      |> mk_TFrees live
-      ||>> mk_TFrees live
-      ||>> mk_TFrees live
-      ||>> mk_TFrees dead
-      ||>> mk_TFrees live
-      ||>> mk_TFrees live
-      ||> fst o mk_TFrees 1
-      ||> the_single
-      ||> `(replicate live);
-
-    val mk_bnf_map = mk_bnf_map_Ds Ds;
-    val mk_bnf_t = mk_bnf_t_Ds Ds;
-    val mk_bnf_T = mk_bnf_T_Ds Ds;
-
-    val pred2RTs = map2 mk_pred2T As' Bs';
-    val pred2RTsAsCs = map2 mk_pred2T As' Cs;
-    val pred2RTsBsCs = map2 mk_pred2T Bs' Cs;
-    val pred2RT's = map2 mk_pred2T Bs' As';
-    val self_pred2RTs = map2 mk_pred2T As' As';
-    val transfer_domRTs = map2 mk_pred2T As' B1Ts;
-    val transfer_ranRTs = map2 mk_pred2T Bs' B2Ts;
-
-    val CA' = mk_bnf_T As' Calpha;
-    val CB' = mk_bnf_T Bs' Calpha;
-    val CC' = mk_bnf_T Cs Calpha;
-    val CB1 = mk_bnf_T B1Ts Calpha;
-    val CB2 = mk_bnf_T B2Ts Calpha;
-
-    val bnf_map_AsAs = mk_bnf_map As' As';
-    val bnf_map_AsBs = mk_bnf_map As' Bs';
-    val bnf_map_AsCs = mk_bnf_map As' Cs;
-    val bnf_map_BsCs = mk_bnf_map Bs' Cs;
-    val bnf_sets_As = map (mk_bnf_t As') bnf_sets;
-    val bnf_sets_Bs = map (mk_bnf_t Bs') bnf_sets;
-    val bnf_bd_As = mk_bnf_t As' bnf_bd;
-    fun mk_bnf_rel RTs CA CB = normalize_rel lthy RTs CA CB bnf_rel;
-
-    val pre_names_lthy = lthy;
-    val (((((((((((((((fs, gs), hs), x), y), zs), ys), As),
-      As_copy), bs), Rs), Rs_copy), Ss),
-      transfer_domRs), transfer_ranRs), names_lthy) = pre_names_lthy
-      |> mk_Frees "f" (map2 (curry op -->) As' Bs')
-      ||>> mk_Frees "g" (map2 (curry op -->) Bs' Cs)
-      ||>> mk_Frees "h" (map2 (curry op -->) As' Ts)
-      ||>> yield_singleton (mk_Frees "x") CA'
-      ||>> yield_singleton (mk_Frees "y") CB'
-      ||>> mk_Frees "z" As'
-      ||>> mk_Frees "y" Bs'
-      ||>> mk_Frees "A" (map HOLogic.mk_setT As')
-      ||>> mk_Frees "A" (map HOLogic.mk_setT As')
-      ||>> mk_Frees "b" As'
-      ||>> mk_Frees "R" pred2RTs
-      ||>> mk_Frees "R" pred2RTs
-      ||>> mk_Frees "S" pred2RTsBsCs
-      ||>> mk_Frees "R" transfer_domRTs
-      ||>> mk_Frees "S" transfer_ranRTs;
-
-    val fs_copy = map2 (retype_free o fastype_of) fs gs;
-    val x_copy = retype_free CA' y;
-
-    val rel = mk_bnf_rel pred2RTs CA' CB';
-    val relAsAs = mk_bnf_rel self_pred2RTs CA' CA';
-    val bnf_wit_As = map (apsnd (mk_bnf_t As')) bnf_wits;
-
-    val map_id0_goal =
-      let val bnf_map_app_id = Term.list_comb (bnf_map_AsAs, map HOLogic.id_const As') in
-        mk_Trueprop_eq (bnf_map_app_id, HOLogic.id_const CA')
-      end;
-
-    val map_comp0_goal =
-      let
-        val bnf_map_app_comp = Term.list_comb (bnf_map_AsCs, map2 (curry HOLogic.mk_comp) gs fs);
-        val comp_bnf_map_app = HOLogic.mk_comp
-          (Term.list_comb (bnf_map_BsCs, gs), Term.list_comb (bnf_map_AsBs, fs));
-      in
-        fold_rev Logic.all (fs @ gs) (mk_Trueprop_eq (bnf_map_app_comp, comp_bnf_map_app))
-      end;
-
-    fun mk_map_cong_prem x z set f f_copy =
-      Logic.all z (Logic.mk_implies
-        (HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set $ x)),
-        mk_Trueprop_eq (f $ z, f_copy $ z)));
-
-    val map_cong0_goal =
-      let
-        val prems = map4 (mk_map_cong_prem x) zs bnf_sets_As fs fs_copy;
-        val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x,
-          Term.list_comb (bnf_map_AsBs, fs_copy) $ x);
-      in
-        fold_rev Logic.all (x :: fs @ fs_copy) (Logic.list_implies (prems, eq))
-      end;
-
-    val set_map0s_goal =
-      let
-        fun mk_goal setA setB f =
-          let
-            val set_comp_map =
-              HOLogic.mk_comp (setB, Term.list_comb (bnf_map_AsBs, fs));
-            val image_comp_set = HOLogic.mk_comp (mk_image f, setA);
-          in
-            fold_rev Logic.all fs (mk_Trueprop_eq (set_comp_map, image_comp_set))
-          end;
-      in
-        map3 mk_goal bnf_sets_As bnf_sets_Bs fs
-      end;
-
-    val card_order_bd_goal = HOLogic.mk_Trueprop (mk_card_order bnf_bd_As);
-
-    val cinfinite_bd_goal = HOLogic.mk_Trueprop (mk_cinfinite bnf_bd_As);
-
-    val set_bds_goal =
-      let
-        fun mk_goal set =
-          Logic.all x (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (set $ x)) bnf_bd_As));
-      in
-        map mk_goal bnf_sets_As
-      end;
-
-    val relAsCs = mk_bnf_rel pred2RTsAsCs CA' CC';
-    val relBsCs = mk_bnf_rel pred2RTsBsCs CB' CC';
-    val rel_OO_lhs = Term.list_comb (relAsCs, map2 (curry mk_rel_compp) Rs Ss);
-    val rel_OO_rhs = mk_rel_compp (Term.list_comb (rel, Rs), Term.list_comb (relBsCs, Ss));
-    val le_rel_OO_goal =
-      fold_rev Logic.all (Rs @ Ss) (HOLogic.mk_Trueprop (mk_leq rel_OO_rhs rel_OO_lhs));
-
-    val rel_OO_Grp_goal = fold_rev Logic.all Rs (mk_Trueprop_eq (Term.list_comb (rel, Rs),
-      Term.list_comb (mk_OO_Grp Ds As' Bs', Rs)));
-
-    val goals = zip_axioms map_id0_goal map_comp0_goal map_cong0_goal set_map0s_goal
-      card_order_bd_goal cinfinite_bd_goal set_bds_goal le_rel_OO_goal rel_OO_Grp_goal;
-
-    fun mk_wit_goals (I, wit) =
-      let
-        val xs = map (nth bs) I;
-        fun wit_goal i =
-          let
-            val z = nth zs i;
-            val set_wit = nth bnf_sets_As i $ Term.list_comb (wit, xs);
-            val concl = HOLogic.mk_Trueprop
-              (if member (op =) I i then HOLogic.mk_eq (z, nth bs i)
-              else @{term False});
-          in
-            fold_rev Logic.all (z :: xs)
-              (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set_wit)), concl))
-          end;
-      in
-        map wit_goal (0 upto live - 1)
-      end;
-
-    val triv_wit_tac = mk_trivial_wit_tac bnf_wit_defs;
-
-    val wit_goalss =
-      (if null raw_wits then SOME triv_wit_tac else NONE, map mk_wit_goals bnf_wit_As);
-
-    fun after_qed mk_wit_thms thms lthy =
-      let
-        val (axioms, nontriv_wit_thms) = apfst (mk_axioms live) (chop (length goals) thms);
-
-        val bd_Card_order = #bd_card_order axioms RS @{thm conjunct2[OF card_order_on_Card_order]};
-        val bd_Cinfinite = @{thm conjI} OF [#bd_cinfinite axioms, bd_Card_order];
-        val bd_Cnotzero = bd_Cinfinite RS @{thm Cinfinite_Cnotzero};
-
-        fun mk_collect_set_map () =
-          let
-            val defT = mk_bnf_T Ts Calpha --> HOLogic.mk_setT T;
-            val collect_map = HOLogic.mk_comp
-              (mk_collect (map (mk_bnf_t Ts) bnf_sets) defT,
-              Term.list_comb (mk_bnf_map As' Ts, hs));
-            val image_collect = mk_collect
-              (map2 (fn h => fn set => HOLogic.mk_comp (mk_image h, set)) hs bnf_sets_As)
-              defT;
-            (*collect {set1 ... setm} o map f1 ... fm = collect {f1` o set1 ... fm` o setm}*)
-            val goal = fold_rev Logic.all hs (mk_Trueprop_eq (collect_map, image_collect));
-          in
-            Goal.prove_sorry lthy [] [] goal (K (mk_collect_set_map_tac (#set_map0 axioms)))
-            |> Thm.close_derivation
-          end;
-
-        val collect_set_map = Lazy.lazy mk_collect_set_map;
-
-        fun mk_in_mono () =
-          let
-            val prems_mono = map2 (HOLogic.mk_Trueprop oo mk_leq) As As_copy;
-            val in_mono_goal =
-              fold_rev Logic.all (As @ As_copy)
-                (Logic.list_implies (prems_mono, HOLogic.mk_Trueprop
-                  (mk_leq (mk_in As bnf_sets_As CA') (mk_in As_copy bnf_sets_As CA'))));
-          in
-            Goal.prove_sorry lthy [] [] in_mono_goal (K (mk_in_mono_tac live))
-            |> Thm.close_derivation
-          end;
-
-        val in_mono = Lazy.lazy mk_in_mono;
-
-        fun mk_in_cong () =
-          let
-            val prems_cong = map2 (curry mk_Trueprop_eq) As As_copy;
-            val in_cong_goal =
-              fold_rev Logic.all (As @ As_copy)
-                (Logic.list_implies (prems_cong,
-                  mk_Trueprop_eq (mk_in As bnf_sets_As CA', mk_in As_copy bnf_sets_As CA')));
-          in
-            Goal.prove_sorry lthy [] [] in_cong_goal
-              (K ((TRY o hyp_subst_tac lthy THEN' rtac refl) 1))
-            |> Thm.close_derivation
-          end;
-
-        val in_cong = Lazy.lazy mk_in_cong;
-
-        val map_id = Lazy.lazy (fn () => mk_map_id (#map_id0 axioms));
-        val map_comp = Lazy.lazy (fn () => mk_map_comp (#map_comp0 axioms));
-
-        fun mk_map_cong () =
-          let
-            val prem0 = mk_Trueprop_eq (x, x_copy);
-            val prems = map4 (mk_map_cong_prem x_copy) zs bnf_sets_As fs fs_copy;
-            val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x,
-              Term.list_comb (bnf_map_AsBs, fs_copy) $ x_copy);
-            val goal = fold_rev Logic.all (x :: x_copy :: fs @ fs_copy)
-              (Logic.list_implies (prem0 :: prems, eq));
-          in
-            Goal.prove_sorry lthy [] [] goal (fn _ => mk_map_cong_tac lthy (#map_cong0 axioms))
-            |> Thm.close_derivation
-          end;
-
-        val map_cong = Lazy.lazy mk_map_cong;
-
-        val set_map = map (fn thm => Lazy.lazy (fn () => mk_set_map thm)) (#set_map0 axioms);
-
-        val wit_thms =
-          if null nontriv_wit_thms then mk_wit_thms (map Lazy.force set_map) else nontriv_wit_thms;
-
-        fun mk_in_bd () =
-          let
-            val bdT = fst (dest_relT (fastype_of bnf_bd_As));
-            val bdTs = replicate live bdT;
-            val bd_bnfT = mk_bnf_T bdTs Calpha;
-            val surj_imp_ordLeq_inst = (if live = 0 then TrueI else
-              let
-                val ranTs = map (fn AT => mk_sumT (AT, HOLogic.unitT)) As';
-                val funTs = map (fn T => bdT --> T) ranTs;
-                val ran_bnfT = mk_bnf_T ranTs Calpha;
-                val (revTs, Ts) = `rev (bd_bnfT :: funTs);
-                val cTs = map (SOME o certifyT lthy) [ran_bnfT, Library.foldr1 HOLogic.mk_prodT Ts];
-                val tinst = fold (fn T => fn t => HOLogic.mk_split (Term.absdummy T t)) (tl revTs)
-                  (Term.absdummy (hd revTs) (Term.list_comb (mk_bnf_map bdTs ranTs,
-                    map Bound (live - 1 downto 0)) $ Bound live));
-                val cts = [NONE, SOME (certify lthy tinst)];
-              in
-                Drule.instantiate' cTs cts @{thm surj_imp_ordLeq}
-              end);
-            val bd = mk_cexp
-              (if live = 0 then ctwo
-                else mk_csum (Library.foldr1 (uncurry mk_csum) (map mk_card_of As)) ctwo)
-              (mk_csum bnf_bd_As (mk_card_of (HOLogic.mk_UNIV bd_bnfT)));
-            val in_bd_goal =
-              fold_rev Logic.all As
-                (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (mk_in As bnf_sets_As CA')) bd));
-          in
-            Goal.prove_sorry lthy [] [] in_bd_goal
-              (mk_in_bd_tac live surj_imp_ordLeq_inst
-                (Lazy.force map_comp) (Lazy.force map_id) (#map_cong0 axioms)
-                (map Lazy.force set_map) (#set_bd axioms) (#bd_card_order axioms)
-                bd_Card_order bd_Cinfinite bd_Cnotzero)
-            |> Thm.close_derivation
-          end;
-
-        val in_bd = Lazy.lazy mk_in_bd;
-
-        val rel_OO_Grp = #rel_OO_Grp axioms;
-        val rel_OO_Grps = no_refl [rel_OO_Grp];
-
-        fun mk_rel_Grp () =
-          let
-            val lhs = Term.list_comb (rel, map2 mk_Grp As fs);
-            val rhs = mk_Grp (mk_in As bnf_sets_As CA') (Term.list_comb (bnf_map_AsBs, fs));
-            val goal = fold_rev Logic.all (As @ fs) (mk_Trueprop_eq (lhs, rhs));
-          in
-            Goal.prove_sorry lthy [] [] goal
-              (mk_rel_Grp_tac rel_OO_Grps (#map_id0 axioms) (#map_cong0 axioms) (Lazy.force map_id)
-                (Lazy.force map_comp) (map Lazy.force set_map))
-            |> Thm.close_derivation
-          end;
-
-        val rel_Grp = Lazy.lazy mk_rel_Grp;
-
-        fun mk_rel_prems f = map2 (HOLogic.mk_Trueprop oo f) Rs Rs_copy
-        fun mk_rel_concl f = HOLogic.mk_Trueprop
-          (f (Term.list_comb (rel, Rs), Term.list_comb (rel, Rs_copy)));
-
-        fun mk_rel_mono () =
-          let
-            val mono_prems = mk_rel_prems mk_leq;
-            val mono_concl = mk_rel_concl (uncurry mk_leq);
-          in
-            Goal.prove_sorry lthy [] []
-              (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (mono_prems, mono_concl)))
-              (K (mk_rel_mono_tac rel_OO_Grps (Lazy.force in_mono)))
-            |> Thm.close_derivation
-          end;
-
-        fun mk_rel_cong () =
-          let
-            val cong_prems = mk_rel_prems (curry HOLogic.mk_eq);
-            val cong_concl = mk_rel_concl HOLogic.mk_eq;
-          in
-            Goal.prove_sorry lthy [] []
-              (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (cong_prems, cong_concl)))
-              (fn _ => (TRY o hyp_subst_tac lthy THEN' rtac refl) 1)
-            |> Thm.close_derivation
-          end;
-
-        val rel_mono = Lazy.lazy mk_rel_mono;
-        val rel_cong = Lazy.lazy mk_rel_cong;
-
-        fun mk_rel_eq () =
-          Goal.prove_sorry lthy [] []
-            (mk_Trueprop_eq (Term.list_comb (relAsAs, map HOLogic.eq_const As'),
-              HOLogic.eq_const CA'))
-            (K (mk_rel_eq_tac live (Lazy.force rel_Grp) (Lazy.force rel_cong) (#map_id0 axioms)))
-          |> Thm.close_derivation;
-
-        val rel_eq = Lazy.lazy mk_rel_eq;
-
-        fun mk_rel_conversep () =
-          let
-            val relBsAs = mk_bnf_rel pred2RT's CB' CA';
-            val lhs = Term.list_comb (relBsAs, map mk_conversep Rs);
-            val rhs = mk_conversep (Term.list_comb (rel, Rs));
-            val le_goal = fold_rev Logic.all Rs (HOLogic.mk_Trueprop (mk_leq lhs rhs));
-            val le_thm = Goal.prove_sorry lthy [] [] le_goal
-              (mk_rel_conversep_le_tac rel_OO_Grps (Lazy.force rel_eq) (#map_cong0 axioms)
-                (Lazy.force map_comp) (map Lazy.force set_map))
-              |> Thm.close_derivation
-            val goal = fold_rev Logic.all Rs (mk_Trueprop_eq (lhs, rhs));
-          in
-            Goal.prove_sorry lthy [] [] goal
-              (K (mk_rel_conversep_tac le_thm (Lazy.force rel_mono)))
-            |> Thm.close_derivation
-          end;
-
-        val rel_conversep = Lazy.lazy mk_rel_conversep;
-
-        fun mk_rel_OO () =
-          Goal.prove_sorry lthy [] []
-            (fold_rev Logic.all (Rs @ Ss) (HOLogic.mk_Trueprop (mk_leq rel_OO_lhs rel_OO_rhs)))
-            (mk_rel_OO_le_tac rel_OO_Grps (Lazy.force rel_eq) (#map_cong0 axioms)
-              (Lazy.force map_comp) (map Lazy.force set_map))
-          |> Thm.close_derivation
-          |> (fn thm => @{thm antisym} OF [thm, #le_rel_OO axioms]);
-
-        val rel_OO = Lazy.lazy mk_rel_OO;
-
-        fun mk_in_rel () = trans OF [rel_OO_Grp, @{thm OO_Grp_alt}] RS @{thm predicate2_eqD};
-
-        val in_rel = Lazy.lazy mk_in_rel;
-
-        fun mk_rel_flip () =
-          let
-            val rel_conversep_thm = Lazy.force rel_conversep;
-            val cts = map (SOME o certify lthy) Rs;
-            val rel_conversep_thm' = cterm_instantiate_pos cts rel_conversep_thm;
-          in
-            unfold_thms lthy @{thms conversep_iff} (rel_conversep_thm' RS @{thm predicate2_eqD})
-            |> singleton (Proof_Context.export names_lthy pre_names_lthy)
-          end;
-
-        val rel_flip = Lazy.lazy mk_rel_flip;
-
-        fun mk_rel_mono_strong () =
-          let
-            fun mk_prem setA setB R S a b =
-              HOLogic.mk_Trueprop
-                (mk_Ball (setA $ x) (Term.absfree (dest_Free a)
-                  (mk_Ball (setB $ y) (Term.absfree (dest_Free b)
-                    (HOLogic.mk_imp (R $ a $ b, S $ a $ b))))));
-            val prems = HOLogic.mk_Trueprop (Term.list_comb (rel, Rs) $ x $ y) :: 
-              map6 mk_prem bnf_sets_As bnf_sets_Bs Rs Rs_copy zs ys;
-            val concl = HOLogic.mk_Trueprop (Term.list_comb (rel, Rs_copy) $ x $ y);
-          in
-            Goal.prove_sorry lthy [] []
-              (fold_rev Logic.all (x :: y :: Rs @ Rs_copy) (Logic.list_implies (prems, concl)))
-              (mk_rel_mono_strong_tac (Lazy.force in_rel) (map Lazy.force set_map))
-            |> Thm.close_derivation
-          end;
-
-        val rel_mono_strong = Lazy.lazy mk_rel_mono_strong;
-
-        fun mk_map_transfer () =
-          let
-            val rels = map2 mk_fun_rel transfer_domRs transfer_ranRs;
-            val rel = mk_fun_rel
-              (Term.list_comb (mk_bnf_rel transfer_domRTs CA' CB1, transfer_domRs))
-              (Term.list_comb (mk_bnf_rel transfer_ranRTs CB' CB2, transfer_ranRs));
-            val concl = HOLogic.mk_Trueprop
-              (fold_rev mk_fun_rel rels rel $ bnf_map_AsBs $ mk_bnf_map B1Ts B2Ts);
-          in
-            Goal.prove_sorry lthy [] []
-              (fold_rev Logic.all (transfer_domRs @ transfer_ranRs) concl)
-              (mk_map_transfer_tac (Lazy.force rel_mono) (Lazy.force in_rel)
-                (map Lazy.force set_map) (#map_cong0 axioms) (Lazy.force map_comp))
-            |> Thm.close_derivation
-          end;
-
-        val map_transfer = Lazy.lazy mk_map_transfer;
-
-        val defs = mk_defs bnf_map_def bnf_set_defs bnf_rel_def;
-
-        val facts = mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_map in_bd in_cong
-          in_mono in_rel map_comp map_cong map_id map_transfer rel_eq rel_flip set_map
-          rel_cong rel_mono rel_mono_strong rel_Grp rel_conversep rel_OO;
-
-        val wits = map2 mk_witness bnf_wits wit_thms;
-
-        val bnf_rel =
-          Term.subst_atomic_types ((Ds ~~ deads) @ (As' ~~ alphas) @ (Bs' ~~ betas)) rel;
-
-        val bnf = mk_bnf bnf_b Calpha live alphas betas dead deads bnf_map bnf_sets bnf_bd axioms
-          defs facts wits bnf_rel;
-      in
-        (bnf, lthy |> note_bnf_thms fact_policy qualify bnf_b bnf)
-      end;
-
-    val one_step_defs =
-      no_reflexive (bnf_map_def :: bnf_bd_def :: bnf_set_defs @ bnf_wit_defs @ [bnf_rel_def]);
-  in
-    (key, goals, wit_goalss, after_qed, lthy, one_step_defs)
-  end;
-
-fun register_bnf key (bnf, lthy) =
-  (bnf, Local_Theory.declaration {syntax = false, pervasive = true}
-    (fn phi => Data.map (Symtab.default (key, morph_bnf phi bnf))) lthy);
-
-fun bnf_def const_policy fact_policy qualify tacs wit_tac Ds map_b rel_b set_bs =
-  (fn (_, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, one_step_defs) =>
-  let
-    fun mk_wits_tac set_maps =
-      K (TRYALL Goal.conjunction_tac) THEN'
-      (case triv_tac_opt of
-        SOME tac => tac set_maps
-      | NONE => fn {context = ctxt, prems} =>
-          unfold_thms_tac ctxt one_step_defs THEN wit_tac {context = ctxt, prems = prems});
-    val wit_goals = map Logic.mk_conjunction_balanced wit_goalss;
-    fun mk_wit_thms set_maps =
-      Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals) (mk_wits_tac set_maps)
-        |> Conjunction.elim_balanced (length wit_goals)
-        |> map2 (Conjunction.elim_balanced o length) wit_goalss
-        |> map (map (Thm.close_derivation o Thm.forall_elim_vars 0));
-  in
-    map2 (Thm.close_derivation oo Goal.prove_sorry lthy [] [])
-      goals (map (fn tac => fn {context = ctxt, prems} =>
-        unfold_thms_tac ctxt one_step_defs THEN tac {context = ctxt, prems = prems}) tacs)
-    |> (fn thms => after_qed mk_wit_thms (map single thms) lthy)
-  end) oo prepare_def const_policy fact_policy qualify (K I) (K I) Ds map_b rel_b set_bs;
-
-val bnf_cmd = (fn (key, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, defs) =>
-  let
-    val wit_goals = map Logic.mk_conjunction_balanced wit_goalss;
-    fun mk_triv_wit_thms tac set_maps =
-      Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals)
-        (K (TRYALL Goal.conjunction_tac) THEN' tac set_maps)
-        |> Conjunction.elim_balanced (length wit_goals)
-        |> map2 (Conjunction.elim_balanced o length) wit_goalss
-        |> map (map (Thm.close_derivation o Thm.forall_elim_vars 0));
-    val (mk_wit_thms, nontriv_wit_goals) = 
-      (case triv_tac_opt of
-        NONE => (fn _ => [], map (map (rpair [])) wit_goalss)
-      | SOME tac => (mk_triv_wit_thms tac, []));
-  in
-    Proof.unfolding ([[(defs, [])]])
-      (Proof.theorem NONE (snd o register_bnf key oo after_qed mk_wit_thms)
-        (map (single o rpair []) goals @ nontriv_wit_goals) lthy)
-  end) oo prepare_def Do_Inline (user_policy Note_Some) I Syntax.read_typ Syntax.read_term NONE
-    Binding.empty Binding.empty [];
-
-fun print_bnfs ctxt =
-  let
-    fun pretty_set sets i = Pretty.block
-      [Pretty.str (mk_setN (i + 1) ^ ":"), Pretty.brk 1,
-          Pretty.quote (Syntax.pretty_term ctxt (nth sets i))];
-
-    fun pretty_bnf (key, BNF {T = T, map = map, sets = sets, bd = bd,
-      live = live, lives = lives, dead = dead, deads = deads, ...}) =
-      Pretty.big_list
-        (Pretty.string_of (Pretty.block [Pretty.str key, Pretty.str ":", Pretty.brk 1,
-          Pretty.quote (Syntax.pretty_typ ctxt T)]))
-        ([Pretty.block [Pretty.str "live:", Pretty.brk 1, Pretty.str (string_of_int live),
-            Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) lives)],
-          Pretty.block [Pretty.str "dead:", Pretty.brk 1, Pretty.str (string_of_int dead),
-            Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) deads)],
-          Pretty.block [Pretty.str (mapN ^ ":"), Pretty.brk 1,
-            Pretty.quote (Syntax.pretty_term ctxt map)]] @
-          List.map (pretty_set sets) (0 upto length sets - 1) @
-          [Pretty.block [Pretty.str (bdN ^ ":"), Pretty.brk 1,
-            Pretty.quote (Syntax.pretty_term ctxt bd)]]);
-  in
-    Pretty.big_list "BNFs:" (map pretty_bnf (Symtab.dest (Data.get (Context.Proof ctxt))))
-    |> Pretty.writeln
-  end;
-
-val _ =
-  Outer_Syntax.improper_command @{command_spec "print_bnfs"}
-    "print all bounded natural functors"
-    (Scan.succeed (Toplevel.keep (print_bnfs o Toplevel.context_of)));
-
-val _ =
-  Outer_Syntax.local_theory_to_proof @{command_spec "bnf"}
-    "register a type as a bounded natural functor"
-    (parse_opt_binding_colon -- Parse.typ --|
-       (Parse.reserved "map" -- @{keyword ":"}) -- Parse.term --
-       (Scan.option ((Parse.reserved "sets" -- @{keyword ":"}) |--
-         Scan.repeat1 (Scan.unless (Parse.reserved "bd") Parse.term)) >> the_default []) --|
-       (Parse.reserved "bd" -- @{keyword ":"}) -- Parse.term --
-       (Scan.option ((Parse.reserved "wits" -- @{keyword ":"}) |--
-         Scan.repeat1 (Scan.unless (Parse.reserved "rel") Parse.term)) >> the_default []) --
-       Scan.option ((Parse.reserved "rel" -- @{keyword ":"}) |-- Parse.term)
-       >> bnf_cmd);
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_def_tactics.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,284 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_def_tactics.ML
-    Author:     Dmitriy Traytel, TU Muenchen
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2012
-
-Tactics for definition of bounded natural functors.
-*)
-
-signature BNF_DEF_TACTICS =
-sig
-  val mk_collect_set_map_tac: thm list -> tactic
-  val mk_map_id: thm -> thm
-  val mk_map_comp: thm -> thm
-  val mk_map_cong_tac: Proof.context -> thm -> tactic
-  val mk_in_mono_tac: int -> tactic
-  val mk_set_map: thm -> thm
-
-  val mk_rel_Grp_tac: thm list -> thm -> thm -> thm -> thm -> thm list ->
-    {prems: thm list, context: Proof.context} -> tactic
-  val mk_rel_eq_tac: int -> thm -> thm -> thm -> tactic
-  val mk_rel_OO_le_tac: thm list -> thm -> thm -> thm -> thm list ->
-    {prems: thm list, context: Proof.context} -> tactic
-  val mk_rel_conversep_tac: thm -> thm -> tactic
-  val mk_rel_conversep_le_tac: thm list -> thm -> thm -> thm -> thm list ->
-    {prems: thm list, context: Proof.context} -> tactic
-  val mk_rel_mono_tac: thm list -> thm -> tactic
-  val mk_rel_mono_strong_tac: thm -> thm list -> {prems: 'a, context: Proof.context} -> tactic
-
-  val mk_map_transfer_tac: thm -> thm -> thm list -> thm -> thm ->
-    {prems: thm list, context: Proof.context} -> tactic
-
-  val mk_in_bd_tac: int -> thm -> thm -> thm -> thm -> thm list -> thm list -> thm -> thm -> thm ->
-    thm -> {prems: thm list, context: Proof.context} -> tactic
-
-  val mk_trivial_wit_tac: thm list -> thm list -> {prems: thm list, context: Proof.context} ->
-    tactic
-end;
-
-structure BNF_Def_Tactics : BNF_DEF_TACTICS =
-struct
-
-open BNF_Util
-open BNF_Tactics
-
-val ord_eq_le_trans = @{thm ord_eq_le_trans};
-val ord_le_eq_trans = @{thm ord_le_eq_trans};
-val conversep_shift = @{thm conversep_le_swap} RS iffD1;
-
-fun mk_map_id id = mk_trans (fun_cong OF [id]) @{thm id_apply};
-fun mk_map_comp comp = @{thm o_eq_dest_lhs} OF [mk_sym comp];
-fun mk_map_cong_tac ctxt cong0 =
-  (hyp_subst_tac ctxt THEN' rtac cong0 THEN'
-   REPEAT_DETERM o (dtac meta_spec THEN' etac meta_mp THEN' atac)) 1;
-fun mk_set_map set_map0 = set_map0 RS @{thm comp_eq_dest};
-fun mk_in_mono_tac n = if n = 0 then rtac subset_UNIV 1
-  else (rtac subsetI THEN'
-  rtac CollectI) 1 THEN
-  REPEAT_DETERM (eresolve_tac [CollectE, conjE] 1) THEN
-  REPEAT_DETERM_N (n - 1)
-    ((rtac conjI THEN' etac subset_trans THEN' atac) 1) THEN
-  (etac subset_trans THEN' atac) 1;
-
-fun mk_collect_set_map_tac set_map0s =
-  (rtac (@{thm collect_o} RS trans) THEN' rtac @{thm arg_cong[of _ _ collect]} THEN'
-  EVERY' (map (fn set_map0 =>
-    rtac (mk_trans @{thm image_insert} @{thm arg_cong2[of _ _ _ _ insert]}) THEN'
-    rtac set_map0) set_map0s) THEN'
-  rtac @{thm image_empty}) 1;
-
-fun mk_rel_Grp_tac rel_OO_Grps map_id0 map_cong0 map_id map_comp set_maps
-  {context = ctxt, prems = _} =
-  let
-    val n = length set_maps;
-    val rel_OO_Grps_tac = if null rel_OO_Grps then K all_tac else rtac (hd rel_OO_Grps RS trans);
-  in
-    if null set_maps then
-      unfold_thms_tac ctxt ((map_id0 RS @{thm Grp_UNIV_id}) :: rel_OO_Grps) THEN
-      rtac @{thm Grp_UNIV_idI[OF refl]} 1
-    else
-      EVERY' [rel_OO_Grps_tac, rtac @{thm antisym}, rtac @{thm predicate2I},
-        REPEAT_DETERM o
-          eresolve_tac [CollectE, exE, conjE, @{thm GrpE}, @{thm relcomppE}, @{thm conversepE}],
-        hyp_subst_tac ctxt, rtac @{thm GrpI}, rtac trans, rtac map_comp, rtac map_cong0,
-        REPEAT_DETERM_N n o EVERY' [rtac @{thm Collect_split_Grp_eqD}, etac @{thm set_mp}, atac],
-        rtac CollectI,
-        CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS ord_eq_le_trans),
-          rtac @{thm image_subsetI}, rtac @{thm Collect_split_Grp_inD}, etac @{thm set_mp}, atac])
-        set_maps,
-        rtac @{thm predicate2I}, REPEAT_DETERM o eresolve_tac [@{thm GrpE}, exE, conjE],
-        hyp_subst_tac ctxt,
-        rtac @{thm relcomppI}, rtac @{thm conversepI},
-        EVERY' (map2 (fn convol => fn map_id0 =>
-          EVERY' [rtac @{thm GrpI}, rtac (box_equals OF [map_cong0, map_comp RS sym, map_id0]),
-            REPEAT_DETERM_N n o rtac (convol RS fun_cong),
-            REPEAT_DETERM o eresolve_tac [CollectE, conjE],
-            rtac CollectI,
-            CONJ_WRAP' (fn thm =>
-              EVERY' [rtac ord_eq_le_trans, rtac thm, rtac @{thm image_subsetI},
-                rtac @{thm convol_mem_GrpI}, etac set_mp, atac])
-            set_maps])
-          @{thms fst_convol snd_convol} [map_id, refl])] 1
-  end;
-
-fun mk_rel_eq_tac n rel_Grp rel_cong map_id0 =
-  (EVERY' (rtac (rel_cong RS trans) :: replicate n (rtac @{thm eq_alt})) THEN'
-  rtac (rel_Grp RSN (2, @{thm box_equals[OF _ sym sym[OF eq_alt]]})) THEN'
-  (if n = 0 then rtac refl
-  else EVERY' [rtac @{thm arg_cong2[of _ _ _ _ "Grp"]},
-    rtac @{thm equalityI}, rtac subset_UNIV, rtac subsetI, rtac CollectI,
-    CONJ_WRAP' (K (rtac subset_UNIV)) (1 upto n), rtac map_id0])) 1;
-
-fun mk_rel_mono_tac rel_OO_Grps in_mono =
-  let
-    val rel_OO_Grps_tac = if null rel_OO_Grps then K all_tac
-      else rtac (hd rel_OO_Grps RS ord_eq_le_trans) THEN'
-        rtac (hd rel_OO_Grps RS sym RSN (2, ord_le_eq_trans));
-  in
-    EVERY' [rel_OO_Grps_tac, rtac @{thm relcompp_mono}, rtac @{thm iffD2[OF conversep_mono]},
-      rtac @{thm Grp_mono}, rtac in_mono, REPEAT_DETERM o etac @{thm Collect_split_mono},
-      rtac @{thm Grp_mono}, rtac in_mono, REPEAT_DETERM o etac @{thm Collect_split_mono}] 1
-  end;
-
-fun mk_rel_conversep_le_tac rel_OO_Grps rel_eq map_cong0 map_comp set_maps
-  {context = ctxt, prems = _} =
-  let
-    val n = length set_maps;
-    val rel_OO_Grps_tac = if null rel_OO_Grps then K all_tac
-      else rtac (hd rel_OO_Grps RS ord_eq_le_trans) THEN'
-        rtac (hd rel_OO_Grps RS sym RS @{thm arg_cong[of _ _ conversep]} RSN (2, ord_le_eq_trans));
-  in
-    if null set_maps then rtac (rel_eq RS @{thm leq_conversepI}) 1
-    else
-      EVERY' [rel_OO_Grps_tac, rtac @{thm predicate2I},
-        REPEAT_DETERM o
-          eresolve_tac [CollectE, exE, conjE, @{thm GrpE}, @{thm relcomppE}, @{thm conversepE}],
-        hyp_subst_tac ctxt, rtac @{thm conversepI}, rtac @{thm relcomppI}, rtac @{thm conversepI},
-        EVERY' (map (fn thm => EVERY' [rtac @{thm GrpI}, rtac sym, rtac trans,
-          rtac map_cong0, REPEAT_DETERM_N n o rtac thm,
-          rtac (map_comp RS sym), rtac CollectI,
-          CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS ord_eq_le_trans),
-            etac @{thm flip_pred}]) set_maps]) [@{thm snd_fst_flip}, @{thm fst_snd_flip}])] 1
-  end;
-
-fun mk_rel_conversep_tac le_conversep rel_mono =
-  EVERY' [rtac @{thm antisym}, rtac le_conversep, rtac @{thm xt1(6)}, rtac conversep_shift,
-    rtac le_conversep, rtac @{thm iffD2[OF conversep_mono]}, rtac rel_mono,
-    REPEAT_DETERM o rtac @{thm eq_refl[OF sym[OF conversep_conversep]]}] 1;
-
-fun mk_rel_OO_le_tac rel_OO_Grps rel_eq map_cong0 map_comp set_maps
-  {context = ctxt, prems = _} =
-  let
-    val n = length set_maps;
-    fun in_tac nthO_in = rtac CollectI THEN'
-        CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS ord_eq_le_trans),
-          rtac @{thm image_subsetI}, rtac nthO_in, etac set_mp, atac]) set_maps;
-    val rel_OO_Grps_tac = if null rel_OO_Grps then K all_tac
-      else rtac (hd rel_OO_Grps RS ord_eq_le_trans) THEN'
-        rtac (@{thm arg_cong2[of _ _ _ _ "op OO"]} OF (replicate 2 (hd rel_OO_Grps RS sym)) RSN
-          (2, ord_le_eq_trans));
-  in
-    if null set_maps then rtac (rel_eq RS @{thm leq_OOI}) 1
-    else
-      EVERY' [rel_OO_Grps_tac, rtac @{thm predicate2I},
-        REPEAT_DETERM o
-          eresolve_tac [CollectE, exE, conjE, @{thm GrpE}, @{thm relcomppE}, @{thm conversepE}],
-        hyp_subst_tac ctxt,
-        rtac @{thm relcomppI}, rtac @{thm relcomppI}, rtac @{thm conversepI}, rtac @{thm GrpI},
-        rtac trans, rtac map_comp, rtac sym, rtac map_cong0,
-        REPEAT_DETERM_N n o rtac @{thm fst_fstOp},
-        in_tac @{thm fstOp_in},
-        rtac @{thm GrpI}, rtac trans, rtac map_comp, rtac map_cong0,
-        REPEAT_DETERM_N n o EVERY' [rtac trans, rtac o_apply, 
-          rtac ballE, rtac subst,
-          rtac @{thm csquare_def}, rtac @{thm csquare_fstOp_sndOp}, atac, etac notE,
-          etac set_mp, atac],
-        in_tac @{thm fstOp_in},
-        rtac @{thm relcomppI}, rtac @{thm conversepI}, rtac @{thm GrpI},
-        rtac trans, rtac map_comp, rtac map_cong0,
-        REPEAT_DETERM_N n o rtac o_apply,
-        in_tac @{thm sndOp_in},
-        rtac @{thm GrpI}, rtac trans, rtac map_comp, rtac sym, rtac map_cong0,
-        REPEAT_DETERM_N n o rtac @{thm snd_sndOp},
-        in_tac @{thm sndOp_in}] 1
-  end;
-
-fun mk_rel_mono_strong_tac in_rel set_maps {context = ctxt, prems = _} =
-  if null set_maps then atac 1
-  else
-    unfold_tac ctxt [in_rel] THEN
-    REPEAT_DETERM (eresolve_tac [exE, CollectE, conjE] 1) THEN
-    hyp_subst_tac ctxt 1 THEN
-    unfold_tac ctxt set_maps THEN
-    EVERY' [rtac exI, rtac @{thm conjI[OF CollectI conjI[OF refl refl]]},
-      CONJ_WRAP' (K (etac @{thm Collect_split_mono_strong} THEN' atac)) set_maps] 1;
-
-fun mk_map_transfer_tac rel_mono in_rel set_maps map_cong0 map_comp
-  {context = ctxt, prems = _} =
-  let
-    val n = length set_maps;
-    val in_tac = if n = 0 then rtac UNIV_I else
-      rtac CollectI THEN' CONJ_WRAP' (fn thm =>
-        etac (thm RS
-          @{thm ord_eq_le_trans[OF _ subset_trans[OF image_mono convol_image_vimage2p]]}))
-      set_maps;
-  in
-    REPEAT_DETERM_N n (HEADGOAL (rtac @{thm fun_relI})) THEN
-    unfold_thms_tac ctxt @{thms fun_rel_iff_leq_vimage2p} THEN
-    HEADGOAL (EVERY' [rtac @{thm order_trans}, rtac rel_mono, REPEAT_DETERM_N n o atac,
-      rtac @{thm predicate2I}, dtac (in_rel RS iffD1),
-      REPEAT_DETERM o eresolve_tac [exE, CollectE, conjE], hyp_subst_tac ctxt,
-      rtac @{thm vimage2pI}, rtac (in_rel RS iffD2), rtac exI, rtac conjI, in_tac,
-      rtac conjI,
-      EVERY' (map (fn convol =>
-        rtac (box_equals OF [map_cong0, map_comp RS sym, map_comp RS sym]) THEN'
-        REPEAT_DETERM_N n o rtac (convol RS fun_cong)) @{thms fst_convol snd_convol})])
-  end;
-
-fun mk_in_bd_tac live surj_imp_ordLeq_inst map_comp map_id map_cong0 set_maps set_bds
-  bd_card_order bd_Card_order bd_Cinfinite bd_Cnotzero {context = ctxt, prems = _} =
-  if live = 0 then
-    rtac @{thm ordLeq_transitive[OF ordLeq_csum2[OF card_of_Card_order]
-      ordLeq_cexp2[OF ordLeq_refl[OF Card_order_ctwo] Card_order_csum]]} 1
-  else
-    let
-      val bd'_Cinfinite = bd_Cinfinite RS @{thm Cinfinite_csum1};
-      val inserts =
-        map (fn set_bd => 
-          iffD2 OF [@{thm card_of_ordLeq}, @{thm ordLeq_ordIso_trans} OF
-            [set_bd, bd_Card_order RS @{thm card_of_Field_ordIso} RS @{thm ordIso_symmetric}]])
-        set_bds;        
-    in
-      EVERY' [rtac (Drule.rotate_prems 1 ctrans), rtac @{thm cprod_cinfinite_bound},
-        rtac (ctrans OF @{thms ordLeq_csum2 ordLeq_cexp2}), rtac @{thm card_of_Card_order},
-        rtac @{thm ordLeq_csum2}, rtac @{thm Card_order_ctwo}, rtac @{thm Card_order_csum},
-        rtac @{thm ordIso_ordLeq_trans}, rtac @{thm cexp_cong1},
-        if live = 1 then rtac @{thm ordIso_refl[OF Card_order_csum]}
-        else
-          REPEAT_DETERM_N (live - 2) o rtac @{thm ordIso_transitive[OF csum_cong2]} THEN'
-          REPEAT_DETERM_N (live - 1) o rtac @{thm csum_csum},
-        rtac bd_Card_order, rtac (@{thm cexp_mono2_Cnotzero} RS ctrans), rtac @{thm ordLeq_csum1},
-        rtac bd_Card_order, rtac @{thm Card_order_csum}, rtac bd_Cnotzero,
-        rtac @{thm csum_Cfinite_cexp_Cinfinite},
-        rtac (if live = 1 then @{thm card_of_Card_order} else @{thm Card_order_csum}),
-        CONJ_WRAP_GEN' (rtac @{thm Cfinite_csum}) (K (rtac @{thm Cfinite_cone})) set_maps,
-        rtac bd'_Cinfinite, rtac @{thm card_of_Card_order},
-        rtac @{thm Card_order_cexp}, rtac @{thm Cinfinite_cexp}, rtac @{thm ordLeq_csum2},
-        rtac @{thm Card_order_ctwo}, rtac bd'_Cinfinite,
-        rtac (Drule.rotate_prems 1 (@{thm cprod_mono2} RSN (2, ctrans))),
-        REPEAT_DETERM_N (live - 1) o
-          (rtac (bd_Cinfinite RS @{thm cprod_cexp_csum_cexp_Cinfinite} RSN (2, ctrans)) THEN'
-           rtac @{thm ordLeq_ordIso_trans[OF cprod_mono2 ordIso_symmetric[OF cprod_cexp]]}),
-        rtac @{thm ordLeq_refl[OF Card_order_cexp]}] 1 THEN
-      unfold_thms_tac ctxt [bd_card_order RS @{thm card_order_csum_cone_cexp_def}] THEN
-      unfold_thms_tac ctxt @{thms cprod_def Field_card_of} THEN
-      EVERY' [rtac (Drule.rotate_prems 1 ctrans), rtac surj_imp_ordLeq_inst, rtac subsetI,
-        Method.insert_tac inserts, REPEAT_DETERM o dtac meta_spec,
-        REPEAT_DETERM o eresolve_tac [exE, Tactic.make_elim conjunct1], etac CollectE,
-        if live = 1 then K all_tac
-        else REPEAT_DETERM_N (live - 2) o (etac conjE THEN' rotate_tac ~1) THEN' etac conjE,
-        rtac (Drule.rotate_prems 1 @{thm image_eqI}), rtac @{thm SigmaI}, rtac @{thm UNIV_I},
-        CONJ_WRAP_GEN' (rtac @{thm SigmaI})
-          (K (etac @{thm If_the_inv_into_in_Func} THEN' atac)) set_maps,
-        rtac sym,
-        rtac (Drule.rotate_prems 1
-           ((box_equals OF [map_cong0 OF replicate live @{thm If_the_inv_into_f_f},
-             map_comp RS sym, map_id]) RSN (2, trans))),
-        REPEAT_DETERM_N (2 * live) o atac,
-        REPEAT_DETERM_N live o rtac (@{thm prod.cases} RS trans),
-        rtac refl,
-        rtac @{thm surj_imp_ordLeq}, rtac subsetI, rtac (Drule.rotate_prems 1 @{thm image_eqI}),
-        REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac CollectI,
-        CONJ_WRAP' (fn thm =>
-          rtac (thm RS ord_eq_le_trans) THEN' etac @{thm subset_trans[OF image_mono Un_upper1]})
-        set_maps,
-        rtac sym,
-        rtac (box_equals OF [map_cong0 OF replicate live @{thm fun_cong[OF sum_case_o_inj(1)]},
-           map_comp RS sym, map_id])] 1
-  end;
-
-fun mk_trivial_wit_tac wit_defs set_maps {context = ctxt, prems = _} =
-  unfold_thms_tac ctxt wit_defs THEN HEADGOAL (EVERY' (map (fn thm =>
-    dtac (thm RS equalityD1 RS set_mp) THEN' etac imageE THEN' atac) set_maps)) THEN ALLGOALS atac;
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_fp_def_sugar.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1523 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_fp_def_sugar.ML
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2012, 2013
-
-Sugared datatype and codatatype constructions.
-*)
-
-signature BNF_FP_DEF_SUGAR =
-sig
-  type fp_sugar =
-    {T: typ,
-     fp: BNF_FP_Util.fp_kind,
-     index: int,
-     pre_bnfs: BNF_Def.bnf list,
-     nested_bnfs: BNF_Def.bnf list,
-     nesting_bnfs: BNF_Def.bnf list,
-     fp_res: BNF_FP_Util.fp_result,
-     ctr_defss: thm list list,
-     ctr_sugars: Ctr_Sugar.ctr_sugar list,
-     co_iterss: term list list,
-     mapss: thm list list,
-     co_inducts: thm list,
-     co_iter_thmsss: thm list list list,
-     disc_co_itersss: thm list list list,
-     sel_co_iterssss: thm list list list list};
-
-  val of_fp_sugar: (fp_sugar -> 'a list) -> fp_sugar -> 'a
-  val eq_fp_sugar: fp_sugar * fp_sugar -> bool
-  val morph_fp_sugar: morphism -> fp_sugar -> fp_sugar
-  val transfer_fp_sugar: Proof.context -> fp_sugar -> fp_sugar
-  val fp_sugar_of: Proof.context -> string -> fp_sugar option
-  val fp_sugars_of: Proof.context -> fp_sugar list
-
-  val co_induct_of: 'a list -> 'a
-  val strong_co_induct_of: 'a list -> 'a
-
-  val tvar_subst: theory -> typ list -> typ list -> ((string * int) * typ) list
-  val exists_subtype_in: typ list -> typ -> bool
-  val flat_rec_arg_args: 'a list list -> 'a list
-  val flat_corec_preds_predsss_gettersss: 'a list -> 'a list list list -> 'a list list list ->
-    'a list
-  val mk_co_iter: theory -> BNF_FP_Util.fp_kind -> typ -> typ list -> term -> term
-  val nesty_bnfs: Proof.context -> typ list list list -> typ list -> BNF_Def.bnf list
-
-  type lfp_sugar_thms =
-    (thm list * thm * Args.src list)
-    * (thm list list * thm list list * Args.src list)
-
-  val morph_lfp_sugar_thms: morphism -> lfp_sugar_thms -> lfp_sugar_thms
-  val transfer_lfp_sugar_thms: Proof.context -> lfp_sugar_thms -> lfp_sugar_thms
-
-  type gfp_sugar_thms =
-    ((thm list * thm) list * Args.src list)
-    * (thm list list * thm list list * Args.src list)
-    * (thm list list * thm list list * Args.src list)
-    * (thm list list * thm list list * Args.src list)
-    * (thm list list list * thm list list list * Args.src list)
-
-  val morph_gfp_sugar_thms: morphism -> gfp_sugar_thms -> gfp_sugar_thms
-  val transfer_gfp_sugar_thms: Proof.context -> gfp_sugar_thms -> gfp_sugar_thms
-
-  val mk_co_iters_prelims: BNF_FP_Util.fp_kind -> typ list list list -> typ list -> typ list ->
-    int list -> int list list -> term list list -> Proof.context ->
-    (term list list
-     * (typ list list * typ list list list list * term list list
-        * term list list list list) list option
-     * (string * term list * term list list
-        * ((term list list * term list list list) * (typ list * typ list list)) list) option)
-    * Proof.context
-  val mk_iter_fun_arg_types: typ list list list -> int list -> int list list -> term ->
-    typ list list list list
-  val mk_coiter_fun_arg_types: typ list list list -> typ list -> int list -> term ->
-    typ list list
-    * (typ list list list list * typ list list list * typ list list list list * typ list)
-  val define_iters: string list ->
-    (typ list list * typ list list list list * term list list * term list list list list) list ->
-    (string -> binding) -> typ list -> typ list -> term list -> Proof.context ->
-    (term list * thm list) * Proof.context
-  val define_coiters: string list -> string * term list * term list list
-    * ((term list list * term list list list) * (typ list * typ list list)) list ->
-    (string -> binding) -> typ list -> typ list -> term list -> Proof.context ->
-    (term list * thm list) * Proof.context
-  val derive_induct_iters_thms_for_types: BNF_Def.bnf list ->
-    (typ list list * typ list list list list * term list list * term list list list list) list ->
-    thm -> thm list list -> BNF_Def.bnf list -> BNF_Def.bnf list -> typ list -> typ list ->
-    typ list -> typ list list list -> term list list -> thm list list -> term list list ->
-    thm list list -> local_theory -> lfp_sugar_thms
-  val derive_coinduct_coiters_thms_for_types: BNF_Def.bnf list ->
-    string * term list * term list list * ((term list list * term list list list)
-      * (typ list * typ list list)) list ->
-    thm -> thm list -> thm list -> thm list list -> BNF_Def.bnf list -> typ list -> typ list ->
-    typ list -> typ list list list -> int list list -> int list list -> int list -> thm list list ->
-    Ctr_Sugar.ctr_sugar list -> term list list -> thm list list -> (thm list -> thm list) ->
-    local_theory -> gfp_sugar_thms
-  val co_datatypes: BNF_FP_Util.fp_kind -> (mixfix list -> binding list -> binding list ->
-      binding list list -> binding list -> (string * sort) list -> typ list * typ list list ->
-      BNF_Def.bnf list -> local_theory -> BNF_FP_Util.fp_result * local_theory) ->
-    (bool * (bool * bool)) * (((((binding * (typ * sort)) list * binding) * (binding * binding))
-      * mixfix) * ((((binding * binding) * (binding * typ) list) * (binding * term) list) *
-        mixfix) list) list ->
-    local_theory -> local_theory
-  val parse_co_datatype_cmd: BNF_FP_Util.fp_kind -> (mixfix list -> binding list -> binding list ->
-      binding list list -> binding list -> (string * sort) list -> typ list * typ list list ->
-      BNF_Def.bnf list -> local_theory -> BNF_FP_Util.fp_result * local_theory) ->
-    (local_theory -> local_theory) parser
-end;
-
-structure BNF_FP_Def_Sugar : BNF_FP_DEF_SUGAR =
-struct
-
-open Ctr_Sugar
-open BNF_Util
-open BNF_Comp
-open BNF_Def
-open BNF_FP_Util
-open BNF_FP_Def_Sugar_Tactics
-
-val EqN = "Eq_";
-
-type fp_sugar =
-  {T: typ,
-   fp: fp_kind,
-   index: int,
-   pre_bnfs: bnf list,
-   nested_bnfs: bnf list,
-   nesting_bnfs: bnf list,
-   fp_res: fp_result,
-   ctr_defss: thm list list,
-   ctr_sugars: ctr_sugar list,
-   co_iterss: term list list,
-   mapss: thm list list,
-   co_inducts: thm list,
-   co_iter_thmsss: thm list list list,
-   disc_co_itersss: thm list list list,
-   sel_co_iterssss: thm list list list list};
-
-fun of_fp_sugar f (fp_sugar as ({index, ...}: fp_sugar)) = nth (f fp_sugar) index;
-
-fun eq_fp_sugar ({T = T1, fp = fp1, index = index1, fp_res = fp_res1, ...} : fp_sugar,
-    {T = T2, fp = fp2, index = index2, fp_res = fp_res2, ...} : fp_sugar) =
-  T1 = T2 andalso fp1 = fp2 andalso index1 = index2 andalso eq_fp_result (fp_res1, fp_res2);
-
-fun morph_fp_sugar phi ({T, fp, index, pre_bnfs, nested_bnfs, nesting_bnfs, fp_res, ctr_defss,
-    ctr_sugars, co_iterss, mapss, co_inducts, co_iter_thmsss, disc_co_itersss, sel_co_iterssss}
-    : fp_sugar) =
-  {T = Morphism.typ phi T, fp = fp, index = index, pre_bnfs = map (morph_bnf phi) pre_bnfs,
-    nested_bnfs = map (morph_bnf phi) nested_bnfs, nesting_bnfs = map (morph_bnf phi) nesting_bnfs,
-   fp_res = morph_fp_result phi fp_res,
-   ctr_defss = map (map (Morphism.thm phi)) ctr_defss,
-   ctr_sugars = map (morph_ctr_sugar phi) ctr_sugars,
-   co_iterss = map (map (Morphism.term phi)) co_iterss,
-   mapss = map (map (Morphism.thm phi)) mapss,
-   co_inducts = map (Morphism.thm phi) co_inducts,
-   co_iter_thmsss = map (map (map (Morphism.thm phi))) co_iter_thmsss,
-   disc_co_itersss = map (map (map (Morphism.thm phi))) disc_co_itersss,
-   sel_co_iterssss = map (map (map (map (Morphism.thm phi)))) sel_co_iterssss};
-
-val transfer_fp_sugar =
-  morph_fp_sugar o Morphism.transfer_morphism o Proof_Context.theory_of;
-
-structure Data = Generic_Data
-(
-  type T = fp_sugar Symtab.table;
-  val empty = Symtab.empty;
-  val extend = I;
-  val merge = Symtab.merge eq_fp_sugar;
-);
-
-fun fp_sugar_of ctxt =
-  Symtab.lookup (Data.get (Context.Proof ctxt))
-  #> Option.map (transfer_fp_sugar ctxt);
-
-fun fp_sugars_of ctxt =
-  Symtab.fold (cons o transfer_fp_sugar ctxt o snd) (Data.get (Context.Proof ctxt)) [];
-
-fun co_induct_of (i :: _) = i;
-fun strong_co_induct_of [_, s] = s;
-
-(* TODO: register "sum" and "prod" as datatypes to enable N2M reduction for them *)
-
-fun register_fp_sugar key fp_sugar =
-  Local_Theory.declaration {syntax = false, pervasive = true}
-    (fn phi => Data.map (Symtab.default (key, morph_fp_sugar phi fp_sugar)));
-
-fun register_fp_sugars fp pre_bnfs nested_bnfs nesting_bnfs (fp_res as {Ts, ...}) ctr_defss
-    ctr_sugars co_iterss mapss co_inducts co_iter_thmsss disc_co_itersss sel_co_iterssss lthy =
-  (0, lthy)
-  |> fold (fn T as Type (s, _) => fn (kk, lthy) => (kk + 1,
-    register_fp_sugar s {T = T, fp = fp, index = kk, pre_bnfs = pre_bnfs,
-        nested_bnfs = nested_bnfs, nesting_bnfs = nesting_bnfs, fp_res = fp_res,
-        ctr_defss = ctr_defss, ctr_sugars = ctr_sugars, co_iterss = co_iterss, mapss = mapss,
-        co_inducts = co_inducts, co_iter_thmsss = co_iter_thmsss, disc_co_itersss = disc_co_itersss,
-        sel_co_iterssss = sel_co_iterssss}
-      lthy)) Ts
-  |> snd;
-
-(* This function could produce clashes in contrived examples (e.g., "x.A", "x.x_A", "y.A"). *)
-fun quasi_unambiguous_case_names names =
-  let
-    val ps = map (`Long_Name.base_name) names;
-    val dups = Library.duplicates (op =) (map fst ps);
-    fun underscore s =
-      let val ss = space_explode Long_Name.separator s in
-        space_implode "_" (drop (length ss - 2) ss)
-      end;
-  in
-    map (fn (base, full) => if member (op =) dups base then underscore full else base) ps
-  end;
-
-val id_def = @{thm id_def};
-val mp_conj = @{thm mp_conj};
-
-val nitpicksimp_attrs = @{attributes [nitpick_simp]};
-val code_nitpicksimp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs;
-val simp_attrs = @{attributes [simp]};
-
-fun tvar_subst thy Ts Us =
-  Vartab.fold (cons o apsnd snd) (fold (Sign.typ_match thy) (Ts ~~ Us) Vartab.empty) [];
-
-val exists_subtype_in = Term.exists_subtype o member (op =);
-
-val lists_bmoc = fold (fn xs => fn t => Term.list_comb (t, xs));
-
-fun flat_rec_arg_args xss =
-  (* FIXME (once the old datatype package is phased out): The first line below gives the preferred
-     order. The second line is for compatibility with the old datatype package. *)
-(*
-  flat xss
-*)
-  map hd xss @ maps tl xss;
-
-fun flat_corec_predss_getterss qss fss = maps (op @) (qss ~~ fss);
-
-fun flat_corec_preds_predsss_gettersss [] [qss] [fss] = flat_corec_predss_getterss qss fss
-  | flat_corec_preds_predsss_gettersss (p :: ps) (qss :: qsss) (fss :: fsss) =
-    p :: flat_corec_predss_getterss qss fss @ flat_corec_preds_predsss_gettersss ps qsss fsss;
-
-fun mk_tupled_fun x f xs =
-  if xs = [x] then f else HOLogic.tupled_lambda x (Term.list_comb (f, xs));
-
-fun mk_uncurried2_fun f xss =
-  mk_tupled_fun (HOLogic.mk_tuple (map HOLogic.mk_tuple xss)) f (flat_rec_arg_args xss);
-
-fun mk_flip (x, Type (_, [T1, Type (_, [T2, T3])])) =
-  Abs ("x", T1, Abs ("y", T2, Var (x, T2 --> T1 --> T3) $ Bound 0 $ Bound 1));
-
-fun flip_rels lthy n thm =
-  let
-    val Rs = Term.add_vars (prop_of thm) [];
-    val Rs' = rev (drop (length Rs - n) Rs);
-    val cRs = map (fn f => (certify lthy (Var f), certify lthy (mk_flip f))) Rs';
-  in
-    Drule.cterm_instantiate cRs thm
-  end;
-
-fun mk_ctor_or_dtor get_T Ts t =
-  let val Type (_, Ts0) = get_T (fastype_of t) in
-    Term.subst_atomic_types (Ts0 ~~ Ts) t
-  end;
-
-val mk_ctor = mk_ctor_or_dtor range_type;
-val mk_dtor = mk_ctor_or_dtor domain_type;
-
-fun mk_co_iter thy fp fpT Cs t =
-  let
-    val (f_Cs, Type (_, [prebody, body])) = strip_fun_type (fastype_of t);
-    val fpT0 = fp_case fp prebody body;
-    val Cs0 = distinct (op =) (map (fp_case fp body_type domain_type) f_Cs);
-    val rho = tvar_subst thy (fpT0 :: Cs0) (fpT :: Cs);
-  in
-    Term.subst_TVars rho t
-  end;
-
-fun mk_co_iters thy fp fpTs Cs ts0 =
-  let
-    val nn = length fpTs;
-    val (fpTs0, Cs0) =
-      map ((fp = Greatest_FP ? swap) o dest_funT o snd o strip_typeN nn o fastype_of) ts0
-      |> split_list;
-    val rho = tvar_subst thy (fpTs0 @ Cs0) (fpTs @ Cs);
-  in
-    map (Term.subst_TVars rho) ts0
-  end;
-
-val mk_fp_iter_fun_types = binder_fun_types o fastype_of;
-
-fun unzip_recT (Type (@{type_name prod}, _)) T = [T]
-  | unzip_recT _ (Type (@{type_name prod}, Ts)) = Ts
-  | unzip_recT _ T = [T];
-
-fun unzip_corecT (Type (@{type_name sum}, _)) T = [T]
-  | unzip_corecT _ (Type (@{type_name sum}, Ts)) = Ts
-  | unzip_corecT _ T = [T];
-
-fun liveness_of_fp_bnf n bnf =
-  (case T_of_bnf bnf of
-    Type (_, Ts) => map (not o member (op =) (deads_of_bnf bnf)) Ts
-  | _ => replicate n false);
-
-fun cannot_merge_types () = error "Mutually recursive types must have the same type parameters";
-
-fun merge_type_arg T T' = if T = T' then T else cannot_merge_types ();
-
-fun merge_type_args (As, As') =
-  if length As = length As' then map2 merge_type_arg As As' else cannot_merge_types ();
-
-fun reassoc_conjs thm =
-  reassoc_conjs (thm RS @{thm conj_assoc[THEN iffD1]})
-  handle THM _ => thm;
-
-fun type_args_named_constrained_of ((((ncAs, _), _), _), _) = ncAs;
-fun type_binding_of ((((_, b), _), _), _) = b;
-fun map_binding_of (((_, (b, _)), _), _) = b;
-fun rel_binding_of (((_, (_, b)), _), _) = b;
-fun mixfix_of ((_, mx), _) = mx;
-fun ctr_specs_of (_, ctr_specs) = ctr_specs;
-
-fun disc_of ((((disc, _), _), _), _) = disc;
-fun ctr_of ((((_, ctr), _), _), _) = ctr;
-fun args_of (((_, args), _), _) = args;
-fun defaults_of ((_, ds), _) = ds;
-fun ctr_mixfix_of (_, mx) = mx;
-
-fun add_nesty_bnf_names Us =
-  let
-    fun add (Type (s, Ts)) ss =
-        let val (needs, ss') = fold_map add Ts ss in
-          if exists I needs then (true, insert (op =) s ss') else (false, ss')
-        end
-      | add T ss = (member (op =) Us T, ss);
-  in snd oo add end;
-
-fun nesty_bnfs ctxt ctr_Tsss Us =
-  map_filter (bnf_of ctxt) (fold (fold (fold (add_nesty_bnf_names Us))) ctr_Tsss []);
-
-fun indexify proj xs f p = f (find_index (curry (op =) (proj p)) xs) p;
-
-type lfp_sugar_thms =
-  (thm list * thm * Args.src list)
-  * (thm list list * thm list list * Args.src list)
-
-fun morph_lfp_sugar_thms phi ((inducts, induct, induct_attrs), (foldss, recss, iter_attrs)) =
-  ((map (Morphism.thm phi) inducts, Morphism.thm phi induct, induct_attrs),
-   (map (map (Morphism.thm phi)) foldss, map (map (Morphism.thm phi)) recss, iter_attrs));
-
-val transfer_lfp_sugar_thms =
-  morph_lfp_sugar_thms o Morphism.transfer_morphism o Proof_Context.theory_of;
-
-type gfp_sugar_thms =
-  ((thm list * thm) list * Args.src list)
-  * (thm list list * thm list list * Args.src list)
-  * (thm list list * thm list list * Args.src list)
-  * (thm list list * thm list list * Args.src list)
-  * (thm list list list * thm list list list * Args.src list);
-
-fun morph_gfp_sugar_thms phi ((coinducts_pairs, coinduct_attrs),
-    (unfoldss, corecss, coiter_attrs), (disc_unfoldss, disc_corecss, disc_iter_attrs),
-    (disc_unfold_iffss, disc_corec_iffss, disc_iter_iff_attrs),
-    (sel_unfoldsss, sel_corecsss, sel_iter_attrs)) =
-  ((map (apfst (map (Morphism.thm phi)) o apsnd (Morphism.thm phi)) coinducts_pairs,
-    coinduct_attrs),
-   (map (map (Morphism.thm phi)) unfoldss, map (map (Morphism.thm phi)) corecss, coiter_attrs),
-   (map (map (Morphism.thm phi)) disc_unfoldss, map (map (Morphism.thm phi)) disc_corecss,
-    disc_iter_attrs),
-   (map (map (Morphism.thm phi)) disc_unfold_iffss, map (map (Morphism.thm phi)) disc_corec_iffss,
-    disc_iter_iff_attrs),
-   (map (map (map (Morphism.thm phi))) sel_unfoldsss,
-    map (map (map (Morphism.thm phi))) sel_corecsss, sel_iter_attrs));
-
-val transfer_gfp_sugar_thms =
-  morph_gfp_sugar_thms o Morphism.transfer_morphism o Proof_Context.theory_of;
-
-fun mk_iter_fun_arg_types0 n ms = map2 dest_tupleT ms o dest_sumTN_balanced n o domain_type;
-
-fun mk_iter_fun_arg_types ctr_Tsss ns mss =
-  mk_fp_iter_fun_types
-  #> map3 mk_iter_fun_arg_types0 ns mss
-  #> map2 (map2 (map2 unzip_recT)) ctr_Tsss;
-
-fun mk_iters_args_types ctr_Tsss Cs ns mss ctor_iter_fun_Tss lthy =
-  let
-    val Css = map2 replicate ns Cs;
-    val y_Tsss = map3 mk_iter_fun_arg_types0 ns mss (map un_fold_of ctor_iter_fun_Tss);
-    val g_Tss = map2 (fn C => map (fn y_Ts => y_Ts ---> C)) Cs y_Tsss;
-
-    val ((gss, ysss), lthy) =
-      lthy
-      |> mk_Freess "f" g_Tss
-      ||>> mk_Freesss "x" y_Tsss;
-
-    val y_Tssss = map (map (map single)) y_Tsss;
-    val yssss = map (map (map single)) ysss;
-
-    val z_Tssss =
-      map4 (fn n => fn ms => fn ctr_Tss => fn ctor_iter_fun_Ts =>
-          map3 (fn m => fn ctr_Ts => fn ctor_iter_fun_T =>
-              map2 unzip_recT ctr_Ts (dest_tupleT m ctor_iter_fun_T))
-            ms ctr_Tss (dest_sumTN_balanced n (domain_type (co_rec_of ctor_iter_fun_Ts))))
-        ns mss ctr_Tsss ctor_iter_fun_Tss;
-
-    val z_Tsss' = map (map flat_rec_arg_args) z_Tssss;
-    val h_Tss = map2 (map2 (curry (op --->))) z_Tsss' Css;
-
-    val hss = map2 (map2 retype_free) h_Tss gss;
-    val zssss_hd = map2 (map2 (map2 (retype_free o hd))) z_Tssss ysss;
-    val (zssss_tl, lthy) =
-      lthy
-      |> mk_Freessss "y" (map (map (map tl)) z_Tssss);
-    val zssss = map2 (map2 (map2 cons)) zssss_hd zssss_tl;
-  in
-    ([(g_Tss, y_Tssss, gss, yssss), (h_Tss, z_Tssss, hss, zssss)], lthy)
-  end;
-
-fun mk_coiter_fun_arg_types0 ctr_Tsss Cs ns fun_Ts =
-  let
-    (*avoid "'a itself" arguments in coiterators*)
-    fun repair_arity [[]] = [[@{typ unit}]]
-      | repair_arity Tss = Tss;
-
-    val ctr_Tsss' = map repair_arity ctr_Tsss;
-    val f_sum_prod_Ts = map range_type fun_Ts;
-    val f_prod_Tss = map2 dest_sumTN_balanced ns f_sum_prod_Ts;
-    val f_Tsss = map2 (map2 (dest_tupleT o length)) ctr_Tsss' f_prod_Tss;
-    val f_Tssss = map3 (fn C => map2 (map2 (map (curry (op -->) C) oo unzip_corecT)))
-      Cs ctr_Tsss' f_Tsss;
-    val q_Tssss = map (map (map (fn [_] => [] | [_, T] => [mk_pred1T (domain_type T)]))) f_Tssss;
-  in
-    (q_Tssss, f_Tsss, f_Tssss, f_sum_prod_Ts)
-  end;
-
-fun mk_coiter_p_pred_types Cs ns = map2 (fn n => replicate (Int.max (0, n - 1)) o mk_pred1T) ns Cs;
-
-fun mk_coiter_fun_arg_types ctr_Tsss Cs ns dtor_coiter =
-  (mk_coiter_p_pred_types Cs ns,
-   mk_fp_iter_fun_types dtor_coiter |> mk_coiter_fun_arg_types0 ctr_Tsss Cs ns);
-
-fun mk_coiters_args_types ctr_Tsss Cs ns dtor_coiter_fun_Tss lthy =
-  let
-    val p_Tss = mk_coiter_p_pred_types Cs ns;
-
-    fun mk_types get_Ts =
-      let
-        val fun_Ts = map get_Ts dtor_coiter_fun_Tss;
-        val (q_Tssss, f_Tsss, f_Tssss, f_sum_prod_Ts) = mk_coiter_fun_arg_types0 ctr_Tsss Cs ns fun_Ts;
-        val pf_Tss = map3 flat_corec_preds_predsss_gettersss p_Tss q_Tssss f_Tssss;
-      in
-        (q_Tssss, f_Tsss, f_Tssss, (f_sum_prod_Ts, pf_Tss))
-      end;
-
-    val (r_Tssss, g_Tsss, g_Tssss, unfold_types) = mk_types un_fold_of;
-    val (s_Tssss, h_Tsss, h_Tssss, corec_types) = mk_types co_rec_of;
-
-    val ((((Free (z, _), cs), pss), gssss), lthy) =
-      lthy
-      |> yield_singleton (mk_Frees "z") dummyT
-      ||>> mk_Frees "a" Cs
-      ||>> mk_Freess "p" p_Tss
-      ||>> mk_Freessss "g" g_Tssss;
-    val rssss = map (map (map (fn [] => []))) r_Tssss;
-
-    val hssss_hd = map2 (map2 (map2 (fn T :: _ => fn [g] => retype_free T g))) h_Tssss gssss;
-    val ((sssss, hssss_tl), lthy) =
-      lthy
-      |> mk_Freessss "q" s_Tssss
-      ||>> mk_Freessss "h" (map (map (map tl)) h_Tssss);
-    val hssss = map2 (map2 (map2 cons)) hssss_hd hssss_tl;
-
-    val cpss = map2 (map o rapp) cs pss;
-
-    fun build_sum_inj mk_inj = build_map lthy (uncurry mk_inj o dest_sumT o snd);
-
-    fun build_dtor_coiter_arg _ [] [cf] = cf
-      | build_dtor_coiter_arg T [cq] [cf, cf'] =
-        mk_If cq (build_sum_inj Inl_const (fastype_of cf, T) $ cf)
-          (build_sum_inj Inr_const (fastype_of cf', T) $ cf');
-
-    fun mk_args qssss fssss f_Tsss =
-      let
-        val pfss = map3 flat_corec_preds_predsss_gettersss pss qssss fssss;
-        val cqssss = map2 (map o map o map o rapp) cs qssss;
-        val cfssss = map2 (map o map o map o rapp) cs fssss;
-        val cqfsss = map3 (map3 (map3 build_dtor_coiter_arg)) f_Tsss cqssss cfssss;
-      in (pfss, cqfsss) end;
-
-    val unfold_args = mk_args rssss gssss g_Tsss;
-    val corec_args = mk_args sssss hssss h_Tsss;
-  in
-    ((z, cs, cpss, [(unfold_args, unfold_types), (corec_args, corec_types)]), lthy)
-  end;
-
-fun mk_co_iters_prelims fp ctr_Tsss fpTs Cs ns mss xtor_co_iterss0 lthy =
-  let
-    val thy = Proof_Context.theory_of lthy;
-
-    val (xtor_co_iter_fun_Tss, xtor_co_iterss) =
-      map (mk_co_iters thy fp fpTs Cs #> `(mk_fp_iter_fun_types o hd)) (transpose xtor_co_iterss0)
-      |> apsnd transpose o apfst transpose o split_list;
-
-    val ((iters_args_types, coiters_args_types), lthy') =
-      if fp = Least_FP then
-        mk_iters_args_types ctr_Tsss Cs ns mss xtor_co_iter_fun_Tss lthy |>> (rpair NONE o SOME)
-      else
-        mk_coiters_args_types ctr_Tsss Cs ns xtor_co_iter_fun_Tss lthy |>> (pair NONE o SOME)
-  in
-    ((xtor_co_iterss, iters_args_types, coiters_args_types), lthy')
-  end;
-
-fun mk_preds_getterss_join c cps sum_prod_T cqfss =
-  let val n = length cqfss in
-    Term.lambda c (mk_IfN sum_prod_T cps
-      (map2 (mk_InN_balanced sum_prod_T n) (map HOLogic.mk_tuple cqfss) (1 upto n)))
-  end;
-
-fun define_co_iters fp fpT Cs binding_specs lthy0 =
-  let
-    val thy = Proof_Context.theory_of lthy0;
-
-    val maybe_conceal_def_binding = Thm.def_binding
-      #> Config.get lthy0 bnf_note_all = false ? Binding.conceal;
-
-    val ((csts, defs), (lthy', lthy)) = lthy0
-      |> apfst split_list o fold_map (fn (b, rhs) =>
-        Local_Theory.define ((b, NoSyn), ((maybe_conceal_def_binding b, []), rhs))
-        #>> apsnd snd) binding_specs
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy lthy';
-
-    val csts' = map (mk_co_iter thy fp fpT Cs o Morphism.term phi) csts;
-    val defs' = map (Morphism.thm phi) defs;
-  in
-    ((csts', defs'), lthy')
-  end;
-
-fun define_iters iterNs iter_args_typess' mk_binding fpTs Cs ctor_iters lthy =
-  let
-    val nn = length fpTs;
-
-    val fpT_to_C as Type (_, [fpT, _]) = snd (strip_typeN nn (fastype_of (hd ctor_iters)));
-
-    fun generate_iter pre (_, _, fss, xssss) ctor_iter =
-      (mk_binding pre,
-       fold_rev (fold_rev Term.lambda) fss (Term.list_comb (ctor_iter,
-         map2 (mk_sum_caseN_balanced oo map2 mk_uncurried2_fun) fss xssss)));
-  in
-    define_co_iters Least_FP fpT Cs (map3 generate_iter iterNs iter_args_typess' ctor_iters) lthy
-  end;
-
-fun define_coiters coiterNs (_, cs, cpss, coiter_args_typess') mk_binding fpTs Cs dtor_coiters
-    lthy =
-  let
-    val nn = length fpTs;
-
-    val C_to_fpT as Type (_, [_, fpT]) = snd (strip_typeN nn (fastype_of (hd dtor_coiters)));
-
-    fun generate_coiter pre ((pfss, cqfsss), (f_sum_prod_Ts, pf_Tss)) dtor_coiter =
-      (mk_binding pre,
-       fold_rev (fold_rev Term.lambda) pfss (Term.list_comb (dtor_coiter,
-         map4 mk_preds_getterss_join cs cpss f_sum_prod_Ts cqfsss)));
-  in
-    define_co_iters Greatest_FP fpT Cs
-      (map3 generate_coiter coiterNs coiter_args_typess' dtor_coiters) lthy
-  end;
-
-fun derive_induct_iters_thms_for_types pre_bnfs [fold_args_types, rec_args_types] ctor_induct
-    ctor_iter_thmss nesting_bnfs nested_bnfs fpTs Cs Xs ctrXs_Tsss ctrss ctr_defss iterss iter_defss
-    lthy =
-  let
-    val iterss' = transpose iterss;
-    val iter_defss' = transpose iter_defss;
-
-    val [folds, recs] = iterss';
-    val [fold_defs, rec_defs] = iter_defss';
-
-    val ctr_Tsss = map (map (binder_types o fastype_of)) ctrss;
-
-    val nn = length pre_bnfs;
-    val ns = map length ctr_Tsss;
-    val mss = map (map length) ctr_Tsss;
-
-    val pre_map_defs = map map_def_of_bnf pre_bnfs;
-    val pre_set_defss = map set_defs_of_bnf pre_bnfs;
-    val nesting_map_idents = map (unfold_thms lthy [id_def] o map_id0_of_bnf) nesting_bnfs;
-    val nested_map_idents = map (unfold_thms lthy [id_def] o map_id0_of_bnf) nested_bnfs;
-    val nested_set_maps = maps set_map_of_bnf nested_bnfs;
-
-    val fp_b_names = map base_name_of_typ fpTs;
-
-    val ((((ps, ps'), xsss), us'), names_lthy) =
-      lthy
-      |> mk_Frees' "P" (map mk_pred1T fpTs)
-      ||>> mk_Freesss "x" ctr_Tsss
-      ||>> Variable.variant_fixes fp_b_names;
-
-    val us = map2 (curry Free) us' fpTs;
-
-    fun mk_sets_nested bnf =
-      let
-        val Type (T_name, Us) = T_of_bnf bnf;
-        val lives = lives_of_bnf bnf;
-        val sets = sets_of_bnf bnf;
-        fun mk_set U =
-          (case find_index (curry (op =) U) lives of
-            ~1 => Term.dummy
-          | i => nth sets i);
-      in
-        (T_name, map mk_set Us)
-      end;
-
-    val setss_nested = map mk_sets_nested nested_bnfs;
-
-    val (induct_thms, induct_thm) =
-      let
-        fun mk_set Ts t =
-          let val Type (_, Ts0) = domain_type (fastype_of t) in
-            Term.subst_atomic_types (Ts0 ~~ Ts) t
-          end;
-
-        fun mk_raw_prem_prems _ (x as Free (_, Type _)) (X as TFree _) =
-            [([], (find_index (curry (op =) X) Xs + 1, x))]
-          | mk_raw_prem_prems names_lthy (x as Free (s, Type (T_name, Ts0))) (Type (_, Xs_Ts0)) =
-            (case AList.lookup (op =) setss_nested T_name of
-              NONE => []
-            | SOME raw_sets0 =>
-              let
-                val (Xs_Ts, (Ts, raw_sets)) =
-                  filter (exists_subtype_in Xs o fst) (Xs_Ts0 ~~ (Ts0 ~~ raw_sets0))
-                  |> split_list ||> split_list;
-                val sets = map (mk_set Ts0) raw_sets;
-                val (ys, names_lthy') = names_lthy |> mk_Frees s Ts;
-                val xysets = map (pair x) (ys ~~ sets);
-                val ppremss = map2 (mk_raw_prem_prems names_lthy') ys Xs_Ts;
-              in
-                flat (map2 (map o apfst o cons) xysets ppremss)
-              end)
-          | mk_raw_prem_prems _ _ _ = [];
-
-        fun close_prem_prem xs t =
-          fold_rev Logic.all (map Free (drop (nn + length xs)
-            (rev (Term.add_frees t (map dest_Free xs @ ps'))))) t;
-
-        fun mk_prem_prem xs (xysets, (j, x)) =
-          close_prem_prem xs (Logic.list_implies (map (fn (x', (y, set)) =>
-              HOLogic.mk_Trueprop (HOLogic.mk_mem (y, set $ x'))) xysets,
-            HOLogic.mk_Trueprop (nth ps (j - 1) $ x)));
-
-        fun mk_raw_prem phi ctr ctr_Ts ctrXs_Ts =
-          let
-            val (xs, names_lthy') = names_lthy |> mk_Frees "x" ctr_Ts;
-            val pprems = flat (map2 (mk_raw_prem_prems names_lthy') xs ctrXs_Ts);
-          in (xs, pprems, HOLogic.mk_Trueprop (phi $ Term.list_comb (ctr, xs))) end;
-
-        fun mk_prem (xs, raw_pprems, concl) =
-          fold_rev Logic.all xs (Logic.list_implies (map (mk_prem_prem xs) raw_pprems, concl));
-
-        val raw_premss = map4 (map3 o mk_raw_prem) ps ctrss ctr_Tsss ctrXs_Tsss;
-
-        val goal =
-          Library.foldr (Logic.list_implies o apfst (map mk_prem)) (raw_premss,
-            HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry (op $)) ps us)));
-
-        val kksss = map (map (map (fst o snd) o #2)) raw_premss;
-
-        val ctor_induct' = ctor_induct OF (map mk_sumEN_tupled_balanced mss);
-
-        val thm =
-          Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
-            mk_induct_tac ctxt nn ns mss kksss (flat ctr_defss) ctor_induct' nested_set_maps
-              pre_set_defss)
-          |> singleton (Proof_Context.export names_lthy lthy)
-          |> Thm.close_derivation;
-      in
-        `(conj_dests nn) thm
-      end;
-
-    val induct_cases = quasi_unambiguous_case_names (maps (map name_of_ctr) ctrss);
-    val induct_case_names_attr = Attrib.internal (K (Rule_Cases.case_names induct_cases));
-
-    val xctrss = map2 (map2 (curry Term.list_comb)) ctrss xsss;
-
-    fun mk_iter_thmss (_, x_Tssss, fss, _) iters iter_defs ctor_iter_thms =
-      let
-        val fiters = map (lists_bmoc fss) iters;
-
-        fun mk_goal fss fiter xctr f xs fxs =
-          fold_rev (fold_rev Logic.all) (xs :: fss)
-            (mk_Trueprop_eq (fiter $ xctr, Term.list_comb (f, fxs)));
-
-        fun maybe_tick (T, U) u f =
-          if try (fst o HOLogic.dest_prodT) U = SOME T then
-            Term.lambda u (HOLogic.mk_prod (u, f $ u))
-          else
-            f;
-
-        fun build_iter (x as Free (_, T)) U =
-          if T = U then
-            x
-          else
-            build_map lthy (indexify (perhaps (try (snd o HOLogic.dest_prodT)) o snd) Cs
-              (fn kk => fn TU => maybe_tick TU (nth us kk) (nth fiters kk))) (T, U) $ x;
-
-        val fxsss = map2 (map2 (flat_rec_arg_args oo map2 (map o build_iter))) xsss x_Tssss;
-
-        val goalss = map5 (map4 o mk_goal fss) fiters xctrss fss xsss fxsss;
-
-        val tacss =
-          map2 (map o mk_iter_tac pre_map_defs (nested_map_idents @ nesting_map_idents) iter_defs)
-            ctor_iter_thms ctr_defss;
-
-        fun prove goal tac =
-          Goal.prove_sorry lthy [] [] goal (tac o #context)
-          |> Thm.close_derivation;
-      in
-        map2 (map2 prove) goalss tacss
-      end;
-
-    val fold_thmss = mk_iter_thmss fold_args_types folds fold_defs (map un_fold_of ctor_iter_thmss);
-    val rec_thmss = mk_iter_thmss rec_args_types recs rec_defs (map co_rec_of ctor_iter_thmss);
-  in
-    ((induct_thms, induct_thm, [induct_case_names_attr]),
-     (fold_thmss, rec_thmss, code_nitpicksimp_attrs @ simp_attrs))
-  end;
-
-fun derive_coinduct_coiters_thms_for_types pre_bnfs (z, cs, cpss,
-      coiters_args_types as [((pgss, crgsss), _), ((phss, cshsss), _)])
-    dtor_coinduct dtor_injects dtor_ctors dtor_coiter_thmss nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss
-    mss ns ctr_defss (ctr_sugars : ctr_sugar list) coiterss coiter_defss export_args lthy =
-  let
-    fun mk_ctor_dtor_coiter_thm dtor_inject dtor_ctor coiter =
-      iffD1 OF [dtor_inject, trans OF [coiter, dtor_ctor RS sym]];
-
-    val ctor_dtor_coiter_thmss =
-      map3 (map oo mk_ctor_dtor_coiter_thm) dtor_injects dtor_ctors dtor_coiter_thmss;
-
-    val coiterss' = transpose coiterss;
-    val coiter_defss' = transpose coiter_defss;
-
-    val [unfold_defs, corec_defs] = coiter_defss';
-
-    val nn = length pre_bnfs;
-
-    val pre_map_defs = map map_def_of_bnf pre_bnfs;
-    val pre_rel_defs = map rel_def_of_bnf pre_bnfs;
-    val nesting_map_idents = map (unfold_thms lthy [id_def] o map_id0_of_bnf) nesting_bnfs;
-    val nesting_rel_eqs = map rel_eq_of_bnf nesting_bnfs;
-
-    val fp_b_names = map base_name_of_typ fpTs;
-
-    val ctrss = map #ctrs ctr_sugars;
-    val discss = map #discs ctr_sugars;
-    val selsss = map #selss ctr_sugars;
-    val exhausts = map #exhaust ctr_sugars;
-    val disc_thmsss = map #disc_thmss ctr_sugars;
-    val discIss = map #discIs ctr_sugars;
-    val sel_thmsss = map #sel_thmss ctr_sugars;
-
-    val (((rs, us'), vs'), names_lthy) =
-      lthy
-      |> mk_Frees "R" (map (fn T => mk_pred2T T T) fpTs)
-      ||>> Variable.variant_fixes fp_b_names
-      ||>> Variable.variant_fixes (map (suffix "'") fp_b_names);
-
-    val us = map2 (curry Free) us' fpTs;
-    val udiscss = map2 (map o rapp) us discss;
-    val uselsss = map2 (map o map o rapp) us selsss;
-
-    val vs = map2 (curry Free) vs' fpTs;
-    val vdiscss = map2 (map o rapp) vs discss;
-    val vselsss = map2 (map o map o rapp) vs selsss;
-
-    val coinduct_thms_pairs =
-      let
-        val uvrs = map3 (fn r => fn u => fn v => r $ u $ v) rs us vs;
-        val uv_eqs = map2 (curry HOLogic.mk_eq) us vs;
-        val strong_rs =
-          map4 (fn u => fn v => fn uvr => fn uv_eq =>
-            fold_rev Term.lambda [u, v] (HOLogic.mk_disj (uvr, uv_eq))) us vs uvrs uv_eqs;
-
-        fun build_the_rel rs' T Xs_T =
-          build_rel lthy (fn (_, X) => nth rs' (find_index (curry (op =) X) Xs)) (T, Xs_T)
-          |> Term.subst_atomic_types (Xs ~~ fpTs);
-
-        fun build_rel_app rs' usel vsel Xs_T =
-          fold rapp [usel, vsel] (build_the_rel rs' (fastype_of usel) Xs_T);
-
-        fun mk_prem_ctr_concls rs' n k udisc usels vdisc vsels ctrXs_Ts =
-          (if k = n then [] else [HOLogic.mk_eq (udisc, vdisc)]) @
-          (if null usels then
-             []
-           else
-             [Library.foldr HOLogic.mk_imp (if n = 1 then [] else [udisc, vdisc],
-                Library.foldr1 HOLogic.mk_conj (map3 (build_rel_app rs') usels vsels ctrXs_Ts))]);
-
-        fun mk_prem_concl rs' n udiscs uselss vdiscs vselss ctrXs_Tss =
-          Library.foldr1 HOLogic.mk_conj (flat (map6 (mk_prem_ctr_concls rs' n)
-            (1 upto n) udiscs uselss vdiscs vselss ctrXs_Tss))
-          handle List.Empty => @{term True};
-
-        fun mk_prem rs' uvr u v n udiscs uselss vdiscs vselss ctrXs_Tss =
-          fold_rev Logic.all [u, v] (Logic.mk_implies (HOLogic.mk_Trueprop uvr,
-            HOLogic.mk_Trueprop (mk_prem_concl rs' n udiscs uselss vdiscs vselss ctrXs_Tss)));
-
-        val concl =
-          HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
-            (map3 (fn uvr => fn u => fn v => HOLogic.mk_imp (uvr, HOLogic.mk_eq (u, v)))
-               uvrs us vs));
-
-        fun mk_goal rs' =
-          Logic.list_implies (map9 (mk_prem rs') uvrs us vs ns udiscss uselsss vdiscss vselsss
-            ctrXs_Tsss, concl);
-
-        val goals = map mk_goal [rs, strong_rs];
-
-        fun prove dtor_coinduct' goal =
-          Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
-            mk_coinduct_tac ctxt nesting_rel_eqs nn ns dtor_coinduct' pre_rel_defs dtor_ctors
-              exhausts ctr_defss disc_thmsss sel_thmsss)
-          |> singleton (Proof_Context.export names_lthy lthy)
-          |> Thm.close_derivation;
-
-        fun postproc nn thm =
-          Thm.permute_prems 0 nn
-            (if nn = 1 then thm RS mp else funpow nn (fn thm => reassoc_conjs (thm RS mp_conj)) thm)
-          |> Drule.zero_var_indexes
-          |> `(conj_dests nn);
-
-        val rel_eqs = map rel_eq_of_bnf pre_bnfs;
-        val rel_monos = map rel_mono_of_bnf pre_bnfs;
-        val dtor_coinducts =
-          [dtor_coinduct, mk_strong_coinduct_thm dtor_coinduct rel_eqs rel_monos lthy];
-      in
-        map2 (postproc nn oo prove) dtor_coinducts goals
-      end;
-
-    fun mk_coinduct_concls ms discs ctrs =
-      let
-        fun mk_disc_concl disc = [name_of_disc disc];
-        fun mk_ctr_concl 0 _ = []
-          | mk_ctr_concl _ ctor = [name_of_ctr ctor];
-        val disc_concls = map mk_disc_concl (fst (split_last discs)) @ [[]];
-        val ctr_concls = map2 mk_ctr_concl ms ctrs;
-      in
-        flat (map2 append disc_concls ctr_concls)
-      end;
-
-    val coinduct_cases = quasi_unambiguous_case_names (map (prefix EqN) fp_b_names);
-    val coinduct_conclss =
-      map3 (quasi_unambiguous_case_names ooo mk_coinduct_concls) mss discss ctrss;
-
-    fun mk_maybe_not pos = not pos ? HOLogic.mk_not;
-
-    val fcoiterss' as [gunfolds, hcorecs] =
-      map2 (fn (pfss, _) => map (lists_bmoc pfss)) (map fst coiters_args_types) coiterss';
-
-    val (unfold_thmss, corec_thmss) =
-      let
-        fun mk_goal pfss c cps fcoiter n k ctr m cfs' =
-          fold_rev (fold_rev Logic.all) ([c] :: pfss)
-            (Logic.list_implies (seq_conds (HOLogic.mk_Trueprop oo mk_maybe_not) n k cps,
-               mk_Trueprop_eq (fcoiter $ c, Term.list_comb (ctr, take m cfs'))));
-
-        fun mk_U maybe_mk_sumT =
-          typ_subst_nonatomic (map2 (fn C => fn fpT => (maybe_mk_sumT fpT C, fpT)) Cs fpTs);
-
-        fun tack z_name (c, u) f =
-          let val z = Free (z_name, mk_sumT (fastype_of u, fastype_of c)) in
-            Term.lambda z (mk_sum_case (Term.lambda u u, Term.lambda c (f $ c)) $ z)
-          end;
-
-        fun build_coiter fcoiters maybe_mk_sumT maybe_tack cqf =
-          let val T = fastype_of cqf in
-            if exists_subtype_in Cs T then
-              let val U = mk_U maybe_mk_sumT T in
-                build_map lthy (indexify snd fpTs (fn kk => fn _ =>
-                  maybe_tack (nth cs kk, nth us kk) (nth fcoiters kk))) (T, U) $ cqf
-              end
-            else
-              cqf
-          end;
-
-        val crgsss' = map (map (map (build_coiter (un_fold_of fcoiterss') (K I) (K I)))) crgsss;
-        val cshsss' = map (map (map (build_coiter (co_rec_of fcoiterss') (curry mk_sumT) (tack z))))
-          cshsss;
-
-        val unfold_goalss = map8 (map4 oooo mk_goal pgss) cs cpss gunfolds ns kss ctrss mss crgsss';
-        val corec_goalss = map8 (map4 oooo mk_goal phss) cs cpss hcorecs ns kss ctrss mss cshsss';
-
-        val unfold_tacss =
-          map3 (map oo mk_coiter_tac unfold_defs nesting_map_idents)
-            (map un_fold_of ctor_dtor_coiter_thmss) pre_map_defs ctr_defss;
-        val corec_tacss =
-          map3 (map oo mk_coiter_tac corec_defs nesting_map_idents)
-            (map co_rec_of ctor_dtor_coiter_thmss) pre_map_defs ctr_defss;
-
-        fun prove goal tac =
-          Goal.prove_sorry lthy [] [] goal (tac o #context)
-          |> Thm.close_derivation;
-
-        val unfold_thmss = map2 (map2 prove) unfold_goalss unfold_tacss;
-        val corec_thmss =
-          map2 (map2 prove) corec_goalss corec_tacss
-          |> map (map (unfold_thms lthy @{thms sum_case_if}));
-      in
-        (unfold_thmss, corec_thmss)
-      end;
-
-    val (disc_unfold_iff_thmss, disc_corec_iff_thmss) =
-      let
-        fun mk_goal c cps fcoiter n k disc =
-          mk_Trueprop_eq (disc $ (fcoiter $ c),
-            if n = 1 then @{const True}
-            else Library.foldr1 HOLogic.mk_conj (seq_conds mk_maybe_not n k cps));
-
-        val unfold_goalss = map6 (map2 oooo mk_goal) cs cpss gunfolds ns kss discss;
-        val corec_goalss = map6 (map2 oooo mk_goal) cs cpss hcorecs ns kss discss;
-
-        fun mk_case_split' cp = Drule.instantiate' [] [SOME (certify lthy cp)] @{thm case_split};
-
-        val case_splitss' = map (map mk_case_split') cpss;
-
-        val unfold_tacss =
-          map3 (map oo mk_disc_coiter_iff_tac) case_splitss' unfold_thmss disc_thmsss;
-        val corec_tacss =
-          map3 (map oo mk_disc_coiter_iff_tac) case_splitss' corec_thmss disc_thmsss;
-
-        fun prove goal tac =
-          Goal.prove_sorry lthy [] [] goal (tac o #context)
-          |> singleton export_args
-          |> singleton (Proof_Context.export names_lthy lthy)
-          |> Thm.close_derivation;
-
-        fun proves [_] [_] = []
-          | proves goals tacs = map2 prove goals tacs;
-      in
-        (map2 proves unfold_goalss unfold_tacss, map2 proves corec_goalss corec_tacss)
-      end;
-
-    fun mk_disc_coiter_thms coiters discIs = map (op RS) (coiters ~~ discIs);
-
-    val disc_unfold_thmss = map2 mk_disc_coiter_thms unfold_thmss discIss;
-    val disc_corec_thmss = map2 mk_disc_coiter_thms corec_thmss discIss;
-
-    fun mk_sel_coiter_thm coiter_thm sel sel_thm =
-      let
-        val (domT, ranT) = dest_funT (fastype_of sel);
-        val arg_cong' =
-          Drule.instantiate' (map (SOME o certifyT lthy) [domT, ranT])
-            [NONE, NONE, SOME (certify lthy sel)] arg_cong
-          |> Thm.varifyT_global;
-        val sel_thm' = sel_thm RSN (2, trans);
-      in
-        coiter_thm RS arg_cong' RS sel_thm'
-      end;
-
-    fun mk_sel_coiter_thms coiter_thmss =
-      map3 (map3 (map2 o mk_sel_coiter_thm)) coiter_thmss selsss sel_thmsss;
-
-    val sel_unfold_thmsss = mk_sel_coiter_thms unfold_thmss;
-    val sel_corec_thmsss = mk_sel_coiter_thms corec_thmss;
-
-    val coinduct_consumes_attr = Attrib.internal (K (Rule_Cases.consumes nn));
-    val coinduct_case_names_attr = Attrib.internal (K (Rule_Cases.case_names coinduct_cases));
-    val coinduct_case_concl_attrs =
-      map2 (fn casex => fn concls =>
-          Attrib.internal (K (Rule_Cases.case_conclusion (casex, concls))))
-        coinduct_cases coinduct_conclss;
-    val coinduct_case_attrs =
-      coinduct_consumes_attr :: coinduct_case_names_attr :: coinduct_case_concl_attrs;
-  in
-    ((coinduct_thms_pairs, coinduct_case_attrs),
-     (unfold_thmss, corec_thmss, code_nitpicksimp_attrs),
-     (disc_unfold_thmss, disc_corec_thmss, []),
-     (disc_unfold_iff_thmss, disc_corec_iff_thmss, simp_attrs),
-     (sel_unfold_thmsss, sel_corec_thmsss, simp_attrs))
-  end;
-
-fun define_co_datatypes prepare_constraint prepare_typ prepare_term fp construct_fp
-    (wrap_opts as (no_discs_sels, (_, rep_compat)), specs) no_defs_lthy0 =
-  let
-    (* TODO: sanity checks on arguments *)
-
-    val _ = if fp = Greatest_FP andalso no_discs_sels then
-        error "Cannot define codatatypes without discriminators and selectors"
-      else
-        ();
-
-    fun qualify mandatory fp_b_name =
-      Binding.qualify mandatory fp_b_name o (rep_compat ? Binding.qualify false rep_compat_prefix);
-
-    val nn = length specs;
-    val fp_bs = map type_binding_of specs;
-    val fp_b_names = map Binding.name_of fp_bs;
-    val fp_common_name = mk_common_name fp_b_names;
-    val map_bs = map map_binding_of specs;
-    val rel_bs = map rel_binding_of specs;
-
-    fun prepare_type_arg (_, (ty, c)) =
-      let val TFree (s, _) = prepare_typ no_defs_lthy0 ty in
-        TFree (s, prepare_constraint no_defs_lthy0 c)
-      end;
-
-    val Ass0 = map (map prepare_type_arg o type_args_named_constrained_of) specs;
-    val unsorted_Ass0 = map (map (resort_tfree HOLogic.typeS)) Ass0;
-    val unsorted_As = Library.foldr1 merge_type_args unsorted_Ass0;
-    val num_As = length unsorted_As;
-    val set_bss = map (map fst o type_args_named_constrained_of) specs;
-
-    val (((Bs0, Cs), Xs), no_defs_lthy) =
-      no_defs_lthy0
-      |> fold (Variable.declare_typ o resort_tfree dummyS) unsorted_As
-      |> mk_TFrees num_As
-      ||>> mk_TFrees nn
-      ||>> variant_tfrees fp_b_names;
-
-    fun add_fake_type spec = Typedecl.basic_typedecl (type_binding_of spec, num_As, mixfix_of spec);
-
-    val (fake_T_names, fake_lthy) = fold_map add_fake_type specs no_defs_lthy0;
-
-    val qsoty = quote o Syntax.string_of_typ fake_lthy;
-
-    val _ = (case Library.duplicates (op =) unsorted_As of [] => ()
-      | A :: _ => error ("Duplicate type parameter " ^ qsoty A ^ " in " ^ co_prefix fp ^
-          "datatype specification"));
-
-    val bad_args =
-      map (Logic.type_map (singleton (Variable.polymorphic no_defs_lthy0))) unsorted_As
-      |> filter_out Term.is_TVar;
-    val _ = null bad_args orelse
-      error ("Locally fixed type argument " ^ qsoty (hd bad_args) ^ " in " ^ co_prefix fp ^
-        "datatype specification");
-
-    val mixfixes = map mixfix_of specs;
-
-    val _ = (case Library.duplicates Binding.eq_name fp_bs of [] => ()
-      | b :: _ => error ("Duplicate type name declaration " ^ quote (Binding.name_of b)));
-
-    val ctr_specss = map ctr_specs_of specs;
-
-    val disc_bindingss = map (map disc_of) ctr_specss;
-    val ctr_bindingss =
-      map2 (fn fp_b_name => map (qualify false fp_b_name o ctr_of)) fp_b_names ctr_specss;
-    val ctr_argsss = map (map args_of) ctr_specss;
-    val ctr_mixfixess = map (map ctr_mixfix_of) ctr_specss;
-
-    val sel_bindingsss = map (map (map fst)) ctr_argsss;
-    val fake_ctr_Tsss0 = map (map (map (prepare_typ fake_lthy o snd))) ctr_argsss;
-    val raw_sel_defaultsss = map (map defaults_of) ctr_specss;
-
-    val (As :: _) :: fake_ctr_Tsss =
-      burrow (burrow (Syntax.check_typs fake_lthy)) (Ass0 :: fake_ctr_Tsss0);
-    val As' = map dest_TFree As;
-
-    val rhs_As' = fold (fold (fold Term.add_tfreesT)) fake_ctr_Tsss [];
-    val _ = (case subtract (op =) As' rhs_As' of [] => ()
-      | extras => error ("Extra type variables on right-hand side: " ^
-          commas (map (qsoty o TFree) extras)));
-
-    val fake_Ts = map (fn s => Type (s, As)) fake_T_names;
-
-    fun eq_fpT_check (T as Type (s, Ts)) (T' as Type (s', Ts')) =
-        s = s' andalso (Ts = Ts' orelse
-          error ("Wrong type arguments in " ^ co_prefix fp ^ "recursive type " ^ qsoty T ^
-            " (expected " ^ qsoty T' ^ ")"))
-      | eq_fpT_check _ _ = false;
-
-    fun freeze_fp (T as Type (s, Ts)) =
-        (case find_index (eq_fpT_check T) fake_Ts of
-          ~1 => Type (s, map freeze_fp Ts)
-        | kk => nth Xs kk)
-      | freeze_fp T = T;
-
-    val unfreeze_fp = Term.typ_subst_atomic (Xs ~~ fake_Ts);
-
-    val ctrXs_Tsss = map (map (map freeze_fp)) fake_ctr_Tsss;
-    val ctrXs_sum_prod_Ts = map (mk_sumTN_balanced o map HOLogic.mk_tupleT) ctrXs_Tsss;
-
-    val fp_eqs =
-      map dest_TFree Xs ~~ map (Term.typ_subst_atomic (As ~~ unsorted_As)) ctrXs_sum_prod_Ts;
-
-    val rhsXs_As' = fold (fold (fold Term.add_tfreesT)) ctrXs_Tsss [];
-    val _ = (case subtract (op =) rhsXs_As' As' of [] => ()
-      | extras => List.app (fn extra => warning ("Unused type variable on right-hand side of " ^
-          co_prefix fp ^ "datatype definition: " ^ qsoty (TFree extra))) extras);
-
-    val (pre_bnfs, (fp_res as {bnfs = fp_bnfs as any_fp_bnf :: _, ctors = ctors0, dtors = dtors0,
-           xtor_co_iterss = xtor_co_iterss0, xtor_co_induct, dtor_ctors, ctor_dtors, ctor_injects,
-           dtor_injects, xtor_map_thms, xtor_set_thmss, xtor_rel_thms, xtor_co_iter_thmss, ...},
-           lthy)) =
-      fp_bnf (construct_fp mixfixes map_bs rel_bs set_bss) fp_bs (map dest_TFree unsorted_As) fp_eqs
-        no_defs_lthy0
-      handle BAD_DEAD (X, X_backdrop) =>
-        (case X_backdrop of
-          Type (bad_tc, _) =>
-          let
-            val fake_T = qsoty (unfreeze_fp X);
-            val fake_T_backdrop = qsoty (unfreeze_fp X_backdrop);
-            fun register_hint () =
-              "\nUse the " ^ quote (fst (fst @{command_spec "bnf"})) ^ " command to register " ^
-              quote bad_tc ^ " as a bounded natural functor to allow nested (co)recursion through \
-              \it";
-          in
-            if is_some (bnf_of no_defs_lthy bad_tc) orelse
-               is_some (fp_sugar_of no_defs_lthy bad_tc) then
-              error ("Inadmissible " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^
-                " in type expression " ^ fake_T_backdrop)
-            else if is_some (Datatype_Data.get_info (Proof_Context.theory_of no_defs_lthy)
-                bad_tc) then
-              error ("Unsupported " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^
-                " via the old-style datatype " ^ quote bad_tc ^ " in type expression " ^
-                fake_T_backdrop ^ register_hint ())
-            else
-              error ("Unsupported " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^
-                " via type constructor " ^ quote bad_tc ^ " in type expression " ^ fake_T_backdrop ^
-                register_hint ())
-          end);
-
-    val time = time lthy;
-    val timer = time (Timer.startRealTimer ());
-
-    val nesting_bnfs = nesty_bnfs lthy ctrXs_Tsss As;
-    val nested_bnfs = nesty_bnfs lthy ctrXs_Tsss Xs;
-
-    val pre_map_defs = map map_def_of_bnf pre_bnfs;
-    val pre_set_defss = map set_defs_of_bnf pre_bnfs;
-    val pre_rel_defs = map rel_def_of_bnf pre_bnfs;
-    val nesting_set_maps = maps set_map_of_bnf nesting_bnfs;
-    val nested_set_maps = maps set_map_of_bnf nested_bnfs;
-
-    val live = live_of_bnf any_fp_bnf;
-    val _ =
-      if live = 0 andalso exists (not o Binding.is_empty) (map_bs @ rel_bs) then
-        warning "Map function and relator names ignored"
-      else
-        ();
-
-    val Bs =
-      map3 (fn alive => fn A as TFree (_, S) => fn B => if alive then resort_tfree S B else A)
-        (liveness_of_fp_bnf num_As any_fp_bnf) As Bs0;
-
-    val B_ify = Term.typ_subst_atomic (As ~~ Bs);
-
-    val ctors = map (mk_ctor As) ctors0;
-    val dtors = map (mk_dtor As) dtors0;
-
-    val fpTs = map (domain_type o fastype_of) dtors;
-
-    fun massage_simple_notes base =
-      filter_out (null o #2)
-      #> map (fn (thmN, thms, attrs) =>
-        ((qualify true base (Binding.name thmN), attrs), [(thms, [])]));
-
-    val massage_multi_notes =
-      maps (fn (thmN, thmss, attrs) =>
-        map3 (fn fp_b_name => fn Type (T_name, _) => fn thms =>
-            ((qualify true fp_b_name (Binding.name thmN), attrs T_name), [(thms, [])]))
-          fp_b_names fpTs thmss)
-      #> filter_out (null o fst o hd o snd);
-
-    val ctr_Tsss = map (map (map (Term.typ_subst_atomic (Xs ~~ fpTs)))) ctrXs_Tsss;
-    val ns = map length ctr_Tsss;
-    val kss = map (fn n => 1 upto n) ns;
-    val mss = map (map length) ctr_Tsss;
-
-    val ((xtor_co_iterss, iters_args_types, coiters_args_types), lthy') =
-      mk_co_iters_prelims fp ctr_Tsss fpTs Cs ns mss xtor_co_iterss0 lthy;
-
-    fun define_ctrs_dtrs_for_type (((((((((((((((((((((((fp_bnf, fp_b), fpT), ctor), dtor),
-            xtor_co_iters), ctor_dtor), dtor_ctor), ctor_inject), pre_map_def), pre_set_defs),
-          pre_rel_def), fp_map_thm), fp_set_thms), fp_rel_thm), n), ks), ms), ctr_bindings),
-        ctr_mixfixes), ctr_Tss), disc_bindings), sel_bindingss), raw_sel_defaultss) no_defs_lthy =
-      let
-        val fp_b_name = Binding.name_of fp_b;
-
-        val dtorT = domain_type (fastype_of ctor);
-        val ctr_prod_Ts = map HOLogic.mk_tupleT ctr_Tss;
-        val ctr_sum_prod_T = mk_sumTN_balanced ctr_prod_Ts;
-
-        val ((((w, xss), yss), u'), names_lthy) =
-          no_defs_lthy
-          |> yield_singleton (mk_Frees "w") dtorT
-          ||>> mk_Freess "x" ctr_Tss
-          ||>> mk_Freess "y" (map (map B_ify) ctr_Tss)
-          ||>> yield_singleton Variable.variant_fixes fp_b_name;
-
-        val u = Free (u', fpT);
-
-        val tuple_xs = map HOLogic.mk_tuple xss;
-        val tuple_ys = map HOLogic.mk_tuple yss;
-
-        val ctr_rhss =
-          map3 (fn k => fn xs => fn tuple_x => fold_rev Term.lambda xs (ctor $
-            mk_InN_balanced ctr_sum_prod_T n tuple_x k)) ks xss tuple_xs;
-
-        val maybe_conceal_def_binding = Thm.def_binding
-          #> Config.get no_defs_lthy bnf_note_all = false ? Binding.conceal;
-
-        val ((raw_ctrs, raw_ctr_defs), (lthy', lthy)) = no_defs_lthy
-          |> apfst split_list o fold_map3 (fn b => fn mx => fn rhs =>
-              Local_Theory.define ((b, mx), ((maybe_conceal_def_binding b, []), rhs)) #>> apsnd snd)
-            ctr_bindings ctr_mixfixes ctr_rhss
-          ||> `Local_Theory.restore;
-
-        val phi = Proof_Context.export_morphism lthy lthy';
-
-        val ctr_defs = map (Morphism.thm phi) raw_ctr_defs;
-        val ctr_defs' =
-          map2 (fn m => fn def => mk_unabs_def m (def RS meta_eq_to_obj_eq)) ms ctr_defs;
-
-        val ctrs0 = map (Morphism.term phi) raw_ctrs;
-        val ctrs = map (mk_ctr As) ctrs0;
-
-        fun wrap_ctrs lthy =
-          let
-            fun exhaust_tac {context = ctxt, prems = _} =
-              let
-                val ctor_iff_dtor_thm =
-                  let
-                    val goal =
-                      fold_rev Logic.all [w, u]
-                        (mk_Trueprop_eq (HOLogic.mk_eq (u, ctor $ w), HOLogic.mk_eq (dtor $ u, w)));
-                  in
-                    Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
-                      mk_ctor_iff_dtor_tac ctxt (map (SOME o certifyT lthy) [dtorT, fpT])
-                        (certify lthy ctor) (certify lthy dtor) ctor_dtor dtor_ctor)
-                    |> Thm.close_derivation
-                    |> Morphism.thm phi
-                  end;
-
-                val sumEN_thm' =
-                  unfold_thms lthy @{thms unit_all_eq1}
-                    (Drule.instantiate' (map (SOME o certifyT lthy) ctr_prod_Ts) []
-                       (mk_sumEN_balanced n))
-                  |> Morphism.thm phi;
-              in
-                mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor_thm sumEN_thm'
-              end;
-
-            val inject_tacss =
-              map2 (fn 0 => K [] | _ => fn ctr_def => [fn {context = ctxt, ...} =>
-                mk_inject_tac ctxt ctr_def ctor_inject]) ms ctr_defs;
-
-            val half_distinct_tacss =
-              map (map (fn (def, def') => fn {context = ctxt, ...} =>
-                mk_half_distinct_tac ctxt ctor_inject [def, def'])) (mk_half_pairss (`I ctr_defs));
-
-            val tacss = [exhaust_tac] :: inject_tacss @ half_distinct_tacss;
-
-            val sel_defaultss = map (map (apsnd (prepare_term lthy))) raw_sel_defaultss
-          in
-            wrap_free_constructors tacss (((wrap_opts, ctrs0), standard_binding), (disc_bindings,
-              (sel_bindingss, sel_defaultss))) lthy
-          end;
-
-        fun derive_maps_sets_rels (ctr_sugar, lthy) =
-          if live = 0 then
-            ((([], [], [], []), ctr_sugar), lthy)
-          else
-            let
-              val rel_flip = rel_flip_of_bnf fp_bnf;
-              val nones = replicate live NONE;
-
-              val ctor_cong =
-                if fp = Least_FP then
-                  Drule.dummy_thm
-                else
-                  let val ctor' = mk_ctor Bs ctor in
-                    cterm_instantiate_pos [NONE, NONE, SOME (certify lthy ctor')] arg_cong
-                  end;
-
-              fun mk_cIn ify =
-                certify lthy o (fp = Greatest_FP ? curry (op $) (map_types ify ctor)) oo
-                mk_InN_balanced (ify ctr_sum_prod_T) n;
-
-              val cxIns = map2 (mk_cIn I) tuple_xs ks;
-              val cyIns = map2 (mk_cIn B_ify) tuple_ys ks;
-
-              fun mk_map_thm ctr_def' cxIn =
-                fold_thms lthy [ctr_def']
-                  (unfold_thms lthy (pre_map_def ::
-                       (if fp = Least_FP then [] else [ctor_dtor, dtor_ctor]) @ sum_prod_thms_map)
-                     (cterm_instantiate_pos (nones @ [SOME cxIn])
-                        (if fp = Least_FP then fp_map_thm else fp_map_thm RS ctor_cong)))
-                |> singleton (Proof_Context.export names_lthy no_defs_lthy);
-
-              fun mk_set_thm fp_set_thm ctr_def' cxIn =
-                fold_thms lthy [ctr_def']
-                  (unfold_thms lthy (pre_set_defs @ nested_set_maps @ nesting_set_maps @
-                       (if fp = Least_FP then [] else [dtor_ctor]) @ sum_prod_thms_set)
-                     (cterm_instantiate_pos [SOME cxIn] fp_set_thm))
-                |> singleton (Proof_Context.export names_lthy no_defs_lthy);
-
-              fun mk_set_thms fp_set_thm = map2 (mk_set_thm fp_set_thm) ctr_defs' cxIns;
-
-              val map_thms = map2 mk_map_thm ctr_defs' cxIns;
-              val set_thmss = map mk_set_thms fp_set_thms;
-
-              val rel_infos = (ctr_defs' ~~ cxIns, ctr_defs' ~~ cyIns);
-
-              fun mk_rel_thm postproc ctr_defs' cxIn cyIn =
-                fold_thms lthy ctr_defs'
-                  (unfold_thms lthy (@{thm Inl_Inr_False} :: pre_rel_def ::
-                       (if fp = Least_FP then [] else [dtor_ctor]) @ sum_prod_thms_rel)
-                     (cterm_instantiate_pos (nones @ [SOME cxIn, SOME cyIn]) fp_rel_thm))
-                |> postproc
-                |> singleton (Proof_Context.export names_lthy no_defs_lthy);
-
-              fun mk_rel_inject_thm ((ctr_def', cxIn), (_, cyIn)) =
-                mk_rel_thm (unfold_thms lthy @{thms eq_sym_Unity_conv}) [ctr_def'] cxIn cyIn;
-
-              val rel_inject_thms = map mk_rel_inject_thm (op ~~ rel_infos);
-
-              fun mk_half_rel_distinct_thm ((xctr_def', cxIn), (yctr_def', cyIn)) =
-                mk_rel_thm (fn thm => thm RS @{thm eq_False[THEN iffD1]}) [xctr_def', yctr_def']
-                  cxIn cyIn;
-
-              fun mk_other_half_rel_distinct_thm thm =
-                flip_rels lthy live thm
-                RS (rel_flip RS sym RS @{thm arg_cong[of _ _ Not]} RS iffD2);
-
-              val half_rel_distinct_thmss =
-                map (map mk_half_rel_distinct_thm) (mk_half_pairss rel_infos);
-              val other_half_rel_distinct_thmss =
-                map (map mk_other_half_rel_distinct_thm) half_rel_distinct_thmss;
-              val (rel_distinct_thms, _) =
-                join_halves n half_rel_distinct_thmss other_half_rel_distinct_thmss;
-
-              val anonymous_notes =
-                [(map (fn th => th RS @{thm eq_False[THEN iffD2]}) rel_distinct_thms,
-                  code_nitpicksimp_attrs),
-                 (map2 (fn th => fn 0 => th RS @{thm eq_True[THEN iffD2]} | _ => th)
-                    rel_inject_thms ms, code_nitpicksimp_attrs)]
-                |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
-
-              val notes =
-                [(mapN, map_thms, code_nitpicksimp_attrs @ simp_attrs),
-                 (rel_distinctN, rel_distinct_thms, simp_attrs),
-                 (rel_injectN, rel_inject_thms, simp_attrs),
-                 (setN, flat set_thmss, code_nitpicksimp_attrs @ simp_attrs)]
-                |> massage_simple_notes fp_b_name;
-            in
-              (((map_thms, rel_inject_thms, rel_distinct_thms, set_thmss), ctr_sugar),
-               lthy |> Local_Theory.notes (anonymous_notes @ notes) |> snd)
-            end;
-
-        fun mk_binding pre = qualify false fp_b_name (Binding.prefix_name (pre ^ "_") fp_b);
-
-        fun massage_res (((maps_sets_rels, ctr_sugar), co_iter_res), lthy) =
-          (((maps_sets_rels, (ctrs, xss, ctr_defs, ctr_sugar)), co_iter_res), lthy);
-      in
-        (wrap_ctrs
-         #> derive_maps_sets_rels
-         ##>>
-           (if fp = Least_FP then define_iters [foldN, recN] (the iters_args_types)
-            else define_coiters [unfoldN, corecN] (the coiters_args_types))
-             mk_binding fpTs Cs xtor_co_iters
-         #> massage_res, lthy')
-      end;
-
-    fun wrap_types_etc (wrap_types_etcs, lthy) =
-      fold_map I wrap_types_etcs lthy
-      |>> apsnd split_list o apfst (apsnd split_list4 o apfst split_list4 o split_list)
-        o split_list;
-
-    fun mk_simp_thms ({injects, distincts, case_thms, ...} : ctr_sugar) un_folds co_recs
-        mapsx rel_injects rel_distincts setss =
-      injects @ distincts @ case_thms @ co_recs @ un_folds @ mapsx @ rel_injects @ rel_distincts
-      @ flat setss;
-
-    fun derive_note_induct_iters_thms_for_types
-        ((((mapss, rel_injects, rel_distincts, setss), (ctrss, _, ctr_defss, ctr_sugars)),
-          (iterss, iter_defss)), lthy) =
-      let
-        val ((induct_thms, induct_thm, induct_attrs), (fold_thmss, rec_thmss, iter_attrs)) =
-          derive_induct_iters_thms_for_types pre_bnfs (the iters_args_types) xtor_co_induct
-            xtor_co_iter_thmss nesting_bnfs nested_bnfs fpTs Cs Xs ctrXs_Tsss ctrss ctr_defss iterss
-            iter_defss lthy;
-
-        val induct_type_attr = Attrib.internal o K o Induct.induct_type;
-
-        val simp_thmss =
-          map7 mk_simp_thms ctr_sugars fold_thmss rec_thmss mapss rel_injects rel_distincts setss;
-
-        val common_notes =
-          (if nn > 1 then [(inductN, [induct_thm], induct_attrs)] else [])
-          |> massage_simple_notes fp_common_name;
-
-        val notes =
-          [(foldN, fold_thmss, K iter_attrs),
-           (inductN, map single induct_thms, fn T_name => induct_attrs @ [induct_type_attr T_name]),
-           (recN, rec_thmss, K iter_attrs),
-           (simpsN, simp_thmss, K [])]
-          |> massage_multi_notes;
-      in
-        lthy
-        |> Local_Theory.notes (common_notes @ notes) |> snd
-        |> register_fp_sugars Least_FP pre_bnfs nested_bnfs nesting_bnfs fp_res ctr_defss ctr_sugars
-          iterss mapss [induct_thm] (transpose [fold_thmss, rec_thmss]) [] []
-      end;
-
-    fun derive_note_coinduct_coiters_thms_for_types
-        ((((mapss, rel_injects, rel_distincts, setss), (_, _, ctr_defss, ctr_sugars)),
-          (coiterss, coiter_defss)), lthy) =
-      let
-        val (([(coinduct_thms, coinduct_thm), (strong_coinduct_thms, strong_coinduct_thm)],
-              coinduct_attrs),
-             (unfold_thmss, corec_thmss, coiter_attrs),
-             (disc_unfold_thmss, disc_corec_thmss, disc_coiter_attrs),
-             (disc_unfold_iff_thmss, disc_corec_iff_thmss, disc_coiter_iff_attrs),
-             (sel_unfold_thmsss, sel_corec_thmsss, sel_coiter_attrs)) =
-          derive_coinduct_coiters_thms_for_types pre_bnfs (the coiters_args_types) xtor_co_induct
-            dtor_injects dtor_ctors xtor_co_iter_thmss nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss ns
-            ctr_defss ctr_sugars coiterss coiter_defss (Proof_Context.export lthy' no_defs_lthy)
-            lthy;
-
-        val sel_unfold_thmss = map flat sel_unfold_thmsss;
-        val sel_corec_thmss = map flat sel_corec_thmsss;
-
-        val coinduct_type_attr = Attrib.internal o K o Induct.coinduct_type;
-
-        val flat_coiter_thms = append oo append;
-
-        val simp_thmss =
-          map7 mk_simp_thms ctr_sugars
-            (map3 flat_coiter_thms disc_unfold_thmss disc_unfold_iff_thmss sel_unfold_thmss)
-            (map3 flat_coiter_thms disc_corec_thmss disc_corec_iff_thmss sel_corec_thmss)
-            mapss rel_injects rel_distincts setss;
-
-        val common_notes =
-          (if nn > 1 then
-             [(coinductN, [coinduct_thm], coinduct_attrs),
-              (strong_coinductN, [strong_coinduct_thm], coinduct_attrs)]
-           else
-             [])
-          |> massage_simple_notes fp_common_name;
-
-        val notes =
-          [(coinductN, map single coinduct_thms,
-            fn T_name => coinduct_attrs @ [coinduct_type_attr T_name]),
-           (corecN, corec_thmss, K coiter_attrs),
-           (disc_corecN, disc_corec_thmss, K disc_coiter_attrs),
-           (disc_corec_iffN, disc_corec_iff_thmss, K disc_coiter_iff_attrs),
-           (disc_unfoldN, disc_unfold_thmss, K disc_coiter_attrs),
-           (disc_unfold_iffN, disc_unfold_iff_thmss, K disc_coiter_iff_attrs),
-           (sel_corecN, sel_corec_thmss, K sel_coiter_attrs),
-           (sel_unfoldN, sel_unfold_thmss, K sel_coiter_attrs),
-           (simpsN, simp_thmss, K []),
-           (strong_coinductN, map single strong_coinduct_thms, K coinduct_attrs),
-           (unfoldN, unfold_thmss, K coiter_attrs)]
-          |> massage_multi_notes;
-
-        fun is_codatatype (Type (s, _)) =
-            (case fp_sugar_of lthy s of SOME {fp = Greatest_FP, ...} => true | _ => false)
-          | is_codatatype _ = false;
-
-        val nitpick_supported = forall (is_codatatype o T_of_bnf) nested_bnfs;
-
-        fun register_nitpick fpT ({ctrs, casex, ...} : ctr_sugar) =
-          Nitpick_HOL.register_codatatype fpT (fst (dest_Const casex))
-            (map (dest_Const o mk_ctr As) ctrs)
-          |> Generic_Target.theory_declaration;
-      in
-        lthy
-        |> Local_Theory.notes (common_notes @ notes) |> snd
-        |> register_fp_sugars Greatest_FP pre_bnfs nested_bnfs nesting_bnfs fp_res ctr_defss
-          ctr_sugars coiterss mapss [coinduct_thm, strong_coinduct_thm]
-          (transpose [unfold_thmss, corec_thmss]) (transpose [disc_unfold_thmss, disc_corec_thmss])
-          (transpose [sel_unfold_thmsss, sel_corec_thmsss])
-        |> nitpick_supported ? fold2 register_nitpick fpTs ctr_sugars
-      end;
-
-    val lthy'' = lthy'
-      |> fold_map define_ctrs_dtrs_for_type (fp_bnfs ~~ fp_bs ~~ fpTs ~~ ctors ~~ dtors ~~
-        xtor_co_iterss ~~ ctor_dtors ~~ dtor_ctors ~~ ctor_injects ~~ pre_map_defs ~~
-        pre_set_defss ~~ pre_rel_defs ~~ xtor_map_thms ~~ xtor_set_thmss ~~ xtor_rel_thms ~~ ns ~~
-        kss ~~ mss ~~ ctr_bindingss ~~ ctr_mixfixess ~~ ctr_Tsss ~~ disc_bindingss ~~
-        sel_bindingsss ~~ raw_sel_defaultsss)
-      |> wrap_types_etc
-      |> fp_case fp derive_note_induct_iters_thms_for_types
-           derive_note_coinduct_coiters_thms_for_types;
-
-    val timer = time (timer ("Constructors, discriminators, selectors, etc., for the new " ^
-      co_prefix fp ^ "datatype"));
-  in
-    timer; lthy''
-  end;
-
-fun co_datatypes x = define_co_datatypes (K I) (K I) (K I) x;
-
-fun co_datatype_cmd x =
-  define_co_datatypes Typedecl.read_constraint Syntax.parse_typ Syntax.parse_term x;
-
-val parse_ctr_arg =
-  @{keyword "("} |-- parse_binding_colon -- Parse.typ --| @{keyword ")"} ||
-  (Parse.typ >> pair Binding.empty);
-
-val parse_defaults =
-  @{keyword "("} |-- Parse.reserved "defaults" |-- Scan.repeat parse_bound_term --| @{keyword ")"};
-
-val parse_type_arg_constrained =
-  Parse.type_ident -- Scan.option (@{keyword "::"} |-- Parse.!!! Parse.sort);
-
-val parse_type_arg_named_constrained = parse_opt_binding_colon -- parse_type_arg_constrained;
-
-(*FIXME: use parse_type_args_named_constrained from BNF_Util and thus 
-  allow users to kill certain arguments of a (co)datatype*)
-val parse_type_args_named_constrained =
-  parse_type_arg_constrained >> (single o pair Binding.empty) ||
-  @{keyword "("} |-- Parse.!!! (Parse.list1 parse_type_arg_named_constrained --| @{keyword ")"}) ||
-  Scan.succeed [];
-
-val parse_ctr_spec =
-  parse_opt_binding_colon -- parse_binding -- Scan.repeat parse_ctr_arg --
-  Scan.optional parse_defaults [] -- Parse.opt_mixfix;
-
-val parse_spec =
-  parse_type_args_named_constrained -- parse_binding -- parse_map_rel_bindings --
-  Parse.opt_mixfix -- (@{keyword "="} |-- Parse.enum1 "|" parse_ctr_spec);
-
-val parse_co_datatype = parse_wrap_free_constructors_options -- Parse.and_list1 parse_spec;
-
-fun parse_co_datatype_cmd fp construct_fp = parse_co_datatype >> co_datatype_cmd fp construct_fp;
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_fp_def_sugar_tactics.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,181 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_fp_def_sugar_tactics.ML
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2012
-
-Tactics for datatype and codatatype sugar.
-*)
-
-signature BNF_FP_DEF_SUGAR_TACTICS =
-sig
-  val sum_prod_thms_map: thm list
-  val sum_prod_thms_set: thm list
-  val sum_prod_thms_rel: thm list
-
-  val mk_coinduct_tac: Proof.context -> thm list -> int -> int list -> thm -> thm list ->
-    thm list -> thm list -> thm list list -> thm list list list -> thm list list list -> tactic
-  val mk_coiter_tac: thm list -> thm list -> thm -> thm -> thm -> Proof.context -> tactic
-  val mk_ctor_iff_dtor_tac: Proof.context -> ctyp option list -> cterm -> cterm -> thm -> thm ->
-    tactic
-  val mk_disc_coiter_iff_tac: thm list -> thm list -> thm list -> Proof.context -> tactic
-  val mk_exhaust_tac: Proof.context -> int -> thm list -> thm -> thm -> tactic
-  val mk_half_distinct_tac: Proof.context -> thm -> thm list -> tactic
-  val mk_induct_tac: Proof.context -> int -> int list -> int list list -> int list list list ->
-    thm list -> thm -> thm list -> thm list list -> tactic
-  val mk_inject_tac: Proof.context -> thm -> thm -> tactic
-  val mk_iter_tac: thm list -> thm list -> thm list -> thm -> thm -> Proof.context -> tactic
-end;
-
-structure BNF_FP_Def_Sugar_Tactics : BNF_FP_DEF_SUGAR_TACTICS =
-struct
-
-open BNF_Tactics
-open BNF_Util
-open BNF_FP_Util
-
-val basic_simp_thms = @{thms simp_thms(7,8,12,14,22,24)};
-val more_simp_thms = basic_simp_thms @ @{thms simp_thms(11,15,16,21)};
-
-val sum_prod_thms_map = @{thms id_apply map_pair_simp prod.cases sum.cases sum_map.simps};
-val sum_prod_thms_set0 =
-  @{thms SUP_empty Sup_empty Sup_insert UN_insert Un_empty_left Un_empty_right Un_iff
-      Union_Un_distrib collect_def[abs_def] image_def o_apply map_pair_simp
-      mem_Collect_eq mem_UN_compreh_eq prod_set_simps sum_map.simps sum_set_simps};
-val sum_prod_thms_set = @{thms UN_compreh_eq_eq} @ sum_prod_thms_set0;
-val sum_prod_thms_rel = @{thms prod_rel_simp sum_rel_simps id_apply};
-
-fun hhf_concl_conv cv ctxt ct =
-  (case Thm.term_of ct of
-    Const (@{const_name all}, _) $ Abs _ =>
-    Conv.arg_conv (Conv.abs_conv (hhf_concl_conv cv o snd) ctxt) ct
-  | _ => Conv.concl_conv ~1 cv ct);
-
-fun co_induct_inst_as_projs ctxt k thm =
-  let
-    val fs = Term.add_vars (prop_of thm) []
-      |> filter (fn (_, Type (@{type_name fun}, [_, T'])) => T' <> HOLogic.boolT | _ => false);
-    fun mk_cfp (f as (_, T)) =
-      (certify ctxt (Var f), certify ctxt (mk_proj T (num_binder_types T) k));
-    val cfps = map mk_cfp fs;
-  in
-    Drule.cterm_instantiate cfps thm
-  end;
-
-val co_induct_inst_as_projs_tac = PRIMITIVE oo co_induct_inst_as_projs;
-
-fun mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor sumEN' =
-  unfold_thms_tac ctxt (ctor_iff_dtor :: ctr_defs) THEN HEADGOAL (rtac sumEN') THEN
-  unfold_thms_tac ctxt @{thms split_paired_all} THEN
-  HEADGOAL (EVERY' (maps (fn k => [select_prem_tac n (rotate_tac 1) k,
-    REPEAT_DETERM o dtac meta_spec, etac meta_mp, atac]) (1 upto n)));
-
-fun mk_ctor_iff_dtor_tac ctxt cTs cctor cdtor ctor_dtor dtor_ctor =
-  HEADGOAL (rtac iffI THEN'
-    EVERY' (map3 (fn cTs => fn cx => fn th =>
-      dtac (Drule.instantiate' cTs [NONE, NONE, SOME cx] arg_cong) THEN'
-      SELECT_GOAL (unfold_thms_tac ctxt [th]) THEN'
-      atac) [rev cTs, cTs] [cdtor, cctor] [dtor_ctor, ctor_dtor]));
-
-fun mk_half_distinct_tac ctxt ctor_inject ctr_defs =
-  unfold_thms_tac ctxt (ctor_inject :: @{thms sum.inject} @ ctr_defs) THEN
-  HEADGOAL (rtac @{thm sum.distinct(1)});
-
-fun mk_inject_tac ctxt ctr_def ctor_inject =
-  unfold_thms_tac ctxt [ctr_def] THEN HEADGOAL (rtac (ctor_inject RS ssubst)) THEN
-  unfold_thms_tac ctxt @{thms sum.inject Pair_eq conj_assoc} THEN HEADGOAL (rtac refl);
-
-val iter_unfold_thms =
-  @{thms comp_def convol_def fst_conv id_def prod_case_Pair_iden snd_conv
-      split_conv unit_case_Unity} @ sum_prod_thms_map;
-
-fun mk_iter_tac pre_map_defs map_idents iter_defs ctor_iter ctr_def ctxt =
-  unfold_thms_tac ctxt (ctr_def :: ctor_iter :: iter_defs @ pre_map_defs @ map_idents @
-    iter_unfold_thms) THEN HEADGOAL (rtac refl);
-
-val coiter_unfold_thms = @{thms id_def} @ sum_prod_thms_map;
-val ss_if_True_False = simpset_of (ss_only @{thms if_True if_False} @{context});
-
-fun mk_coiter_tac coiter_defs map_idents ctor_dtor_coiter pre_map_def ctr_def ctxt =
-  unfold_thms_tac ctxt (ctr_def :: coiter_defs) THEN
-  HEADGOAL (rtac (ctor_dtor_coiter RS trans) THEN'
-    asm_simp_tac (put_simpset ss_if_True_False ctxt)) THEN_MAYBE
-  (unfold_thms_tac ctxt (pre_map_def :: map_idents @ coiter_unfold_thms) THEN
-   HEADGOAL (rtac refl ORELSE' rtac (@{thm unit_eq} RS arg_cong)));
-
-fun mk_disc_coiter_iff_tac case_splits' coiters discs ctxt =
-  EVERY (map3 (fn case_split_tac => fn coiter_thm => fn disc =>
-      HEADGOAL case_split_tac THEN unfold_thms_tac ctxt [coiter_thm] THEN
-      HEADGOAL (asm_simp_tac (ss_only basic_simp_thms ctxt)) THEN
-      (if is_refl disc then all_tac else HEADGOAL (rtac disc)))
-    (map rtac case_splits' @ [K all_tac]) coiters discs);
-
-fun solve_prem_prem_tac ctxt =
-  REPEAT o (eresolve_tac @{thms bexE rev_bexI} ORELSE' rtac @{thm rev_bexI[OF UNIV_I]} ORELSE'
-    hyp_subst_tac ctxt ORELSE' resolve_tac @{thms disjI1 disjI2}) THEN'
-  (rtac refl ORELSE' atac ORELSE' rtac @{thm singletonI});
-
-fun mk_induct_leverage_prem_prems_tac ctxt nn kks set_maps pre_set_defs =
-  HEADGOAL (EVERY' (maps (fn kk => [select_prem_tac nn (dtac meta_spec) kk, etac meta_mp,
-    SELECT_GOAL (unfold_thms_tac ctxt (pre_set_defs @ set_maps @ sum_prod_thms_set0)),
-    solve_prem_prem_tac ctxt]) (rev kks)));
-
-fun mk_induct_discharge_prem_tac ctxt nn n set_maps pre_set_defs m k kks =
-  let val r = length kks in
-    HEADGOAL (EVERY' [select_prem_tac n (rotate_tac 1) k, rotate_tac ~1, hyp_subst_tac ctxt,
-      REPEAT_DETERM_N m o (dtac meta_spec THEN' rotate_tac ~1)]) THEN
-    EVERY [REPEAT_DETERM_N r
-        (HEADGOAL (rotate_tac ~1 THEN' dtac meta_mp THEN' rotate_tac 1) THEN prefer_tac 2),
-      if r > 0 then ALLGOALS (Goal.norm_hhf_tac ctxt) else all_tac, HEADGOAL atac,
-      mk_induct_leverage_prem_prems_tac ctxt nn kks set_maps pre_set_defs]
-  end;
-
-fun mk_induct_tac ctxt nn ns mss kkss ctr_defs ctor_induct' set_maps pre_set_defss =
-  let val n = Integer.sum ns in
-    unfold_thms_tac ctxt ctr_defs THEN HEADGOAL (rtac ctor_induct') THEN
-    co_induct_inst_as_projs_tac ctxt 0 THEN
-    EVERY (map4 (EVERY oooo map3 o mk_induct_discharge_prem_tac ctxt nn n set_maps) pre_set_defss
-      mss (unflat mss (1 upto n)) kkss)
-  end;
-
-fun mk_coinduct_same_ctr_tac ctxt rel_eqs pre_rel_def dtor_ctor ctr_def discs sels =
-  hyp_subst_tac ctxt THEN'
-  CONVERSION (hhf_concl_conv
-    (Conv.top_conv (K (Conv.try_conv (Conv.rewr_conv ctr_def))) ctxt) ctxt) THEN'
-  SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: dtor_ctor :: sels)) THEN'
-  SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: dtor_ctor :: sels @ sum_prod_thms_rel)) THEN'
-  (atac ORELSE' REPEAT o etac conjE THEN'
-     full_simp_tac
-       (ss_only (@{thm prod.inject} :: no_refl discs @ rel_eqs @ more_simp_thms) ctxt) THEN'
-     REPEAT o etac conjE THEN_MAYBE' REPEAT o hyp_subst_tac ctxt THEN'
-     REPEAT o (resolve_tac [refl, conjI] ORELSE' atac));
-
-fun mk_coinduct_distinct_ctrs_tac ctxt discs discs' =
-  let
-    val discs'' = map (perhaps (try (fn th => th RS @{thm notnotD}))) (discs @ discs')
-      |> distinct Thm.eq_thm_prop;
-  in
-    hyp_subst_tac ctxt THEN' REPEAT o etac conjE THEN'
-    full_simp_tac (ss_only (refl :: no_refl discs'' @ basic_simp_thms) ctxt)
-  end;
-
-fun mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn kk n pre_rel_def dtor_ctor exhaust ctr_defs
-    discss selss =
-  let val ks = 1 upto n in
-    EVERY' ([rtac allI, rtac allI, rtac impI, select_prem_tac nn (dtac meta_spec) kk,
-        dtac meta_spec, dtac meta_mp, atac, rtac exhaust, K (co_induct_inst_as_projs_tac ctxt 0),
-        hyp_subst_tac ctxt] @
-      map4 (fn k => fn ctr_def => fn discs => fn sels =>
-        EVERY' ([rtac exhaust, K (co_induct_inst_as_projs_tac ctxt 1)] @
-          map2 (fn k' => fn discs' =>
-            if k' = k then
-              mk_coinduct_same_ctr_tac ctxt rel_eqs' pre_rel_def dtor_ctor ctr_def discs sels
-            else
-              mk_coinduct_distinct_ctrs_tac ctxt discs discs') ks discss)) ks ctr_defs discss selss)
-  end;
-
-fun mk_coinduct_tac ctxt rel_eqs' nn ns dtor_coinduct' pre_rel_defs dtor_ctors exhausts ctr_defss
-    discsss selsss =
-  HEADGOAL (rtac dtor_coinduct' THEN'
-    EVERY' (map8 (mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn)
-      (1 upto nn) ns pre_rel_defs dtor_ctors exhausts ctr_defss discsss selsss));
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_fp_n2m.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,378 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_fp_n2m.ML
-    Author:     Dmitriy Traytel, TU Muenchen
-    Copyright   2013
-
-Flattening of nested to mutual (co)recursion.
-*)
-
-signature BNF_FP_N2M =
-sig
-  val construct_mutualized_fp: BNF_FP_Util.fp_kind  -> typ list -> BNF_FP_Def_Sugar.fp_sugar list ->
-    binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list ->
-    local_theory -> BNF_FP_Util.fp_result * local_theory
-end;
-
-structure BNF_FP_N2M : BNF_FP_N2M =
-struct
-
-open BNF_Def
-open BNF_Util
-open BNF_FP_Util
-open BNF_FP_Def_Sugar
-open BNF_Tactics
-open BNF_FP_N2M_Tactics
-
-fun force_typ ctxt T =
-  map_types Type_Infer.paramify_vars
-  #> Type.constraint T
-  #> Syntax.check_term ctxt
-  #> singleton (Variable.polymorphic ctxt);
-
-fun mk_prod_map f g =
-  let
-    val ((fAT, fBT), fT) = `dest_funT (fastype_of f);
-    val ((gAT, gBT), gT) = `dest_funT (fastype_of g);
-  in
-    Const (@{const_name map_pair},
-      fT --> gT --> HOLogic.mk_prodT (fAT, gAT) --> HOLogic.mk_prodT (fBT, gBT)) $ f $ g
-  end;
-
-fun mk_sum_map f g =
-  let
-    val ((fAT, fBT), fT) = `dest_funT (fastype_of f);
-    val ((gAT, gBT), gT) = `dest_funT (fastype_of g);
-  in
-    Const (@{const_name sum_map}, fT --> gT --> mk_sumT (fAT, gAT) --> mk_sumT (fBT, gBT)) $ f $ g
-  end;
-
-fun construct_mutualized_fp fp fpTs fp_sugars bs resBs (resDs, Dss) bnfs lthy =
-  let
-    fun steal get = map (of_fp_sugar (get o #fp_res)) fp_sugars;
-
-    val n = length bnfs;
-    val deads = fold (union (op =)) Dss resDs;
-    val As = subtract (op =) deads (map TFree resBs);
-    val names_lthy = fold Variable.declare_typ (As @ deads) lthy;
-    val m = length As;
-    val live = m + n;
-    val ((Xs, Bs), names_lthy) = names_lthy
-      |> mk_TFrees n
-      ||>> mk_TFrees m;
-    val allAs = As @ Xs;
-    val phiTs = map2 mk_pred2T As Bs;
-    val theta = As ~~ Bs;
-    val fpTs' = map (Term.typ_subst_atomic theta) fpTs;
-    val pre_phiTs = map2 mk_pred2T fpTs fpTs';
-
-    fun mk_co_algT T U = fp_case fp (T --> U) (U --> T);
-    fun co_swap pair = fp_case fp I swap pair;
-    val dest_co_algT = co_swap o dest_funT;
-    val co_alg_argT = fp_case fp range_type domain_type;
-    val co_alg_funT = fp_case fp domain_type range_type;
-    val mk_co_product = curry (fp_case fp mk_convol mk_sum_case);
-    val mk_map_co_product = fp_case fp mk_prod_map mk_sum_map;
-    val co_proj1_const = fp_case fp (fst_const o fst) (uncurry Inl_const o dest_sumT o snd);
-    val mk_co_productT = curry (fp_case fp HOLogic.mk_prodT mk_sumT);
-    val dest_co_productT = fp_case fp HOLogic.dest_prodT dest_sumT;
-
-    val ((ctors, dtors), (xtor's, xtors)) =
-      let
-        val ctors = map2 (force_typ names_lthy o (fn T => dummyT --> T)) fpTs (steal #ctors);
-        val dtors = map2 (force_typ names_lthy o (fn T => T --> dummyT)) fpTs (steal #dtors);
-      in
-        ((ctors, dtors), `(map (Term.subst_atomic_types theta)) (fp_case fp ctors dtors))
-      end;
-
-    val xTs = map (domain_type o fastype_of) xtors;
-    val yTs = map (domain_type o fastype_of) xtor's;
-
-    val (((((phis, phis'), pre_phis), xs), ys), names_lthy) = names_lthy
-      |> mk_Frees' "R" phiTs
-      ||>> mk_Frees "S" pre_phiTs
-      ||>> mk_Frees "x" xTs
-      ||>> mk_Frees "y" yTs;
-
-    val fp_bnfs = steal #bnfs;
-    val pre_bnfs = map (of_fp_sugar #pre_bnfs) fp_sugars;
-    val pre_bnfss = map #pre_bnfs fp_sugars;
-    val nesty_bnfss = map (fn sugar => #nested_bnfs sugar @ #nesting_bnfs sugar) fp_sugars;
-    val fp_nesty_bnfss = fp_bnfs :: nesty_bnfss;
-    val fp_nesty_bnfs = distinct eq_bnf (flat fp_nesty_bnfss);
-
-    val rels =
-      let
-        fun find_rel T As Bs = fp_nesty_bnfss
-          |> map (filter_out (curry eq_bnf BNF_Comp.DEADID_bnf))
-          |> get_first (find_first (fn bnf => Type.could_unify (T_of_bnf bnf, T)))
-          |> Option.map (fn bnf =>
-            let val live = live_of_bnf bnf;
-            in (mk_rel live As Bs (rel_of_bnf bnf), live) end)
-          |> the_default (HOLogic.eq_const T, 0);
-
-        fun mk_rel (T as Type (_, Ts)) (Type (_, Us)) =
-              let
-                val (rel, live) = find_rel T Ts Us;
-                val (Ts', Us') = fastype_of rel |> strip_typeN live |> fst |> map_split dest_pred2T;
-                val rels = map2 mk_rel Ts' Us';
-              in
-                Term.list_comb (rel, rels)
-              end
-          | mk_rel (T as TFree _) _ = (nth phis (find_index (curry op = T) As)
-              handle General.Subscript => HOLogic.eq_const T)
-          | mk_rel _ _ = raise Fail "fpTs contains schematic type variables";
-      in
-        map2 (fold_rev Term.absfree phis' oo mk_rel) fpTs fpTs'
-      end;
-
-    val pre_rels = map2 (fn Ds => mk_rel_of_bnf Ds (As @ fpTs) (Bs @ fpTs')) Dss bnfs;
-
-    val rel_unfoldss = map (maps (fn bnf => no_refl [rel_def_of_bnf bnf])) pre_bnfss;
-    val rel_xtor_co_inducts = steal (split_conj_thm o #rel_xtor_co_induct_thm)
-      |> map2 (fn unfs => unfold_thms lthy (id_apply :: unfs)) rel_unfoldss;
-
-    val rel_defs = map rel_def_of_bnf bnfs;
-    val rel_monos = map rel_mono_of_bnf bnfs;
-
-    val rel_xtor_co_induct_thm =
-      mk_rel_xtor_co_induct_thm fp pre_rels pre_phis rels phis xs ys xtors xtor's
-        (mk_rel_xtor_co_induct_tactic fp rel_xtor_co_inducts rel_defs rel_monos) lthy;
-
-    val rel_eqs = no_refl (map rel_eq_of_bnf fp_nesty_bnfs);
-    val map_id0s = no_refl (map map_id0_of_bnf bnfs);
-
-    val xtor_co_induct_thm =
-      (case fp of
-        Least_FP =>
-          let
-            val (Ps, names_lthy) = names_lthy
-              |> mk_Frees "P" (map (fn T => T --> HOLogic.boolT) fpTs);
-            fun mk_Grp_id P =
-              let val T = domain_type (fastype_of P);
-              in mk_Grp (HOLogic.Collect_const T $ P) (HOLogic.id_const T) end;
-            val cts = map (SOME o certify lthy) (map HOLogic.eq_const As @ map mk_Grp_id Ps);
-          in
-            cterm_instantiate_pos cts rel_xtor_co_induct_thm
-            |> singleton (Proof_Context.export names_lthy lthy)
-            |> unfold_thms lthy (@{thms eq_le_Grp_id_iff all_simps(1,2)[symmetric]} @ rel_eqs)
-            |> funpow n (fn thm => thm RS spec)
-            |> unfold_thms lthy (@{thm eq_alt} :: map rel_Grp_of_bnf bnfs @ map_id0s)
-            |> unfold_thms lthy @{thms Grp_id_mono_subst eqTrueI[OF subset_UNIV] simp_thms(22)}
-            |> unfold_thms lthy @{thms subset_iff mem_Collect_eq
-               atomize_conjL[symmetric] atomize_all[symmetric] atomize_imp[symmetric]}
-            |> unfold_thms lthy (maps set_defs_of_bnf bnfs)
-          end
-      | Greatest_FP =>
-          let
-            val cts = NONE :: map (SOME o certify lthy) (map HOLogic.eq_const As);
-          in
-            cterm_instantiate_pos cts rel_xtor_co_induct_thm
-            |> unfold_thms lthy (@{thms le_fun_def le_bool_def all_simps(1,2)[symmetric]} @ rel_eqs)
-            |> funpow (2 * n) (fn thm => thm RS spec)
-            |> Conv.fconv_rule (Object_Logic.atomize lthy)
-            |> funpow n (fn thm => thm RS mp)
-          end);
-
-    val fold_preTs = map2 (fn Ds => mk_T_of_bnf Ds allAs) Dss bnfs;
-    val fold_pre_deads_only_Ts = map2 (fn Ds => mk_T_of_bnf Ds (replicate live dummyT)) Dss bnfs;
-    val rec_theta = Xs ~~ map2 mk_co_productT fpTs Xs;
-    val rec_preTs = map (Term.typ_subst_atomic rec_theta) fold_preTs;
-
-    val fold_strTs = map2 mk_co_algT fold_preTs Xs;
-    val rec_strTs = map2 mk_co_algT rec_preTs Xs;
-    val resTs = map2 mk_co_algT fpTs Xs;
-
-    val (((fold_strs, fold_strs'), (rec_strs, rec_strs')), names_lthy) = names_lthy
-      |> mk_Frees' "s" fold_strTs
-      ||>> mk_Frees' "s" rec_strTs;
-
-    val co_iters = steal #xtor_co_iterss;
-    val ns = map (length o #pre_bnfs) fp_sugars;
-    fun substT rho (Type (@{type_name "fun"}, [T, U])) = substT rho T --> substT rho U
-      | substT rho (Type (s, Ts)) = Type (s, map (typ_subst_nonatomic rho) Ts)
-      | substT _ T = T;
-    fun force_iter is_rec i TU TU_rec raw_iters =
-      let
-        val approx_fold = un_fold_of raw_iters
-          |> force_typ names_lthy
-            (replicate (nth ns i) dummyT ---> (if is_rec then TU_rec else TU));
-        val TUs = binder_fun_types (Term.typ_subst_atomic (Xs ~~ fpTs) (fastype_of approx_fold));
-        val js = find_indices Type.could_unify
-          TUs (map (Term.typ_subst_atomic (Xs ~~ fpTs)) fold_strTs);
-        val Tpats = map (fn j => mk_co_algT (nth fold_pre_deads_only_Ts j) (nth Xs j)) js;
-        val iter = raw_iters |> (if is_rec then co_rec_of else un_fold_of);
-      in
-        force_typ names_lthy (Tpats ---> TU) iter
-      end;
-
-    fun mk_iter b_opt is_rec iters lthy TU =
-      let
-        val x = co_alg_argT TU;
-        val i = find_index (fn T => x = T) Xs;
-        val TUiter =
-          (case find_first (fn f => body_fun_type (fastype_of f) = TU) iters of
-            NONE => nth co_iters i
-              |> force_iter is_rec i
-                (TU |> (is_none b_opt andalso not is_rec) ? substT (fpTs ~~ Xs))
-                (TU |> (is_none b_opt) ? substT (map2 mk_co_productT fpTs Xs ~~ Xs))
-          | SOME f => f);
-        val TUs = binder_fun_types (fastype_of TUiter);
-        val iter_preTs = if is_rec then rec_preTs else fold_preTs;
-        val iter_strs = if is_rec then rec_strs else fold_strs;
-        fun mk_s TU' =
-          let
-            val i = find_index (fn T => co_alg_argT TU' = T) Xs;
-            val sF = co_alg_funT TU';
-            val F = nth iter_preTs i;
-            val s = nth iter_strs i;
-          in
-            (if sF = F then s
-            else
-              let
-                val smapT = replicate live dummyT ---> mk_co_algT sF F;
-                fun hidden_to_unit t =
-                  Term.subst_TVars (map (rpair HOLogic.unitT) (Term.add_tvar_names t [])) t;
-                val smap = map_of_bnf (nth bnfs i)
-                  |> force_typ names_lthy smapT
-                  |> hidden_to_unit;
-                val smap_argTs = strip_typeN live (fastype_of smap) |> fst;
-                fun mk_smap_arg TU =
-                  (if domain_type TU = range_type TU then
-                    HOLogic.id_const (domain_type TU)
-                  else if is_rec then
-                    let
-                      val (TY, (U, X)) = TU |> dest_co_algT ||> dest_co_productT;
-                      val T = mk_co_algT TY U;
-                    in
-                      (case try (force_typ lthy T o build_map lthy co_proj1_const o dest_funT) T of
-                        SOME f => mk_co_product f
-                          (fst (fst (mk_iter NONE is_rec iters lthy (mk_co_algT TY X))))
-                      | NONE => mk_map_co_product
-                          (build_map lthy co_proj1_const
-                            (dest_funT (mk_co_algT (dest_co_productT TY |> fst) U)))
-                          (HOLogic.id_const X))
-                    end
-                  else
-                    fst (fst (mk_iter NONE is_rec iters lthy TU)))
-                val smap_args = map mk_smap_arg smap_argTs;
-              in
-                HOLogic.mk_comp (co_swap (s, Term.list_comb (smap, smap_args)))
-              end)
-          end;
-        val t = Term.list_comb (TUiter, map mk_s TUs);
-      in
-        (case b_opt of
-          NONE => ((t, Drule.dummy_thm), lthy)
-        | SOME b => Local_Theory.define ((b, NoSyn), ((Binding.conceal (Thm.def_binding b), []),
-            fold_rev Term.absfree (if is_rec then rec_strs' else fold_strs') t)) lthy |>> apsnd snd)
-      end;
-
-    fun mk_iters is_rec name lthy =
-      fold2 (fn TU => fn b => fn ((iters, defs), lthy) =>
-        mk_iter (SOME b) is_rec iters lthy TU |>> (fn (f, d) => (f :: iters, d :: defs)))
-      resTs (map (Binding.suffix_name ("_" ^ name)) bs) (([], []), lthy)
-      |>> apfst rev o apsnd rev;
-    val foldN = fp_case fp ctor_foldN dtor_unfoldN;
-    val recN = fp_case fp ctor_recN dtor_corecN;
-    val (((raw_un_folds, raw_un_fold_defs), (raw_co_recs, raw_co_rec_defs)), (lthy, raw_lthy)) =
-      lthy
-      |> mk_iters false foldN
-      ||>> mk_iters true recN
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism raw_lthy lthy;
-
-    val un_folds = map (Morphism.term phi) raw_un_folds;
-    val co_recs = map (Morphism.term phi) raw_co_recs;
-
-    val (xtor_un_fold_thms, xtor_co_rec_thms) =
-      let
-        val folds = map (fn f => Term.list_comb (f, fold_strs)) raw_un_folds;
-        val recs = map (fn r => Term.list_comb (r, rec_strs)) raw_co_recs;
-        val fold_mapTs = co_swap (As @ fpTs, As @ Xs);
-        val rec_mapTs = co_swap (As @ fpTs, As @ map2 mk_co_productT fpTs Xs);
-        val pre_fold_maps =
-          map2 (fn Ds => fn bnf =>
-            Term.list_comb (uncurry (mk_map_of_bnf Ds) fold_mapTs bnf,
-              map HOLogic.id_const As @ folds))
-          Dss bnfs;
-        val pre_rec_maps =
-          map2 (fn Ds => fn bnf =>
-            Term.list_comb (uncurry (mk_map_of_bnf Ds) rec_mapTs bnf,
-              map HOLogic.id_const As @ map2 (mk_co_product o HOLogic.id_const) fpTs recs))
-          Dss bnfs;
-
-        fun mk_goals f xtor s smap =
-          ((f, xtor), (s, smap))
-          |> pairself (HOLogic.mk_comp o co_swap)
-          |> HOLogic.mk_eq;
-
-        val fold_goals = map4 mk_goals folds xtors fold_strs pre_fold_maps
-        val rec_goals = map4 mk_goals recs xtors rec_strs pre_rec_maps;
-
-        fun mk_thms ss goals tac =
-          Library.foldr1 HOLogic.mk_conj goals
-          |> HOLogic.mk_Trueprop
-          |> fold_rev Logic.all ss
-          |> (fn goal => Goal.prove_sorry raw_lthy [] [] goal tac)
-          |> Thm.close_derivation
-          |> Morphism.thm phi
-          |> split_conj_thm
-          |> map (fn thm => thm RS @{thm comp_eq_dest});
-
-        val pre_map_defs = no_refl (map map_def_of_bnf bnfs);
-        val fp_pre_map_defs = no_refl (map map_def_of_bnf pre_bnfs);
-
-        val map_unfoldss = map (maps (fn bnf => no_refl [map_def_of_bnf bnf])) pre_bnfss;
-        val unfold_map = map2 (fn unfs => unfold_thms lthy (id_apply :: unfs)) map_unfoldss;
-
-        val fp_xtor_co_iterss = steal #xtor_co_iter_thmss;
-        val fp_xtor_un_folds = map (mk_pointfree lthy o un_fold_of) fp_xtor_co_iterss |> unfold_map;
-        val fp_xtor_co_recs = map (mk_pointfree lthy o co_rec_of) fp_xtor_co_iterss |> unfold_map;
-
-        val fp_co_iter_o_mapss = steal #xtor_co_iter_o_map_thmss;
-        val fp_fold_o_maps = map un_fold_of fp_co_iter_o_mapss |> unfold_map;
-        val fp_rec_o_maps = map co_rec_of fp_co_iter_o_mapss |> unfold_map;
-        val fold_thms = fp_case fp @{thm o_assoc[symmetric]} @{thm o_assoc} ::
-          @{thms id_apply o_apply o_id id_o map_pair.comp map_pair.id sum_map.comp sum_map.id};
-        val rec_thms = fold_thms @ fp_case fp
-          @{thms fst_convol map_pair_o_convol convol_o}
-          @{thms sum_case_o_inj(1) sum_case_o_sum_map o_sum_case};
-        val map_thms = no_refl (maps (fn bnf =>
-          [map_comp0_of_bnf bnf RS sym, map_id0_of_bnf bnf]) fp_nesty_bnfs);
-
-        fun mk_tac defs o_map_thms xtor_thms thms {context = ctxt, prems = _} =
-          unfold_thms_tac ctxt
-            (flat [thms, defs, pre_map_defs, fp_pre_map_defs, xtor_thms, o_map_thms, map_thms]) THEN
-          CONJ_WRAP (K (HEADGOAL (rtac refl))) bnfs;
-
-        val fold_tac = mk_tac raw_un_fold_defs fp_fold_o_maps fp_xtor_un_folds fold_thms;
-        val rec_tac = mk_tac raw_co_rec_defs fp_rec_o_maps fp_xtor_co_recs rec_thms;
-      in
-        (mk_thms fold_strs fold_goals fold_tac, mk_thms rec_strs rec_goals rec_tac)
-      end;
-
-    (* These results are half broken. This is deliberate. We care only about those fields that are
-       used by "primrec_new", "primcorecursive", and "datatype_new_compat". *)
-    val fp_res =
-      ({Ts = fpTs,
-        bnfs = steal #bnfs,
-        dtors = dtors,
-        ctors = ctors,
-        xtor_co_iterss = transpose [un_folds, co_recs],
-        xtor_co_induct = xtor_co_induct_thm,
-        dtor_ctors = steal #dtor_ctors (*too general types*),
-        ctor_dtors = steal #ctor_dtors (*too general types*),
-        ctor_injects = steal #ctor_injects (*too general types*),
-        dtor_injects = steal #dtor_injects (*too general types*),
-        xtor_map_thms = steal #xtor_map_thms (*too general types and terms*),
-        xtor_set_thmss = steal #xtor_set_thmss (*too general types and terms*),
-        xtor_rel_thms = steal #xtor_rel_thms (*too general types and terms*),
-        xtor_co_iter_thmss = transpose [xtor_un_fold_thms, xtor_co_rec_thms],
-        xtor_co_iter_o_map_thmss = steal #xtor_co_iter_o_map_thmss (*theorem about old constant*),
-        rel_xtor_co_induct_thm = rel_xtor_co_induct_thm}
-       |> morph_fp_result (Morphism.term_morphism "BNF" (singleton (Variable.polymorphic lthy))));
-  in
-    (fp_res, lthy)
-  end;
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_fp_n2m_sugar.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,394 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_fp_n2m_sugar.ML
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2013
-
-Suggared flattening of nested to mutual (co)recursion.
-*)
-
-signature BNF_FP_N2M_SUGAR =
-sig
-  val unfold_let: term -> term
-  val dest_map: Proof.context -> string -> term -> term * term list
-
-  val mutualize_fp_sugars: BNF_FP_Util.fp_kind -> binding list -> typ list -> (term -> int list) ->
-    term list list list list -> BNF_FP_Def_Sugar.fp_sugar list -> local_theory ->
-    (BNF_FP_Def_Sugar.fp_sugar list
-     * (BNF_FP_Def_Sugar.lfp_sugar_thms option * BNF_FP_Def_Sugar.gfp_sugar_thms option))
-    * local_theory
-  val indexify_callsss: BNF_FP_Def_Sugar.fp_sugar -> (term * term list list) list ->
-    term list list list
-  val nested_to_mutual_fps: BNF_FP_Util.fp_kind -> binding list -> typ list -> (term -> int list) ->
-    (term * term list list) list list -> local_theory ->
-    (typ list * int list * BNF_FP_Def_Sugar.fp_sugar list
-     * (BNF_FP_Def_Sugar.lfp_sugar_thms option * BNF_FP_Def_Sugar.gfp_sugar_thms option))
-    * local_theory
-end;
-
-structure BNF_FP_N2M_Sugar : BNF_FP_N2M_SUGAR =
-struct
-
-open Ctr_Sugar
-open BNF_Util
-open BNF_Def
-open BNF_FP_Util
-open BNF_FP_Def_Sugar
-open BNF_FP_N2M
-
-val n2mN = "n2m_"
-
-type n2m_sugar = fp_sugar list * (lfp_sugar_thms option * gfp_sugar_thms option);
-
-structure Data = Generic_Data
-(
-  type T = n2m_sugar Typtab.table;
-  val empty = Typtab.empty;
-  val extend = I;
-  val merge = Typtab.merge (eq_fst (eq_list eq_fp_sugar));
-);
-
-fun morph_n2m_sugar phi (fp_sugars, (lfp_sugar_thms_opt, gfp_sugar_thms_opt)) =
-  (map (morph_fp_sugar phi) fp_sugars,
-   (Option.map (morph_lfp_sugar_thms phi) lfp_sugar_thms_opt,
-    Option.map (morph_gfp_sugar_thms phi) gfp_sugar_thms_opt));
-
-val transfer_n2m_sugar =
-  morph_n2m_sugar o Morphism.transfer_morphism o Proof_Context.theory_of;
-
-fun n2m_sugar_of ctxt =
-  Typtab.lookup (Data.get (Context.Proof ctxt))
-  #> Option.map (transfer_n2m_sugar ctxt);
-
-fun register_n2m_sugar key n2m_sugar =
-  Local_Theory.declaration {syntax = false, pervasive = false}
-    (fn phi => Data.map (Typtab.default (key, morph_n2m_sugar phi n2m_sugar)));
-
-fun unfold_let (Const (@{const_name Let}, _) $ arg1 $ arg2) = unfold_let (betapply (arg2, arg1))
-  | unfold_let (Const (@{const_name prod_case}, _) $ t) =
-    (case unfold_let t of
-      t' as Abs (s1, T1, Abs (s2, T2, _)) =>
-      let val v = Var ((s1 ^ s2, Term.maxidx_of_term t' + 1), HOLogic.mk_prodT (T1, T2)) in
-        lambda v (incr_boundvars 1 (betapplys (t', [HOLogic.mk_fst v, HOLogic.mk_snd v])))
-      end
-    | _ => t)
-  | unfold_let (t $ u) = betapply (unfold_let t, unfold_let u)
-  | unfold_let (Abs (s, T, t)) = Abs (s, T, unfold_let t)
-  | unfold_let t = t;
-
-fun mk_map_pattern ctxt s =
-  let
-    val bnf = the (bnf_of ctxt s);
-    val mapx = map_of_bnf bnf;
-    val live = live_of_bnf bnf;
-    val (f_Ts, _) = strip_typeN live (fastype_of mapx);
-    val fs = map_index (fn (i, T) => Var (("?f", i), T)) f_Ts;
-  in
-    (mapx, betapplys (mapx, fs))
-  end;
-
-fun dest_map ctxt s call =
-  let
-    val (map0, pat) = mk_map_pattern ctxt s;
-    val (_, tenv) = fo_match ctxt call pat;
-  in
-    (map0, Vartab.fold_rev (fn (_, (_, f)) => cons f) tenv [])
-  end;
-
-fun dest_abs_or_applied_map _ _ (Abs (_, _, t)) = (Term.dummy, [t])
-  | dest_abs_or_applied_map ctxt s (t1 $ _) = dest_map ctxt s t1;
-
-fun map_partition f xs =
-  fold_rev (fn x => fn (ys, (good, bad)) =>
-      case f x of SOME y => (y :: ys, (x :: good, bad)) | NONE => (ys, (good, x :: bad)))
-    xs ([], ([], []));
-
-fun key_of_fp_eqs fp fpTs fp_eqs =
-  Type (fp_case fp "l" "g", fpTs @ maps (fn (x, T) => [TFree x, T]) fp_eqs);
-
-(* TODO: test with sort constraints on As *)
-fun mutualize_fp_sugars fp bs fpTs get_indices callssss fp_sugars0 no_defs_lthy0 =
-  let
-    val thy = Proof_Context.theory_of no_defs_lthy0;
-
-    val qsotm = quote o Syntax.string_of_term no_defs_lthy0;
-
-    fun incompatible_calls t1 t2 =
-      error ("Incompatible " ^ co_prefix fp ^ "recursive calls: " ^ qsotm t1 ^ " vs. " ^ qsotm t2);
-    fun nested_self_call t =
-      error ("Unsupported nested self-call " ^ qsotm t);
-
-    val b_names = map Binding.name_of bs;
-    val fp_b_names = map base_name_of_typ fpTs;
-
-    val nn = length fpTs;
-
-    fun target_ctr_sugar_of_fp_sugar fpT ({T, index, ctr_sugars, ...} : fp_sugar) =
-      let
-        val rho = Vartab.fold (cons o apsnd snd) (Sign.typ_match thy (T, fpT) Vartab.empty) [];
-        val phi = Morphism.term_morphism "BNF" (Term.subst_TVars rho);
-      in
-        morph_ctr_sugar phi (nth ctr_sugars index)
-      end;
-
-    val ctr_defss = map (of_fp_sugar #ctr_defss) fp_sugars0;
-    val mapss = map (of_fp_sugar #mapss) fp_sugars0;
-    val ctr_sugars = map2 target_ctr_sugar_of_fp_sugar fpTs fp_sugars0;
-
-    val ctrss = map #ctrs ctr_sugars;
-    val ctr_Tss = map (map fastype_of) ctrss;
-
-    val As' = fold (fold Term.add_tfreesT) ctr_Tss [];
-    val As = map TFree As';
-
-    val ((Cs, Xs), no_defs_lthy) =
-      no_defs_lthy0
-      |> fold Variable.declare_typ As
-      |> mk_TFrees nn
-      ||>> variant_tfrees fp_b_names;
-
-    fun check_call_dead live_call call =
-      if null (get_indices call) then () else incompatible_calls live_call call;
-
-    fun freeze_fpTs_simple (T as Type (s, Ts)) =
-        (case find_index (curry (op =) T) fpTs of
-          ~1 => Type (s, map freeze_fpTs_simple Ts)
-        | kk => nth Xs kk)
-      | freeze_fpTs_simple T = T;
-
-    fun freeze_fpTs_map (fpT as Type (_, Ts')) (callss, (live_call :: _, dead_calls))
-        (T as Type (s, Ts)) =
-      if Ts' = Ts then
-        nested_self_call live_call
-      else
-        (List.app (check_call_dead live_call) dead_calls;
-         Type (s, map2 (freeze_fpTs fpT) (flatten_type_args_of_bnf (the (bnf_of no_defs_lthy s)) []
-           (transpose callss)) Ts))
-    and freeze_fpTs fpT calls (T as Type (s, _)) =
-        (case map_partition (try (snd o dest_map no_defs_lthy s)) calls of
-          ([], _) =>
-          (case map_partition (try (snd o dest_abs_or_applied_map no_defs_lthy s)) calls of
-            ([], _) => freeze_fpTs_simple T
-          | callsp => freeze_fpTs_map fpT callsp T)
-        | callsp => freeze_fpTs_map fpT callsp T)
-      | freeze_fpTs _ _ T = T;
-
-    val ctr_Tsss = map (map binder_types) ctr_Tss;
-    val ctrXs_Tsss = map3 (map2 o map2 o freeze_fpTs) fpTs callssss ctr_Tsss;
-    val ctrXs_sum_prod_Ts = map (mk_sumTN_balanced o map HOLogic.mk_tupleT) ctrXs_Tsss;
-    val ctr_Ts = map (body_type o hd) ctr_Tss;
-
-    val ns = map length ctr_Tsss;
-    val kss = map (fn n => 1 upto n) ns;
-    val mss = map (map length) ctr_Tsss;
-
-    val fp_eqs = map dest_TFree Xs ~~ ctrXs_sum_prod_Ts;
-    val key = key_of_fp_eqs fp fpTs fp_eqs;
-  in
-    (case n2m_sugar_of no_defs_lthy key of
-      SOME n2m_sugar => (n2m_sugar, no_defs_lthy)
-    | NONE =>
-      let
-        val base_fp_names = Name.variant_list [] fp_b_names;
-        val fp_bs = map2 (fn b_name => fn base_fp_name =>
-            Binding.qualify true b_name (Binding.name (n2mN ^ base_fp_name)))
-          b_names base_fp_names;
-
-        val (pre_bnfs, (fp_res as {xtor_co_iterss = xtor_co_iterss0, xtor_co_induct, dtor_injects,
-               dtor_ctors, xtor_co_iter_thmss, ...}, lthy)) =
-          fp_bnf (construct_mutualized_fp fp fpTs fp_sugars0) fp_bs As' fp_eqs no_defs_lthy;
-
-        val nesting_bnfs = nesty_bnfs lthy ctrXs_Tsss As;
-        val nested_bnfs = nesty_bnfs lthy ctrXs_Tsss Xs;
-
-        val ((xtor_co_iterss, iters_args_types, coiters_args_types), _) =
-          mk_co_iters_prelims fp ctr_Tsss fpTs Cs ns mss xtor_co_iterss0 lthy;
-
-        fun mk_binding b suf = Binding.suffix_name ("_" ^ suf) b;
-
-        val ((co_iterss, co_iter_defss), lthy) =
-          fold_map2 (fn b =>
-            (if fp = Least_FP then define_iters [foldN, recN] (the iters_args_types)
-             else define_coiters [unfoldN, corecN] (the coiters_args_types))
-              (mk_binding b) fpTs Cs) fp_bs xtor_co_iterss lthy
-          |>> split_list;
-
-        val ((co_inducts, un_fold_thmss, co_rec_thmss, disc_unfold_thmss, disc_corec_thmss,
-              sel_unfold_thmsss, sel_corec_thmsss), fp_sugar_thms) =
-          if fp = Least_FP then
-            derive_induct_iters_thms_for_types pre_bnfs (the iters_args_types) xtor_co_induct
-              xtor_co_iter_thmss nesting_bnfs nested_bnfs fpTs Cs Xs ctrXs_Tsss ctrss ctr_defss
-              co_iterss co_iter_defss lthy
-            |> `(fn ((_, induct, _), (fold_thmss, rec_thmss, _)) =>
-              ([induct], fold_thmss, rec_thmss, [], [], [], []))
-            ||> (fn info => (SOME info, NONE))
-          else
-            derive_coinduct_coiters_thms_for_types pre_bnfs (the coiters_args_types) xtor_co_induct
-              dtor_injects dtor_ctors xtor_co_iter_thmss nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss
-              ns ctr_defss ctr_sugars co_iterss co_iter_defss
-              (Proof_Context.export lthy no_defs_lthy) lthy
-            |> `(fn ((coinduct_thms_pairs, _), (unfold_thmss, corec_thmss, _),
-                    (disc_unfold_thmss, disc_corec_thmss, _), _,
-                    (sel_unfold_thmsss, sel_corec_thmsss, _)) =>
-              (map snd coinduct_thms_pairs, unfold_thmss, corec_thmss, disc_unfold_thmss,
-               disc_corec_thmss, sel_unfold_thmsss, sel_corec_thmsss))
-            ||> (fn info => (NONE, SOME info));
-
-        val phi = Proof_Context.export_morphism no_defs_lthy no_defs_lthy0;
-
-        fun mk_target_fp_sugar (kk, T) =
-          {T = T, fp = fp, index = kk, pre_bnfs = pre_bnfs, nested_bnfs = nested_bnfs,
-           nesting_bnfs = nesting_bnfs, fp_res = fp_res, ctr_defss = ctr_defss,
-           ctr_sugars = ctr_sugars, co_iterss = co_iterss, mapss = mapss, co_inducts = co_inducts,
-           co_iter_thmsss = transpose [un_fold_thmss, co_rec_thmss],
-           disc_co_itersss = transpose [disc_unfold_thmss, disc_corec_thmss],
-           sel_co_iterssss = transpose [sel_unfold_thmsss, sel_corec_thmsss]}
-          |> morph_fp_sugar phi;
-
-        val n2m_sugar = (map_index mk_target_fp_sugar fpTs, fp_sugar_thms);
-      in
-        (n2m_sugar, lthy |> register_n2m_sugar key n2m_sugar)
-      end)
-  end;
-
-fun indexify_callsss fp_sugar callsss =
-  let
-    val {ctrs, ...} = of_fp_sugar #ctr_sugars fp_sugar;
-    fun indexify_ctr ctr =
-      (case AList.lookup Term.aconv_untyped callsss ctr of
-        NONE => replicate (num_binder_types (fastype_of ctr)) []
-      | SOME callss => map (map (Envir.beta_eta_contract o unfold_let)) callss);
-  in
-    map indexify_ctr ctrs
-  end;
-
-fun retypargs tyargs (Type (s, _)) = Type (s, tyargs);
-
-fun fold_subtype_pairs f (T as Type (s, Ts), U as Type (s', Us)) =
-    f (T, U) #> (if s = s' then fold (fold_subtype_pairs f) (Ts ~~ Us) else I)
-  | fold_subtype_pairs f TU = f TU;
-
-fun nested_to_mutual_fps fp actual_bs actual_Ts get_indices actual_callssss0 lthy =
-  let
-    val qsoty = quote o Syntax.string_of_typ lthy;
-    val qsotys = space_implode " or " o map qsoty;
-
-    fun duplicate_datatype T = error (qsoty T ^ " is not mutually recursive with itself");
-    fun not_co_datatype0 T = error (qsoty T ^ " is not a " ^ co_prefix fp ^ "datatype");
-    fun not_co_datatype (T as Type (s, _)) =
-        if fp = Least_FP andalso
-           is_some (Datatype_Data.get_info (Proof_Context.theory_of lthy) s) then
-          error (qsoty T ^ " is not a new-style datatype (cf. \"datatype_new\")")
-        else
-          not_co_datatype0 T
-      | not_co_datatype T = not_co_datatype0 T;
-    fun not_mutually_nested_rec Ts1 Ts2 =
-      error (qsotys Ts1 ^ " is neither mutually recursive with " ^ qsotys Ts2 ^
-        " nor nested recursive via " ^ qsotys Ts2);
-
-    val _ = (case Library.duplicates (op =) actual_Ts of [] => () | T :: _ => duplicate_datatype T);
-
-    val perm_actual_Ts =
-      sort (prod_ord int_ord Term_Ord.typ_ord o pairself (`Term.size_of_typ)) actual_Ts;
-
-    fun the_ctrs_of (Type (s, Ts)) = map (mk_ctr Ts) (#ctrs (the (ctr_sugar_of lthy s)));
-
-    fun the_fp_sugar_of (T as Type (T_name, _)) =
-      (case fp_sugar_of lthy T_name of
-        SOME (fp_sugar as {fp = fp', ...}) => if fp = fp' then fp_sugar else not_co_datatype T
-      | NONE => not_co_datatype T);
-
-    fun gen_rhss_in gen_Ts rho subTs =
-      let
-        fun maybe_insert (T, Type (_, gen_tyargs)) =
-            if member (op =) subTs T then insert (op =) gen_tyargs else I
-          | maybe_insert _ = I;
-
-        val ctrs = maps the_ctrs_of gen_Ts;
-        val gen_ctr_Ts = maps (binder_types o fastype_of) ctrs;
-        val ctr_Ts = map (Term.typ_subst_atomic rho) gen_ctr_Ts;
-      in
-        fold (fold_subtype_pairs maybe_insert) (ctr_Ts ~~ gen_ctr_Ts) []
-      end;
-
-    fun gather_types _ _ num_groups seen gen_seen [] = (num_groups, seen, gen_seen)
-      | gather_types lthy rho num_groups seen gen_seen ((T as Type (_, tyargs)) :: Ts) =
-        let
-          val {fp_res = {Ts = mutual_Ts0, ...}, ...} = the_fp_sugar_of T;
-          val mutual_Ts = map (retypargs tyargs) mutual_Ts0;
-
-          val _ = seen = [] orelse exists (exists_subtype_in seen) mutual_Ts orelse
-            not_mutually_nested_rec mutual_Ts seen;
-
-          fun fresh_tyargs () =
-            let
-              (* The name "'z" is unlikely to clash with the context, yielding more cache hits. *)
-              val (gen_tyargs, lthy') =
-                variant_tfrees (replicate (length tyargs) "z") lthy
-                |>> map Logic.varifyT_global;
-              val rho' = (gen_tyargs ~~ tyargs) @ rho;
-            in
-              (rho', gen_tyargs, gen_seen, lthy')
-            end;
-
-          val (rho', gen_tyargs, gen_seen', lthy') =
-            if exists (exists_subtype_in seen) mutual_Ts then
-              (case gen_rhss_in gen_seen rho mutual_Ts of
-                [] => fresh_tyargs ()
-              | gen_tyargs :: gen_tyargss_tl =>
-                let
-                  val unify_pairs = split_list (maps (curry (op ~~) gen_tyargs) gen_tyargss_tl);
-                  val mgu = Type.raw_unifys unify_pairs Vartab.empty;
-                  val gen_tyargs' = map (Envir.subst_type mgu) gen_tyargs;
-                  val gen_seen' = map (Envir.subst_type mgu) gen_seen;
-                in
-                  (rho, gen_tyargs', gen_seen', lthy)
-                end)
-            else
-              fresh_tyargs ();
-
-          val gen_mutual_Ts = map (retypargs gen_tyargs) mutual_Ts0;
-          val Ts' = filter_out (member (op =) mutual_Ts) Ts;
-        in
-          gather_types lthy' rho' (num_groups + 1) (seen @ mutual_Ts) (gen_seen' @ gen_mutual_Ts)
-            Ts'
-        end
-      | gather_types _ _ _ _ _ (T :: _) = not_co_datatype T;
-
-    val (num_groups, perm_Ts, perm_gen_Ts) = gather_types lthy [] 0 [] [] perm_actual_Ts;
-    val perm_frozen_gen_Ts = map Logic.unvarifyT_global perm_gen_Ts;
-
-    val missing_Ts = perm_Ts |> subtract (op =) actual_Ts;
-    val Ts = actual_Ts @ missing_Ts;
-
-    val nn = length Ts;
-    val kks = 0 upto nn - 1;
-
-    val callssss0 = pad_list [] nn actual_callssss0;
-
-    val common_name = mk_common_name (map Binding.name_of actual_bs);
-    val bs = pad_list (Binding.name common_name) nn actual_bs;
-
-    fun permute xs = permute_like (op =) Ts perm_Ts xs;
-    fun unpermute perm_xs = permute_like (op =) perm_Ts Ts perm_xs;
-
-    val perm_bs = permute bs;
-    val perm_kks = permute kks;
-    val perm_callssss0 = permute callssss0;
-    val perm_fp_sugars0 = map (the o fp_sugar_of lthy o fst o dest_Type) perm_Ts;
-
-    val perm_callssss = map2 indexify_callsss perm_fp_sugars0 perm_callssss0;
-
-    val get_perm_indices = map (fn kk => find_index (curry (op =) kk) perm_kks) o get_indices;
-
-    val ((perm_fp_sugars, fp_sugar_thms), lthy) =
-      if num_groups > 1 then
-        mutualize_fp_sugars fp perm_bs perm_frozen_gen_Ts get_perm_indices perm_callssss
-          perm_fp_sugars0 lthy
-      else
-        ((perm_fp_sugars0, (NONE, NONE)), lthy);
-
-    val fp_sugars = unpermute perm_fp_sugars;
-  in
-    ((missing_Ts, perm_kks, fp_sugars, fp_sugar_thms), lthy)
-  end;
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_fp_n2m_tactics.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,41 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_fp_n2m_tactics.ML
-    Author:     Dmitriy Traytel, TU Muenchen
-    Copyright   2013
-
-Tactics for mutualization of nested (co)datatypes.
-*)
-
-signature BNF_FP_N2M_TACTICS =
-sig
-  val mk_rel_xtor_co_induct_tactic: BNF_FP_Util.fp_kind -> thm list -> thm list -> thm list ->
-    {prems: thm list, context: Proof.context} -> tactic
-end;
-
-structure BNF_FP_N2M_Tactics : BNF_FP_N2M_TACTICS =
-struct
-
-open BNF_Util
-open BNF_FP_Util
-
-fun mk_rel_xtor_co_induct_tactic fp co_inducts rel_defs rel_monos
-  {context = ctxt, prems = raw_C_IHs} =
-  let
-    val unfolds = map (fn def => unfold_thms ctxt (id_apply :: no_reflexive [def])) rel_defs;
-    val folded_C_IHs = map (fn thm => thm RS @{thm spec2} RS mp) raw_C_IHs;
-    val C_IHs = map2 (curry op |>) folded_C_IHs unfolds;
-    val C_IH_monos =
-      map3 (fn C_IH => fn mono => fn unfold =>
-        (mono RSN (2, @{thm rev_predicate2D}), C_IH)
-        |> fp = Greatest_FP ? swap
-        |> op RS
-        |> unfold)
-      folded_C_IHs rel_monos unfolds;
-  in
-    HEADGOAL (CONJ_WRAP_GEN' (rtac @{thm context_conjI})
-      (fn thm => rtac thm THEN_ALL_NEW (rotate_tac ~1 THEN'
-         REPEAT_ALL_NEW (FIRST' [eresolve_tac C_IHs, eresolve_tac C_IH_monos,
-           rtac @{thm order_refl}, atac, resolve_tac co_inducts])))
-    co_inducts)
-  end;
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_fp_rec_sugar_util.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,67 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_fp_rec_sugar_util.ML
-    Author:     Lorenz Panny, TU Muenchen
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2013
-
-Library for recursor and corecursor sugar.
-*)
-
-signature BNF_FP_REC_SUGAR_UTIL =
-sig
-  val indexed: 'a list -> int -> int list * int
-  val indexedd: 'a list list -> int -> int list list * int
-  val indexeddd: 'a list list list -> int -> int list list list * int
-  val indexedddd: 'a list list list list -> int -> int list list list list * int
-  val find_index_eq: ''a list -> ''a -> int
-  val finds: ('a * 'b -> bool) -> 'a list -> 'b list -> ('a * 'b list) list * 'b list
-
-  val drop_all: term -> term
-
-  val mk_partial_compN: int -> typ -> term -> term
-  val mk_partial_comp: typ -> typ -> term -> term
-  val mk_compN: int -> typ list -> term * term -> term
-  val mk_comp: typ list -> term * term -> term
-
-  val get_indices: ((binding * typ) * 'a) list -> term -> int list
-end;
-
-structure BNF_FP_Rec_Sugar_Util : BNF_FP_REC_SUGAR_UTIL =
-struct
-
-fun indexe _ h = (h, h + 1);
-fun indexed xs = fold_map indexe xs;
-fun indexedd xss = fold_map indexed xss;
-fun indexeddd xsss = fold_map indexedd xsss;
-fun indexedddd xssss = fold_map indexeddd xssss;
-
-fun find_index_eq hs h = find_index (curry (op =) h) hs;
-
-fun finds eq = fold_map (fn x => List.partition (curry eq x) #>> pair x);
-
-fun drop_all t =
-  subst_bounds (strip_qnt_vars @{const_name all} t |> map Free |> rev,
-    strip_qnt_body @{const_name all} t);
-
-fun mk_partial_comp gT fT g =
-  let val T = domain_type fT --> range_type gT in
-    Const (@{const_name Fun.comp}, gT --> fT --> T) $ g
-  end;
-
-fun mk_partial_compN 0 _ g = g
-  | mk_partial_compN n fT g =
-    let val g' = mk_partial_compN (n - 1) (range_type fT) g in
-      mk_partial_comp (fastype_of g') fT g'
-    end;
-
-fun mk_compN n bound_Ts (g, f) =
-  let val typof = curry fastype_of1 bound_Ts in
-    mk_partial_compN n (typof f) g $ f
-  end;
-
-val mk_comp = mk_compN 1;
-
-fun get_indices fixes t = map (fst #>> Binding.name_of #> Free) fixes
-  |> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE)
-  |> map_filter I;
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_fp_util.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,635 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_fp_util.ML
-    Author:     Dmitriy Traytel, TU Muenchen
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2012, 2013
-
-Shared library for the datatype and codatatype constructions.
-*)
-
-signature BNF_FP_UTIL =
-sig
-  datatype fp_kind = Least_FP | Greatest_FP
-  val fp_case: fp_kind -> 'a -> 'a -> 'a
-
-  type fp_result =
-    {Ts: typ list,
-     bnfs: BNF_Def.bnf list,
-     ctors: term list,
-     dtors: term list,
-     xtor_co_iterss: term list list,
-     xtor_co_induct: thm,
-     dtor_ctors: thm list,
-     ctor_dtors: thm list,
-     ctor_injects: thm list,
-     dtor_injects: thm list,
-     xtor_map_thms: thm list,
-     xtor_set_thmss: thm list list,
-     xtor_rel_thms: thm list,
-     xtor_co_iter_thmss: thm list list,
-     xtor_co_iter_o_map_thmss: thm list list,
-     rel_xtor_co_induct_thm: thm}
-
-  val morph_fp_result: morphism -> fp_result -> fp_result
-  val eq_fp_result: fp_result * fp_result -> bool
-  val un_fold_of: 'a list -> 'a
-  val co_rec_of: 'a list -> 'a
-
-  val time: Proof.context -> Timer.real_timer -> string -> Timer.real_timer
-
-  val IITN: string
-  val LevN: string
-  val algN: string
-  val behN: string
-  val bisN: string
-  val carTN: string
-  val caseN: string
-  val coN: string
-  val coinductN: string
-  val corecN: string
-  val ctorN: string
-  val ctor_dtorN: string
-  val ctor_exhaustN: string
-  val ctor_induct2N: string
-  val ctor_inductN: string
-  val ctor_injectN: string
-  val ctor_foldN: string
-  val ctor_fold_o_mapN: string
-  val ctor_fold_transferN: string
-  val ctor_fold_uniqueN: string
-  val ctor_mapN: string
-  val ctor_map_uniqueN: string
-  val ctor_recN: string
-  val ctor_rec_o_mapN: string
-  val ctor_rec_uniqueN: string
-  val ctor_relN: string
-  val ctor_set_inclN: string
-  val ctor_set_set_inclN: string
-  val disc_unfoldN: string
-  val disc_unfold_iffN: string
-  val disc_corecN: string
-  val disc_corec_iffN: string
-  val dtorN: string
-  val dtor_coinductN: string
-  val dtor_corecN: string
-  val dtor_corec_o_mapN: string
-  val dtor_corec_uniqueN: string
-  val dtor_ctorN: string
-  val dtor_exhaustN: string
-  val dtor_injectN: string
-  val dtor_mapN: string
-  val dtor_map_coinductN: string
-  val dtor_map_strong_coinductN: string
-  val dtor_map_uniqueN: string
-  val dtor_relN: string
-  val dtor_set_inclN: string
-  val dtor_set_set_inclN: string
-  val dtor_strong_coinductN: string
-  val dtor_unfoldN: string
-  val dtor_unfold_o_mapN: string
-  val dtor_unfold_transferN: string
-  val dtor_unfold_uniqueN: string
-  val exhaustN: string
-  val foldN: string
-  val hsetN: string
-  val hset_recN: string
-  val inductN: string
-  val injectN: string
-  val isNodeN: string
-  val lsbisN: string
-  val mapN: string
-  val map_uniqueN: string
-  val min_algN: string
-  val morN: string
-  val nchotomyN: string
-  val recN: string
-  val rel_coinductN: string
-  val rel_inductN: string
-  val rel_injectN: string
-  val rel_distinctN: string
-  val rvN: string
-  val sel_corecN: string
-  val set_inclN: string
-  val set_set_inclN: string
-  val sel_unfoldN: string
-  val setN: string
-  val simpsN: string
-  val strTN: string
-  val str_initN: string
-  val strong_coinductN: string
-  val sum_bdN: string
-  val sum_bdTN: string
-  val unfoldN: string
-  val uniqueN: string
-
-  (* TODO: Don't index set facts. Isabelle packages traditionally generate uniform names. *)
-  val mk_ctor_setN: int -> string
-  val mk_dtor_setN: int -> string
-  val mk_dtor_set_inductN: int -> string
-  val mk_set_inductN: int -> string
-
-  val co_prefix: fp_kind -> string
-
-  val base_name_of_typ: typ -> string
-  val mk_common_name: string list -> string
-
-  val split_conj_thm: thm -> thm list
-  val split_conj_prems: int -> thm -> thm
-
-  val mk_sumTN: typ list -> typ
-  val mk_sumTN_balanced: typ list -> typ
-
-  val mk_proj: typ -> int -> int -> term
-
-  val mk_convol: term * term -> term
-
-  val Inl_const: typ -> typ -> term
-  val Inr_const: typ -> typ -> term
-
-  val mk_Inl: typ -> term -> term
-  val mk_Inr: typ -> term -> term
-  val mk_InN: typ list -> term -> int -> term
-  val mk_InN_balanced: typ -> int -> term -> int -> term
-  val mk_sum_case: term * term -> term
-  val mk_sum_caseN: term list -> term
-  val mk_sum_caseN_balanced: term list -> term
-
-  val dest_sumT: typ -> typ * typ
-  val dest_sumTN: int -> typ -> typ list
-  val dest_sumTN_balanced: int -> typ -> typ list
-  val dest_tupleT: int -> typ -> typ list
-
-  val If_const: typ -> term
-
-  val mk_Field: term -> term
-  val mk_If: term -> term -> term -> term
-  val mk_union: term * term -> term
-
-  val mk_sumEN: int -> thm
-  val mk_sumEN_balanced: int -> thm
-  val mk_sumEN_tupled_balanced: int list -> thm
-  val mk_sum_casesN: int -> int -> thm
-  val mk_sum_casesN_balanced: int -> int -> thm
-
-  val fixpoint: ('a * 'a -> bool) -> ('a list -> 'a list) -> 'a list -> 'a list
-
-  val mk_rel_xtor_co_induct_thm: fp_kind -> term list -> term list -> term list -> term list ->
-    term list -> term list -> term list -> term list ->
-    ({prems: thm list, context: Proof.context} -> tactic) -> Proof.context -> thm
-  val mk_un_fold_transfer_thms: fp_kind -> term list -> term list -> term list -> term list ->
-    term list -> term list -> ({prems: thm list, context: Proof.context} -> tactic) ->
-    Proof.context -> thm list
-  val mk_xtor_un_fold_o_map_thms: fp_kind -> bool -> int -> thm -> thm list -> thm list ->
-    thm list -> thm list -> thm list
-
-  val mk_strong_coinduct_thm: thm -> thm list -> thm list -> Proof.context -> thm
-
-  val fp_bnf: (binding list -> (string * sort) list -> typ list * typ list list ->
-      BNF_Def.bnf list -> local_theory -> 'a) ->
-    binding list -> (string * sort) list -> ((string * sort) * typ) list -> local_theory ->
-    BNF_Def.bnf list * 'a
-end;
-
-structure BNF_FP_Util : BNF_FP_UTIL =
-struct
-
-open BNF_Comp
-open BNF_Def
-open BNF_Util
-
-datatype fp_kind = Least_FP | Greatest_FP;
-
-fun fp_case Least_FP l _ = l
-  | fp_case Greatest_FP _ g = g;
-
-type fp_result =
-  {Ts: typ list,
-   bnfs: BNF_Def.bnf list,
-   ctors: term list,
-   dtors: term list,
-   xtor_co_iterss: term list list,
-   xtor_co_induct: thm,
-   dtor_ctors: thm list,
-   ctor_dtors: thm list,
-   ctor_injects: thm list,
-   dtor_injects: thm list,
-   xtor_map_thms: thm list,
-   xtor_set_thmss: thm list list,
-   xtor_rel_thms: thm list,
-   xtor_co_iter_thmss: thm list list,
-   xtor_co_iter_o_map_thmss: thm list list,
-   rel_xtor_co_induct_thm: thm};
-
-fun morph_fp_result phi {Ts, bnfs, ctors, dtors, xtor_co_iterss, xtor_co_induct, dtor_ctors,
-    ctor_dtors, ctor_injects, dtor_injects, xtor_map_thms, xtor_set_thmss, xtor_rel_thms,
-    xtor_co_iter_thmss, xtor_co_iter_o_map_thmss, rel_xtor_co_induct_thm} =
-  {Ts = map (Morphism.typ phi) Ts,
-   bnfs = map (morph_bnf phi) bnfs,
-   ctors = map (Morphism.term phi) ctors,
-   dtors = map (Morphism.term phi) dtors,
-   xtor_co_iterss = map (map (Morphism.term phi)) xtor_co_iterss,
-   xtor_co_induct = Morphism.thm phi xtor_co_induct,
-   dtor_ctors = map (Morphism.thm phi) dtor_ctors,
-   ctor_dtors = map (Morphism.thm phi) ctor_dtors,
-   ctor_injects = map (Morphism.thm phi) ctor_injects,
-   dtor_injects = map (Morphism.thm phi) dtor_injects,
-   xtor_map_thms = map (Morphism.thm phi) xtor_map_thms,
-   xtor_set_thmss = map (map (Morphism.thm phi)) xtor_set_thmss,
-   xtor_rel_thms = map (Morphism.thm phi) xtor_rel_thms,
-   xtor_co_iter_thmss = map (map (Morphism.thm phi)) xtor_co_iter_thmss,
-   xtor_co_iter_o_map_thmss = map (map (Morphism.thm phi)) xtor_co_iter_o_map_thmss,
-   rel_xtor_co_induct_thm = Morphism.thm phi rel_xtor_co_induct_thm};
-
-fun eq_fp_result ({bnfs = bnfs1, ...} : fp_result, {bnfs = bnfs2, ...} : fp_result) =
-  eq_list eq_bnf (bnfs1, bnfs2);
-
-fun un_fold_of [f, _] = f;
-fun co_rec_of [_, r] = r;
-
-
-fun time ctxt timer msg = (if Config.get ctxt bnf_timing
-  then warning (msg ^ ": " ^ ATP_Util.string_of_time (Timer.checkRealTimer timer))
-  else (); Timer.startRealTimer ());
-
-val preN = "pre_"
-val rawN = "raw_"
-
-val coN = "co"
-val unN = "un"
-val algN = "alg"
-val IITN = "IITN"
-val foldN = "fold"
-val unfoldN = unN ^ foldN
-val uniqueN = "_unique"
-val transferN = "_transfer"
-val simpsN = "simps"
-val ctorN = "ctor"
-val dtorN = "dtor"
-val ctor_foldN = ctorN ^ "_" ^ foldN
-val dtor_unfoldN = dtorN ^ "_" ^ unfoldN
-val ctor_fold_uniqueN = ctor_foldN ^ uniqueN
-val ctor_fold_o_mapN = ctor_foldN ^ "_o_" ^ mapN
-val dtor_unfold_uniqueN = dtor_unfoldN ^ uniqueN
-val dtor_unfold_o_mapN = dtor_unfoldN ^ "_o_" ^ mapN
-val ctor_fold_transferN = ctor_foldN ^ transferN
-val dtor_unfold_transferN = dtor_unfoldN ^ transferN
-val ctor_mapN = ctorN ^ "_" ^ mapN
-val dtor_mapN = dtorN ^ "_" ^ mapN
-val map_uniqueN = mapN ^ uniqueN
-val ctor_map_uniqueN = ctorN ^ "_" ^ map_uniqueN
-val dtor_map_uniqueN = dtorN ^ "_" ^ map_uniqueN
-val min_algN = "min_alg"
-val morN = "mor"
-val bisN = "bis"
-val lsbisN = "lsbis"
-val sum_bdTN = "sbdT"
-val sum_bdN = "sbd"
-val carTN = "carT"
-val strTN = "strT"
-val isNodeN = "isNode"
-val LevN = "Lev"
-val rvN = "recover"
-val behN = "beh"
-val setN = "set"
-val mk_ctor_setN = prefix (ctorN ^ "_") o mk_setN
-val mk_dtor_setN = prefix (dtorN ^ "_") o mk_setN
-fun mk_set_inductN i = mk_setN i ^ "_induct"
-val mk_dtor_set_inductN = prefix (dtorN ^ "_") o mk_set_inductN
-
-val str_initN = "str_init"
-val recN = "rec"
-val corecN = coN ^ recN
-val ctor_recN = ctorN ^ "_" ^ recN
-val ctor_rec_o_mapN = ctor_recN ^ "_o_" ^ mapN
-val ctor_rec_uniqueN = ctor_recN ^ uniqueN
-val dtor_corecN = dtorN ^ "_" ^ corecN
-val dtor_corec_o_mapN = dtor_corecN ^ "_o_" ^ mapN
-val dtor_corec_uniqueN = dtor_corecN ^ uniqueN
-
-val ctor_dtorN = ctorN ^ "_" ^ dtorN
-val dtor_ctorN = dtorN ^ "_" ^ ctorN
-val nchotomyN = "nchotomy"
-val injectN = "inject"
-val exhaustN = "exhaust"
-val ctor_injectN = ctorN ^ "_" ^ injectN
-val ctor_exhaustN = ctorN ^ "_" ^ exhaustN
-val dtor_injectN = dtorN ^ "_" ^ injectN
-val dtor_exhaustN = dtorN ^ "_" ^ exhaustN
-val ctor_relN = ctorN ^ "_" ^ relN
-val dtor_relN = dtorN ^ "_" ^ relN
-val inductN = "induct"
-val coinductN = coN ^ inductN
-val ctor_inductN = ctorN ^ "_" ^ inductN
-val ctor_induct2N = ctor_inductN ^ "2"
-val dtor_map_coinductN = dtor_mapN ^ "_" ^ coinductN
-val dtor_coinductN = dtorN ^ "_" ^ coinductN
-val strong_coinductN = "strong_" ^ coinductN
-val dtor_map_strong_coinductN = dtor_mapN ^ "_" ^ strong_coinductN
-val dtor_strong_coinductN = dtorN ^ "_" ^ strong_coinductN
-val hsetN = "Hset"
-val hset_recN = hsetN ^ "_rec"
-val set_inclN = "set_incl"
-val ctor_set_inclN = ctorN ^ "_" ^ set_inclN
-val dtor_set_inclN = dtorN ^ "_" ^ set_inclN
-val set_set_inclN = "set_set_incl"
-val ctor_set_set_inclN = ctorN ^ "_" ^ set_set_inclN
-val dtor_set_set_inclN = dtorN ^ "_" ^ set_set_inclN
-
-val caseN = "case"
-val discN = "disc"
-val disc_unfoldN = discN ^ "_" ^ unfoldN
-val disc_corecN = discN ^ "_" ^ corecN
-val iffN = "_iff"
-val disc_unfold_iffN = discN ^ "_" ^ unfoldN ^ iffN
-val disc_corec_iffN = discN ^ "_" ^ corecN ^ iffN
-val distinctN = "distinct"
-val rel_distinctN = relN ^ "_" ^ distinctN
-val injectN = "inject"
-val rel_injectN = relN ^ "_" ^ injectN
-val rel_coinductN = relN ^ "_" ^ coinductN
-val rel_inductN = relN ^ "_" ^ inductN
-val selN = "sel"
-val sel_unfoldN = selN ^ "_" ^ unfoldN
-val sel_corecN = selN ^ "_" ^ corecN
-
-fun co_prefix fp = (if fp = Greatest_FP then "co" else "");
-
-fun add_components_of_typ (Type (s, Ts)) =
-    cons (Long_Name.base_name s) #> fold_rev add_components_of_typ Ts
-  | add_components_of_typ _ = I;
-
-fun base_name_of_typ T = space_implode "_" (add_components_of_typ T []);
-
-val mk_common_name = space_implode "_";
-
-fun dest_sumT (Type (@{type_name sum}, [T, T'])) = (T, T');
-
-fun dest_sumTN 1 T = [T]
-  | dest_sumTN n (Type (@{type_name sum}, [T, T'])) = T :: dest_sumTN (n - 1) T';
-
-val dest_sumTN_balanced = Balanced_Tree.dest dest_sumT;
-
-(* TODO: move something like this to "HOLogic"? *)
-fun dest_tupleT 0 @{typ unit} = []
-  | dest_tupleT 1 T = [T]
-  | dest_tupleT n (Type (@{type_name prod}, [T, T'])) = T :: dest_tupleT (n - 1) T';
-
-val mk_sumTN = Library.foldr1 mk_sumT;
-val mk_sumTN_balanced = Balanced_Tree.make mk_sumT;
-
-fun mk_proj T n k =
-  let val (binders, _) = strip_typeN n T in
-    fold_rev (fn T => fn t => Abs (Name.uu, T, t)) binders (Bound (n - k - 1))
-  end;
-
-fun mk_convol (f, g) =
-  let
-    val (fU, fTU) = `range_type (fastype_of f);
-    val ((gT, gU), gTU) = `dest_funT (fastype_of g);
-    val convolT = fTU --> gTU --> gT --> HOLogic.mk_prodT (fU, gU);
-  in Const (@{const_name convol}, convolT) $ f $ g end;
-
-fun Inl_const LT RT = Const (@{const_name Inl}, LT --> mk_sumT (LT, RT));
-fun mk_Inl RT t = Inl_const (fastype_of t) RT $ t;
-
-fun Inr_const LT RT = Const (@{const_name Inr}, RT --> mk_sumT (LT, RT));
-fun mk_Inr LT t = Inr_const LT (fastype_of t) $ t;
-
-fun mk_InN [_] t 1 = t
-  | mk_InN (_ :: Ts) t 1 = mk_Inl (mk_sumTN Ts) t
-  | mk_InN (LT :: Ts) t m = mk_Inr LT (mk_InN Ts t (m - 1))
-  | mk_InN Ts t _ = raise (TYPE ("mk_InN", Ts, [t]));
-
-fun mk_InN_balanced sum_T n t k =
-  let
-    fun repair_types T (Const (s as @{const_name Inl}, _) $ t) = repair_inj_types T s fst t
-      | repair_types T (Const (s as @{const_name Inr}, _) $ t) = repair_inj_types T s snd t
-      | repair_types _ t = t
-    and repair_inj_types T s get t =
-      let val T' = get (dest_sumT T) in
-        Const (s, T' --> T) $ repair_types T' t
-      end;
-  in
-    Balanced_Tree.access {left = mk_Inl dummyT, right = mk_Inr dummyT, init = t} n k
-    |> repair_types sum_T
-  end;
-
-fun mk_sum_case (f, g) =
-  let
-    val fT = fastype_of f;
-    val gT = fastype_of g;
-  in
-    Const (@{const_name sum_case},
-      fT --> gT --> mk_sumT (domain_type fT, domain_type gT) --> range_type fT) $ f $ g
-  end;
-
-val mk_sum_caseN = Library.foldr1 mk_sum_case;
-val mk_sum_caseN_balanced = Balanced_Tree.make mk_sum_case;
-
-fun If_const T = Const (@{const_name If}, HOLogic.boolT --> T --> T --> T);
-fun mk_If p t f = let val T = fastype_of t in If_const T $ p $ t $ f end;
-
-fun mk_Field r =
-  let val T = fst (dest_relT (fastype_of r));
-  in Const (@{const_name Field}, mk_relT (T, T) --> HOLogic.mk_setT T) $ r end;
-
-val mk_union = HOLogic.mk_binop @{const_name sup};
-
-(*dangerous; use with monotonic, converging functions only!*)
-fun fixpoint eq f X = if subset eq (f X, X) then X else fixpoint eq f (f X);
-
-(* stolen from "~~/src/HOL/Tools/Datatype/datatype_aux.ML" *)
-fun split_conj_thm th =
-  ((th RS conjunct1) :: split_conj_thm (th RS conjunct2)) handle THM _ => [th];
-
-fun split_conj_prems limit th =
-  let
-    fun split n i th =
-      if i = n then th else split n (i + 1) (conjI RSN (i, th)) handle THM _ => th;
-  in split limit 1 th end;
-
-fun mk_sumEN 1 = @{thm one_pointE}
-  | mk_sumEN 2 = @{thm sumE}
-  | mk_sumEN n =
-    (fold (fn i => fn thm => @{thm obj_sumE_f} RSN (i, thm)) (2 upto n - 1) @{thm obj_sumE}) OF
-      replicate n (impI RS allI);
-
-fun mk_obj_sumEN_balanced n =
-  Balanced_Tree.make (fn (thm1, thm2) => thm1 RSN (1, thm2 RSN (2, @{thm obj_sumE_f})))
-    (replicate n asm_rl);
-
-fun mk_sumEN_balanced' n all_impIs = mk_obj_sumEN_balanced n OF all_impIs RS @{thm obj_one_pointE};
-
-fun mk_sumEN_balanced 1 = @{thm one_pointE} (*optimization*)
-  | mk_sumEN_balanced 2 = @{thm sumE} (*optimization*)
-  | mk_sumEN_balanced n = mk_sumEN_balanced' n (replicate n (impI RS allI));
-
-fun mk_tupled_allIN 0 = @{thm unit_all_impI}
-  | mk_tupled_allIN 1 = @{thm impI[THEN allI]}
-  | mk_tupled_allIN 2 = @{thm prod_all_impI} (*optimization*)
-  | mk_tupled_allIN n = mk_tupled_allIN (n - 1) RS @{thm prod_all_impI_step};
-
-fun mk_sumEN_tupled_balanced ms =
-  let val n = length ms in
-    if forall (curry op = 1) ms then mk_sumEN_balanced n
-    else mk_sumEN_balanced' n (map mk_tupled_allIN ms)
-  end;
-
-fun mk_sum_casesN 1 1 = refl
-  | mk_sum_casesN _ 1 = @{thm sum.cases(1)}
-  | mk_sum_casesN 2 2 = @{thm sum.cases(2)}
-  | mk_sum_casesN n k = trans OF [@{thm sum_case_step(2)}, mk_sum_casesN (n - 1) (k - 1)];
-
-fun mk_sum_step base step thm =
-  if Thm.eq_thm_prop (thm, refl) then base else trans OF [step, thm];
-
-fun mk_sum_casesN_balanced 1 1 = refl
-  | mk_sum_casesN_balanced n k =
-    Balanced_Tree.access {left = mk_sum_step @{thm sum.cases(1)} @{thm sum_case_step(1)},
-      right = mk_sum_step @{thm sum.cases(2)} @{thm sum_case_step(2)}, init = refl} n k;
-
-fun mk_rel_xtor_co_induct_thm fp pre_rels pre_phis rels phis xs ys xtors xtor's tac lthy =
-  let
-    val pre_relphis = map (fn rel => Term.list_comb (rel, phis @ pre_phis)) pre_rels;
-    val relphis = map (fn rel => Term.list_comb (rel, phis)) rels;
-    fun mk_xtor fp' xtor x = if fp = fp' then xtor $ x else x;
-    val dtor = mk_xtor Greatest_FP;
-    val ctor = mk_xtor Least_FP;
-    fun flip f x y = if fp = Greatest_FP then f y x else f x y;
-
-    fun mk_prem pre_relphi phi x y xtor xtor' =
-      HOLogic.mk_Trueprop (list_all_free [x, y] (flip (curry HOLogic.mk_imp)
-        (pre_relphi $ (dtor xtor x) $ (dtor xtor' y)) (phi $ (ctor xtor x) $ (ctor xtor' y))));
-    val prems = map6 mk_prem pre_relphis pre_phis xs ys xtors xtor's;
-
-    val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
-      (map2 (flip mk_leq) relphis pre_phis));
-  in
-    Goal.prove_sorry lthy (map (fst o dest_Free) (phis @ pre_phis)) prems concl tac
-    |> Thm.close_derivation
-    |> (fn thm => thm OF (replicate (length pre_rels) @{thm allI[OF allI[OF impI]]}))
-  end;
-
-fun mk_un_fold_transfer_thms fp pre_rels pre_phis rels phis un_folds un_folds' tac lthy =
-  let
-    val pre_relphis = map (fn rel => Term.list_comb (rel, phis @ pre_phis)) pre_rels;
-    val relphis = map (fn rel => Term.list_comb (rel, phis)) rels;
-    fun flip f x y = if fp = Greatest_FP then f y x else f x y;
-
-    val arg_rels = map2 (flip mk_fun_rel) pre_relphis pre_phis;
-    fun mk_transfer relphi pre_phi un_fold un_fold' =
-      fold_rev mk_fun_rel arg_rels (flip mk_fun_rel relphi pre_phi) $ un_fold $ un_fold';
-    val transfers = map4 mk_transfer relphis pre_phis un_folds un_folds';
-
-    val goal = fold_rev Logic.all (phis @ pre_phis)
-      (HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj transfers));
-  in
-    Goal.prove_sorry lthy [] [] goal tac
-    |> Thm.close_derivation
-    |> split_conj_thm
-  end;
-
-fun mk_xtor_un_fold_o_map_thms fp is_rec m un_fold_unique xtor_maps xtor_un_folds sym_map_comps
-    map_cong0s =
-  let
-    val n = length sym_map_comps;
-    val rewrite_comp_comp2 = fp_case fp @{thm rewriteR_comp_comp2} @{thm rewriteL_comp_comp2};
-    val rewrite_comp_comp = fp_case fp @{thm rewriteR_comp_comp} @{thm rewriteL_comp_comp};
-    val map_cong_passive_args1 = replicate m (fp_case fp @{thm id_o} @{thm o_id} RS fun_cong);
-    val map_cong_active_args1 = replicate n (if is_rec
-      then fp_case fp @{thm convol_o} @{thm o_sum_case} RS fun_cong
-      else refl);
-    val map_cong_passive_args2 = replicate m (fp_case fp @{thm o_id} @{thm id_o} RS fun_cong);
-    val map_cong_active_args2 = replicate n (if is_rec
-      then fp_case fp @{thm map_pair_o_convol_id} @{thm sum_case_o_sum_map_id}
-      else fp_case fp @{thm id_o} @{thm o_id} RS fun_cong);
-    fun mk_map_congs passive active = map (fn thm => thm OF (passive @ active) RS ext) map_cong0s;
-    val map_cong1s = mk_map_congs map_cong_passive_args1 map_cong_active_args1;
-    val map_cong2s = mk_map_congs map_cong_passive_args2 map_cong_active_args2;
-    
-    fun mk_rewrites map_congs = map2 (fn sym_map_comp => fn map_cong =>
-      mk_trans sym_map_comp map_cong RS rewrite_comp_comp) sym_map_comps map_congs;
-    val rewrite1s = mk_rewrites map_cong1s;
-    val rewrite2s = mk_rewrites map_cong2s;
-    val unique_prems =
-      map4 (fn xtor_map => fn un_fold => fn rewrite1 => fn rewrite2 =>
-        mk_trans (rewrite_comp_comp2 OF [xtor_map, un_fold])
-          (mk_trans rewrite1 (mk_sym rewrite2)))
-      xtor_maps xtor_un_folds rewrite1s rewrite2s;
-  in
-    split_conj_thm (un_fold_unique OF map (fp_case fp I mk_sym) unique_prems)
-  end;
-
-fun mk_strong_coinduct_thm coind rel_eqs rel_monos ctxt =
-  let
-    val n = Thm.nprems_of coind;
-    val m = Thm.nprems_of (hd rel_monos) - n;
-    fun mk_inst phi = (phi, mk_union (phi, HOLogic.eq_const (fst (dest_pred2T (fastype_of phi)))))
-      |> pairself (certify ctxt);
-    val insts = Term.add_vars (Thm.prop_of coind) [] |> rev |> take n |> map (mk_inst o Var);
-    fun mk_unfold rel_eq rel_mono =
-      let
-        val eq = iffD2 OF [rel_eq RS @{thm predicate2_eqD}, refl];
-        val mono = rel_mono OF (replicate m @{thm order_refl} @ replicate n @{thm eq_subset});
-      in eq RS (mono RS @{thm predicate2D}) RS @{thm eqTrueI} end;
-    val unfolds = map2 mk_unfold rel_eqs rel_monos @ @{thms sup_fun_def sup_bool_def
-      imp_disjL all_conj_distrib subst_eq_imp simp_thms(18,21,35)};
-  in
-    Thm.instantiate ([], insts) coind
-    |> unfold_thms ctxt unfolds
-  end;
-
-fun fp_bnf construct_fp bs resBs fp_eqs lthy =
-  let
-    val time = time lthy;
-    val timer = time (Timer.startRealTimer ());
-    val (Xs, rhsXs) = split_list fp_eqs;
-
-    (* FIXME: because of "@ Xs", the output could contain type variables that are not in the
-       input; also, "fp_sort" should put the "resBs" first and in the order in which they appear *)
-    fun fp_sort Ass =
-      subtract (op =) Xs (filter (fn T => exists (fn Ts => member (op =) Ts T) Ass) resBs) @ Xs;
-
-    fun raw_qualify base_b =
-      let val (_, qs, n) = Binding.dest base_b;
-      in
-        Binding.prefix_name rawN
-        #> fold_rev (fn (s, mand) => Binding.qualify mand s) (qs @ [(n, true)])
-        #> Binding.conceal
-      end;
-
-    val ((bnfs, (deadss, livess)), (unfold_set, lthy)) = apfst (apsnd split_list o split_list)
-      (fold_map2 (fn b => bnf_of_typ Smart_Inline (raw_qualify b) fp_sort Xs) bs rhsXs
-        (empty_unfolds, lthy));
-
-    fun norm_qualify i = Binding.qualify true (Binding.name_of (nth bs (Int.max (0, i - 1))))
-      #> Binding.conceal;
-
-    val Ass = map (map dest_TFree) livess;
-    val resDs = fold (subtract (op =)) Ass resBs;
-    val Ds = fold (fold Term.add_tfreesT) deadss [];
-
-    val timer = time (timer "Construction of BNFs");
-
-    val ((kill_poss, _), (bnfs', (unfold_set', lthy'))) =
-      normalize_bnfs norm_qualify Ass Ds fp_sort bnfs unfold_set lthy;
-
-    val Dss = map3 (append oo map o nth) livess kill_poss deadss;
-
-    fun pre_qualify b = Binding.qualify false (Binding.name_of b)
-      #> Config.get lthy' bnf_note_all = false ? Binding.conceal;
-
-    val ((pre_bnfs, deadss), lthy'') =
-      fold_map3 (fn b => seal_bnf (pre_qualify b) unfold_set' (Binding.prefix_name preN b))
-        bs Dss bnfs' lthy'
-      |>> split_list;
-
-    val timer = time (timer "Normalization & sealing of BNFs");
-
-    val res = construct_fp bs resBs (map TFree resDs, deadss) pre_bnfs lthy'';
-
-    val timer = time (timer "FP construction in total");
-  in
-    timer; (pre_bnfs, res)
-  end;
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_gfp.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2827 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_gfp.ML
-    Author:     Dmitriy Traytel, TU Muenchen
-    Author:     Andrei Popescu, TU Muenchen
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2012
-
-Codatatype construction.
-*)
-
-signature BNF_GFP =
-sig
-  val construct_gfp: mixfix list -> binding list -> binding list -> binding list list ->
-    binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list ->
-    local_theory -> BNF_FP_Util.fp_result * local_theory
-end;
-
-structure BNF_GFP : BNF_GFP =
-struct
-
-open BNF_Def
-open BNF_Util
-open BNF_Tactics
-open BNF_Comp
-open BNF_FP_Util
-open BNF_FP_Def_Sugar
-open BNF_GFP_Rec_Sugar
-open BNF_GFP_Util
-open BNF_GFP_Tactics
-
-datatype wit_tree = Wit_Leaf of int | Wit_Node of (int * int * int list) * wit_tree list;
-
-fun mk_tree_args (I, T) (I', Ts) = (sort_distinct int_ord (I @ I'), T :: Ts);
-
-fun finish Iss m seen i (nwit, I) =
-  let
-    val treess = map (fn j =>
-        if j < m orelse member (op =) seen j then [([j], Wit_Leaf j)]
-        else
-          map_index (finish Iss m (insert (op =) j seen) j) (nth Iss (j - m))
-          |> flat
-          |> minimize_wits)
-      I;
-  in
-    map (fn (I, t) => (I, Wit_Node ((i - m, nwit, filter (fn i => i < m) I), t)))
-      (fold_rev (map_product mk_tree_args) treess [([], [])])
-    |> minimize_wits
-  end;
-
-fun tree_to_ctor_wit vars _ _ (Wit_Leaf j) = ([j], nth vars j)
-  | tree_to_ctor_wit vars ctors witss (Wit_Node ((i, nwit, I), subtrees)) =
-     (I, nth ctors i $ (Term.list_comb (snd (nth (nth witss i) nwit),
-       map (snd o tree_to_ctor_wit vars ctors witss) subtrees)));
-
-fun tree_to_coind_wits _ (Wit_Leaf _) = []
-  | tree_to_coind_wits lwitss (Wit_Node ((i, nwit, I), subtrees)) =
-     ((i, I), nth (nth lwitss i) nwit) :: maps (tree_to_coind_wits lwitss) subtrees;
-
-(*all BNFs have the same lives*)
-fun construct_gfp mixfixes map_bs rel_bs set_bss0 bs resBs (resDs, Dss) bnfs lthy =
-  let
-    val time = time lthy;
-    val timer = time (Timer.startRealTimer ());
-
-    val live = live_of_bnf (hd bnfs);
-    val n = length bnfs; (*active*)
-    val ks = 1 upto n;
-    val m = live - n; (*passive, if 0 don't generate a new BNF*)
-    val ls = 1 upto m;
-
-    val note_all = Config.get lthy bnf_note_all;
-    val b_names = map Binding.name_of bs;
-    val b_name = mk_common_name b_names;
-    val b = Binding.name b_name;
-    val mk_internal_b = Binding.name #> Binding.prefix true b_name #> Binding.conceal;
-    fun mk_internal_bs name =
-      map (fn b =>
-        Binding.prefix true b_name (Binding.prefix_name (name ^ "_") b) |> Binding.conceal) bs;
-    val external_bs = map2 (Binding.prefix false) b_names bs
-      |> note_all = false ? map Binding.conceal;
-
-    (* TODO: check if m, n, etc., are sane *)
-
-    val deads = fold (union (op =)) Dss resDs;
-    val names_lthy = fold Variable.declare_typ deads lthy;
-    val passives = map fst (subtract (op = o apsnd TFree) deads resBs);
-
-    (* tvars *)
-    val ((((((passiveAs, activeAs), passiveBs), activeBs), passiveCs), activeCs), idxT) = names_lthy
-      |> variant_tfrees passives
-      ||>> mk_TFrees n
-      ||>> variant_tfrees passives
-      ||>> mk_TFrees n
-      ||>> mk_TFrees m
-      ||>> mk_TFrees n
-      ||> fst o mk_TFrees 1
-      ||> the_single;
-
-    val allAs = passiveAs @ activeAs;
-    val allBs' = passiveBs @ activeBs;
-    val Ass = replicate n allAs;
-    val allBs = passiveAs @ activeBs;
-    val Bss = replicate n allBs;
-    val allCs = passiveAs @ activeCs;
-    val allCs' = passiveBs @ activeCs;
-    val Css' = replicate n allCs';
-
-    (* types *)
-    val dead_poss =
-      map (fn x => if member (op =) deads (TFree x) then SOME (TFree x) else NONE) resBs;
-    fun mk_param NONE passive = (hd passive, tl passive)
-      | mk_param (SOME a) passive = (a, passive);
-    val mk_params = fold_map mk_param dead_poss #> fst;
-
-    fun mk_FTs Ts = map2 (fn Ds => mk_T_of_bnf Ds Ts) Dss bnfs;
-    val (params, params') = `(map Term.dest_TFree) (mk_params passiveAs);
-    val (dead_params, dead_params') = `(map Term.dest_TFree) (subtract (op =) passiveAs params');
-    val FTsAs = mk_FTs allAs;
-    val FTsBs = mk_FTs allBs;
-    val FTsCs = mk_FTs allCs;
-    val ATs = map HOLogic.mk_setT passiveAs;
-    val BTs = map HOLogic.mk_setT activeAs;
-    val B'Ts = map HOLogic.mk_setT activeBs;
-    val B''Ts = map HOLogic.mk_setT activeCs;
-    val sTs = map2 (fn T => fn U => T --> U) activeAs FTsAs;
-    val s'Ts = map2 (fn T => fn U => T --> U) activeBs FTsBs;
-    val s''Ts = map2 (fn T => fn U => T --> U) activeCs FTsCs;
-    val fTs = map2 (fn T => fn U => T --> U) activeAs activeBs;
-    val self_fTs = map (fn T => T --> T) activeAs;
-    val gTs = map2 (fn T => fn U => T --> U) activeBs activeCs;
-    val all_gTs = map2 (fn T => fn U => T --> U) allBs allCs';
-    val RTs = map2 (fn T => fn U => HOLogic.mk_prodT (T, U)) activeAs activeBs;
-    val sRTs = map2 (fn T => fn U => HOLogic.mk_prodT (T, U)) activeAs activeAs;
-    val R'Ts = map2 (fn T => fn U => HOLogic.mk_prodT (T, U)) activeBs activeCs;
-    val setsRTs = map HOLogic.mk_setT sRTs;
-    val setRTs = map HOLogic.mk_setT RTs;
-    val all_sbisT = HOLogic.mk_tupleT setsRTs;
-    val setR'Ts = map HOLogic.mk_setT R'Ts;
-    val FRTs = mk_FTs (passiveAs @ RTs);
-    val sumBsAs = map2 (curry mk_sumT) activeBs activeAs;
-    val sumFTs = mk_FTs (passiveAs @ sumBsAs);
-    val sum_sTs = map2 (fn T => fn U => T --> U) activeAs sumFTs;
-
-    (* terms *)
-    val mapsAsAs = map4 mk_map_of_bnf Dss Ass Ass bnfs;
-    val mapsAsBs = map4 mk_map_of_bnf Dss Ass Bss bnfs;
-    val mapsBsCs' = map4 mk_map_of_bnf Dss Bss Css' bnfs;
-    val mapsAsCs' = map4 mk_map_of_bnf Dss Ass Css' bnfs;
-    val map_Inls = map4 mk_map_of_bnf Dss Bss (replicate n (passiveAs @ sumBsAs)) bnfs;
-    val map_Inls_rev = map4 mk_map_of_bnf Dss (replicate n (passiveAs @ sumBsAs)) Bss bnfs;
-    val map_fsts = map4 mk_map_of_bnf Dss (replicate n (passiveAs @ RTs)) Ass bnfs;
-    val map_snds = map4 mk_map_of_bnf Dss (replicate n (passiveAs @ RTs)) Bss bnfs;
-    fun mk_setss Ts = map3 mk_sets_of_bnf (map (replicate live) Dss)
-      (map (replicate live) (replicate n Ts)) bnfs;
-    val setssAs = mk_setss allAs;
-    val setssAs' = transpose setssAs;
-    val bis_setss = mk_setss (passiveAs @ RTs);
-    val relsAsBs = map4 mk_rel_of_bnf Dss Ass Bss bnfs;
-    val bds = map3 mk_bd_of_bnf Dss Ass bnfs;
-    val sum_bd = Library.foldr1 (uncurry mk_csum) bds;
-    val sum_bdT = fst (dest_relT (fastype_of sum_bd));
-
-    val emptys = map (fn T => HOLogic.mk_set T []) passiveAs;
-    val Zeros = map (fn empty =>
-     HOLogic.mk_tuple (map (fn U => absdummy U empty) activeAs)) emptys;
-    val hrecTs = map fastype_of Zeros;
-    val hsetTs = map (fn hrecT => Library.foldr (op -->) (sTs, HOLogic.natT --> hrecT)) hrecTs;
-
-    val (((((((((((((((((((((((((((((((((((zs, zs'), zs_copy), zs_copy2), z's), (ys, ys')),
-      As), Bs), Bs_copy), B's), B''s), ss), sum_ss), s's), s''s), fs), fs_copy),
-      self_fs), gs), all_gs), xFs), yFs), yFs_copy), RFs), (Rtuple, Rtuple')), (hrecs, hrecs')),
-      (nat, nat')), Rs), Rs_copy), R's), sRs), (idx, idx')), Idx), Ris), Kss), names_lthy) = lthy
-      |> mk_Frees' "b" activeAs
-      ||>> mk_Frees "b" activeAs
-      ||>> mk_Frees "b" activeAs
-      ||>> mk_Frees "b" activeBs
-      ||>> mk_Frees' "y" passiveAs
-      ||>> mk_Frees "A" ATs
-      ||>> mk_Frees "B" BTs
-      ||>> mk_Frees "B" BTs
-      ||>> mk_Frees "B'" B'Ts
-      ||>> mk_Frees "B''" B''Ts
-      ||>> mk_Frees "s" sTs
-      ||>> mk_Frees "sums" sum_sTs
-      ||>> mk_Frees "s'" s'Ts
-      ||>> mk_Frees "s''" s''Ts
-      ||>> mk_Frees "f" fTs
-      ||>> mk_Frees "f" fTs
-      ||>> mk_Frees "f" self_fTs
-      ||>> mk_Frees "g" gTs
-      ||>> mk_Frees "g" all_gTs
-      ||>> mk_Frees "x" FTsAs
-      ||>> mk_Frees "y" FTsBs
-      ||>> mk_Frees "y" FTsBs
-      ||>> mk_Frees "x" FRTs
-      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "Rtuple") all_sbisT
-      ||>> mk_Frees' "rec" hrecTs
-      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "n") HOLogic.natT
-      ||>> mk_Frees "R" setRTs
-      ||>> mk_Frees "R" setRTs
-      ||>> mk_Frees "R'" setR'Ts
-      ||>> mk_Frees "R" setsRTs
-      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "i") idxT
-      ||>> yield_singleton (mk_Frees "I") (HOLogic.mk_setT idxT)
-      ||>> mk_Frees "Ri" (map (fn T => idxT --> T) setRTs)
-      ||>> mk_Freess "K" (map (fn AT => map (fn T => T --> AT) activeAs) ATs);
-
-    val passive_UNIVs = map HOLogic.mk_UNIV passiveAs;
-    val passive_Id_ons = map mk_Id_on As;
-    val active_UNIVs = map HOLogic.mk_UNIV activeAs;
-    val sum_UNIVs = map HOLogic.mk_UNIV sumBsAs;
-    val passive_ids = map HOLogic.id_const passiveAs;
-    val active_ids = map HOLogic.id_const activeAs;
-    val Inls = map2 Inl_const activeBs activeAs;
-    val fsts = map fst_const RTs;
-    val snds = map snd_const RTs;
-
-    (* thms *)
-    val bd_card_orders = map bd_card_order_of_bnf bnfs;
-    val bd_card_order = hd bd_card_orders
-    val bd_Card_orders = map bd_Card_order_of_bnf bnfs;
-    val bd_Card_order = hd bd_Card_orders;
-    val bd_Cinfinites = map bd_Cinfinite_of_bnf bnfs;
-    val bd_Cinfinite = hd bd_Cinfinites;
-    val in_monos = map in_mono_of_bnf bnfs;
-    val map_comp0s = map map_comp0_of_bnf bnfs;
-    val sym_map_comps = map mk_sym map_comp0s;
-    val map_comps = map map_comp_of_bnf bnfs;
-    val map_cong0s = map map_cong0_of_bnf bnfs;
-    val map_id0s = map map_id0_of_bnf bnfs;
-    val map_ids = map map_id_of_bnf bnfs;
-    val set_bdss = map set_bd_of_bnf bnfs;
-    val set_mapss = map set_map_of_bnf bnfs;
-    val rel_congs = map rel_cong_of_bnf bnfs;
-    val rel_converseps = map rel_conversep_of_bnf bnfs;
-    val rel_Grps = map rel_Grp_of_bnf bnfs;
-    val rel_OOs = map rel_OO_of_bnf bnfs;
-    val rel_OO_Grps = map rel_OO_Grp_of_bnf bnfs;
-
-    val timer = time (timer "Extracted terms & thms");
-
-    (* derived thms *)
-
-    (*map g1 ... gm g(m+1) ... g(m+n) (map id ... id f(m+1) ... f(m+n) x) =
-      map g1 ... gm (g(m+1) o f(m+1)) ... (g(m+n) o f(m+n)) x*)
-    fun mk_map_comp_id x mapAsBs mapBsCs mapAsCs map_comp0 =
-      let
-        val lhs = Term.list_comb (mapBsCs, all_gs) $
-          (Term.list_comb (mapAsBs, passive_ids @ fs) $ x);
-        val rhs =
-          Term.list_comb (mapAsCs, take m all_gs @ map HOLogic.mk_comp (drop m all_gs ~~ fs)) $ x;
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (x :: fs @ all_gs) (mk_Trueprop_eq (lhs, rhs)))
-          (K (mk_map_comp_id_tac map_comp0))
-        |> Thm.close_derivation
-      end;
-
-    val map_comp_id_thms = map5 mk_map_comp_id xFs mapsAsBs mapsBsCs' mapsAsCs' map_comps;
-
-    (*forall a : set(m+1) x. f(m+1) a = a; ...; forall a : set(m+n) x. f(m+n) a = a ==>
-      map id ... id f(m+1) ... f(m+n) x = x*)
-    fun mk_map_cong0L x mapAsAs sets map_cong0 map_id =
-      let
-        fun mk_prem set f z z' =
-          HOLogic.mk_Trueprop
-            (mk_Ball (set $ x) (Term.absfree z' (HOLogic.mk_eq (f $ z, z))));
-        val prems = map4 mk_prem (drop m sets) self_fs zs zs';
-        val goal = mk_Trueprop_eq (Term.list_comb (mapAsAs, passive_ids @ self_fs) $ x, x);
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (x :: self_fs) (Logic.list_implies (prems, goal)))
-          (K (mk_map_cong0L_tac m map_cong0 map_id))
-        |> Thm.close_derivation
-      end;
-
-    val map_cong0L_thms = map5 mk_map_cong0L xFs mapsAsAs setssAs map_cong0s map_ids;
-    val in_mono'_thms = map (fn thm =>
-      (thm OF (replicate m subset_refl)) RS @{thm set_mp}) in_monos;
-
-    val map_arg_cong_thms =
-      let
-        val prems = map2 (curry mk_Trueprop_eq) yFs yFs_copy;
-        val maps = map (fn mapx => Term.list_comb (mapx, all_gs)) mapsBsCs';
-        val concls =
-          map3 (fn x => fn y => fn mapx => mk_Trueprop_eq (mapx $ x, mapx $ y)) yFs yFs_copy maps;
-        val goals =
-          map4 (fn prem => fn concl => fn x => fn y =>
-            fold_rev Logic.all (x :: y :: all_gs) (Logic.mk_implies (prem, concl)))
-          prems concls yFs yFs_copy;
-      in
-        map (fn goal => Goal.prove_sorry lthy [] [] goal
-          (K ((hyp_subst_tac lthy THEN' rtac refl) 1)) |> Thm.close_derivation) goals
-      end;
-
-    val timer = time (timer "Derived simple theorems");
-
-    (* coalgebra *)
-
-    val coalg_bind = mk_internal_b (coN ^ algN) ;
-    val coalg_name = Binding.name_of coalg_bind;
-    val coalg_def_bind = (Thm.def_binding coalg_bind, []);
-
-    (*forall i = 1 ... n: (\<forall>x \<in> Bi. si \<in> Fi_in A1 .. Am B1 ... Bn)*)
-    val coalg_spec =
-      let
-        val coalgT = Library.foldr (op -->) (ATs @ BTs @ sTs, HOLogic.boolT);
-
-        val ins = map3 mk_in (replicate n (As @ Bs)) setssAs FTsAs;
-        fun mk_coalg_conjunct B s X z z' =
-          mk_Ball B (Term.absfree z' (HOLogic.mk_mem (s $ z, X)));
-
-        val lhs = Term.list_comb (Free (coalg_name, coalgT), As @ Bs @ ss);
-        val rhs = Library.foldr1 HOLogic.mk_conj (map5 mk_coalg_conjunct Bs ss ins zs zs')
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((coalg_free, (_, coalg_def_free)), (lthy, lthy_old)) =
-      lthy
-      |> Specification.definition (SOME (coalg_bind, NONE, NoSyn), (coalg_def_bind, coalg_spec))
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-    val coalg = fst (Term.dest_Const (Morphism.term phi coalg_free));
-    val coalg_def = Morphism.thm phi coalg_def_free;
-
-    fun mk_coalg As Bs ss =
-      let
-        val args = As @ Bs @ ss;
-        val Ts = map fastype_of args;
-        val coalgT = Library.foldr (op -->) (Ts, HOLogic.boolT);
-      in
-        Term.list_comb (Const (coalg, coalgT), args)
-      end;
-
-    val coalg_prem = HOLogic.mk_Trueprop (mk_coalg As Bs ss);
-
-    val coalg_in_thms = map (fn i =>
-      coalg_def RS iffD1 RS mk_conjunctN n i RS bspec) ks
-
-    val coalg_set_thmss =
-      let
-        val coalg_prem = HOLogic.mk_Trueprop (mk_coalg As Bs ss);
-        fun mk_prem x B = HOLogic.mk_Trueprop (HOLogic.mk_mem (x, B));
-        fun mk_concl s x B set = HOLogic.mk_Trueprop (mk_leq (set $ (s $ x)) B);
-        val prems = map2 mk_prem zs Bs;
-        val conclss = map3 (fn s => fn x => fn sets => map2 (mk_concl s x) (As @ Bs) sets)
-          ss zs setssAs;
-        val goalss = map3 (fn x => fn prem => fn concls => map (fn concl =>
-          fold_rev Logic.all (x :: As @ Bs @ ss)
-            (Logic.list_implies (coalg_prem :: [prem], concl))) concls) zs prems conclss;
-      in
-        map (fn goals => map (fn goal => Goal.prove_sorry lthy [] [] goal
-          (K (mk_coalg_set_tac coalg_def)) |> Thm.close_derivation) goals) goalss
-      end;
-
-    fun mk_tcoalg ATs BTs = mk_coalg (map HOLogic.mk_UNIV ATs) (map HOLogic.mk_UNIV BTs);
-
-    val tcoalg_thm =
-      let
-        val goal = fold_rev Logic.all ss
-          (HOLogic.mk_Trueprop (mk_tcoalg passiveAs activeAs ss))
-      in
-        Goal.prove_sorry lthy [] [] goal
-          (K (stac coalg_def 1 THEN CONJ_WRAP
-            (K (EVERY' [rtac ballI, rtac CollectI,
-              CONJ_WRAP' (K (EVERY' [rtac @{thm subset_UNIV}])) allAs] 1)) ss))
-        |> Thm.close_derivation
-      end;
-
-    val timer = time (timer "Coalgebra definition & thms");
-
-    (* morphism *)
-
-    val mor_bind = mk_internal_b morN;
-    val mor_name = Binding.name_of mor_bind;
-    val mor_def_bind = (Thm.def_binding mor_bind, []);
-
-    (*fbetw) forall i = 1 ... n: (\<forall>x \<in> Bi. fi x \<in> B'i)*)
-    (*mor) forall i = 1 ... n: (\<forall>x \<in> Bi.
-       Fi_map id ... id f1 ... fn (si x) = si' (fi x)*)
-    val mor_spec =
-      let
-        val morT = Library.foldr (op -->) (BTs @ sTs @ B'Ts @ s'Ts @ fTs, HOLogic.boolT);
-
-        fun mk_fbetw f B1 B2 z z' =
-          mk_Ball B1 (Term.absfree z' (HOLogic.mk_mem (f $ z, B2)));
-        fun mk_mor B mapAsBs f s s' z z' =
-          mk_Ball B (Term.absfree z' (HOLogic.mk_eq
-            (Term.list_comb (mapAsBs, passive_ids @ fs @ [s $ z]), s' $ (f $ z))));
-        val lhs = Term.list_comb (Free (mor_name, morT), Bs @ ss @ B's @ s's @ fs);
-        val rhs = HOLogic.mk_conj
-          (Library.foldr1 HOLogic.mk_conj (map5 mk_fbetw fs Bs B's zs zs'),
-           Library.foldr1 HOLogic.mk_conj (map7 mk_mor Bs mapsAsBs fs ss s's zs zs'))
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((mor_free, (_, mor_def_free)), (lthy, lthy_old)) =
-      lthy
-      |> Specification.definition (SOME (mor_bind, NONE, NoSyn), (mor_def_bind, mor_spec))
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-    val mor = fst (Term.dest_Const (Morphism.term phi mor_free));
-    val mor_def = Morphism.thm phi mor_def_free;
-
-    fun mk_mor Bs1 ss1 Bs2 ss2 fs =
-      let
-        val args = Bs1 @ ss1 @ Bs2 @ ss2 @ fs;
-        val Ts = map fastype_of (Bs1 @ ss1 @ Bs2 @ ss2 @ fs);
-        val morT = Library.foldr (op -->) (Ts, HOLogic.boolT);
-      in
-        Term.list_comb (Const (mor, morT), args)
-      end;
-
-    val mor_prem = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs);
-
-    val (mor_image_thms, morE_thms) =
-      let
-        val prem = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs);
-        fun mk_image_goal f B1 B2 = fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs)
-          (Logic.mk_implies (prem, HOLogic.mk_Trueprop (mk_leq (mk_image f $ B1) B2)));
-        val image_goals = map3 mk_image_goal fs Bs B's;
-        fun mk_elim_goal B mapAsBs f s s' x =
-          fold_rev Logic.all (x :: Bs @ ss @ B's @ s's @ fs)
-            (Logic.list_implies ([prem, HOLogic.mk_Trueprop (HOLogic.mk_mem (x, B))],
-              mk_Trueprop_eq (Term.list_comb (mapAsBs, passive_ids @ fs @ [s $ x]), s' $ (f $ x))));
-        val elim_goals = map6 mk_elim_goal Bs mapsAsBs fs ss s's zs;
-        fun prove goal =
-          Goal.prove_sorry lthy [] [] goal (K (mk_mor_elim_tac mor_def))
-          |> Thm.close_derivation;
-      in
-        (map prove image_goals, map prove elim_goals)
-      end;
-
-    val mor_image'_thms = map (fn thm => @{thm set_mp} OF [thm, imageI]) mor_image_thms;
-
-    val mor_incl_thm =
-      let
-        val prems = map2 (HOLogic.mk_Trueprop oo mk_leq) Bs Bs_copy;
-        val concl = HOLogic.mk_Trueprop (mk_mor Bs ss Bs_copy ss active_ids);
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (Bs @ ss @ Bs_copy) (Logic.list_implies (prems, concl)))
-          (K (mk_mor_incl_tac mor_def map_ids))
-        |> Thm.close_derivation
-      end;
-
-    val mor_id_thm = mor_incl_thm OF (replicate n subset_refl);
-
-    val mor_comp_thm =
-      let
-        val prems =
-          [HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs),
-           HOLogic.mk_Trueprop (mk_mor B's s's B''s s''s gs)];
-        val concl =
-          HOLogic.mk_Trueprop (mk_mor Bs ss B''s s''s (map2 (curry HOLogic.mk_comp) gs fs));
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (Bs @ ss @ B's @ s's @ B''s @ s''s @ fs @ gs)
-            (Logic.list_implies (prems, concl)))
-          (K (mk_mor_comp_tac mor_def mor_image'_thms morE_thms map_comp_id_thms))
-        |> Thm.close_derivation
-      end;
-
-    val mor_cong_thm =
-      let
-        val prems = map HOLogic.mk_Trueprop
-         (map2 (curry HOLogic.mk_eq) fs_copy fs @ [mk_mor Bs ss B's s's fs])
-        val concl = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs_copy);
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs @ fs_copy)
-            (Logic.list_implies (prems, concl)))
-          (K ((hyp_subst_tac lthy THEN' atac) 1))
-        |> Thm.close_derivation
-      end;
-
-    val mor_UNIV_thm =
-      let
-        fun mk_conjunct mapAsBs f s s' = HOLogic.mk_eq
-            (HOLogic.mk_comp (Term.list_comb (mapAsBs, passive_ids @ fs), s),
-            HOLogic.mk_comp (s', f));
-        val lhs = mk_mor active_UNIVs ss (map HOLogic.mk_UNIV activeBs) s's fs;
-        val rhs = Library.foldr1 HOLogic.mk_conj (map4 mk_conjunct mapsAsBs fs ss s's);
-      in
-        Goal.prove_sorry lthy [] [] (fold_rev Logic.all (ss @ s's @ fs) (mk_Trueprop_eq (lhs, rhs)))
-          (K (mk_mor_UNIV_tac morE_thms mor_def))
-        |> Thm.close_derivation
-      end;
-
-    val mor_str_thm =
-      let
-        val maps = map2 (fn Ds => fn bnf => Term.list_comb
-          (mk_map_of_bnf Ds allAs (passiveAs @ FTsAs) bnf, passive_ids @ ss)) Dss bnfs;
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all ss (HOLogic.mk_Trueprop
-            (mk_mor active_UNIVs ss (map HOLogic.mk_UNIV FTsAs) maps ss)))
-          (K (mk_mor_str_tac ks mor_UNIV_thm))
-        |> Thm.close_derivation
-      end;
-
-    val mor_sum_case_thm =
-      let
-        val maps = map3 (fn s => fn sum_s => fn mapx =>
-          mk_sum_case (HOLogic.mk_comp (Term.list_comb (mapx, passive_ids @ Inls), s), sum_s))
-          s's sum_ss map_Inls;
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (s's @ sum_ss) (HOLogic.mk_Trueprop
-            (mk_mor (map HOLogic.mk_UNIV activeBs) s's sum_UNIVs maps Inls)))
-          (K (mk_mor_sum_case_tac ks mor_UNIV_thm))
-        |> Thm.close_derivation
-      end;
-
-    val timer = time (timer "Morphism definition & thms");
-
-    fun hset_rec_bind j = mk_internal_b (hset_recN ^ (if m = 1 then "" else string_of_int j));
-    val hset_rec_name = Binding.name_of o hset_rec_bind;
-    val hset_rec_def_bind = rpair [] o Thm.def_binding o hset_rec_bind;
-
-    fun hset_rec_spec j Zero hsetT hrec hrec' =
-      let
-        fun mk_Suc s setsAs z z' =
-          let
-            val (set, sets) = apfst (fn xs => nth xs (j - 1)) (chop m setsAs);
-            fun mk_UN set k = mk_UNION (set $ (s $ z)) (mk_nthN n hrec k);
-          in
-            Term.absfree z'
-              (mk_union (set $ (s $ z), Library.foldl1 mk_union (map2 mk_UN sets ks)))
-          end;
-
-        val Suc = Term.absdummy HOLogic.natT (Term.absfree hrec'
-          (HOLogic.mk_tuple (map4 mk_Suc ss setssAs zs zs')));
-
-        val lhs = Term.list_comb (Free (hset_rec_name j, hsetT), ss);
-        val rhs = mk_nat_rec Zero Suc;
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((hset_rec_frees, (_, hset_rec_def_frees)), (lthy, lthy_old)) =
-      lthy
-      |> fold_map5 (fn j => fn Zero => fn hsetT => fn hrec => fn hrec' => Specification.definition
-        (SOME (hset_rec_bind j, NONE, NoSyn),
-          (hset_rec_def_bind j, hset_rec_spec j Zero hsetT hrec hrec')))
-        ls Zeros hsetTs hrecs hrecs'
-      |>> apsnd split_list o split_list
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-
-    val hset_rec_defs = map (Morphism.thm phi) hset_rec_def_frees;
-    val hset_recs = map (fst o Term.dest_Const o Morphism.term phi) hset_rec_frees;
-
-    fun mk_hset_rec ss nat i j T =
-      let
-        val args = ss @ [nat];
-        val Ts = map fastype_of ss;
-        val bTs = map domain_type Ts;
-        val hrecT = HOLogic.mk_tupleT (map (fn U => U --> HOLogic.mk_setT T) bTs)
-        val hset_recT = Library.foldr (op -->) (Ts, HOLogic.natT --> hrecT);
-      in
-        mk_nthN n (Term.list_comb (Const (nth hset_recs (j - 1), hset_recT), args)) i
-      end;
-
-    val hset_rec_0ss = mk_rec_simps n @{thm nat_rec_0} hset_rec_defs;
-    val hset_rec_Sucss = mk_rec_simps n @{thm nat_rec_Suc} hset_rec_defs;
-    val hset_rec_0ss' = transpose hset_rec_0ss;
-    val hset_rec_Sucss' = transpose hset_rec_Sucss;
-
-    fun hset_binds j = mk_internal_bs (hsetN ^ (if m = 1 then "" else string_of_int j))
-    fun hset_bind i j = nth (hset_binds j) (i - 1);
-    val hset_name = Binding.name_of oo hset_bind;
-    val hset_def_bind = rpair [] o Thm.def_binding oo hset_bind;
-
-    fun hset_spec i j =
-      let
-        val U = nth activeAs (i - 1);
-        val z = nth zs (i - 1);
-        val T = nth passiveAs (j - 1);
-        val setT = HOLogic.mk_setT T;
-        val hsetT = Library.foldr (op -->) (sTs, U --> setT);
-
-        val lhs = Term.list_comb (Free (hset_name i j, hsetT), ss @ [z]);
-        val rhs = mk_UNION (HOLogic.mk_UNIV HOLogic.natT)
-          (Term.absfree nat' (mk_hset_rec ss nat i j T $ z));
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((hset_frees, (_, hset_def_frees)), (lthy, lthy_old)) =
-      lthy
-      |> fold_map (fn i => fold_map (fn j => Specification.definition
-        (SOME (hset_bind i j, NONE, NoSyn), (hset_def_bind i j, hset_spec i j))) ls) ks
-      |>> map (apsnd split_list o split_list)
-      |>> apsnd split_list o split_list
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-
-    val hset_defss = map (map (Morphism.thm phi)) hset_def_frees;
-    val hset_defss' = transpose hset_defss;
-    val hset_namess = map (map (fst o Term.dest_Const o Morphism.term phi)) hset_frees;
-
-    fun mk_hset ss i j T =
-      let
-        val Ts = map fastype_of ss;
-        val bTs = map domain_type Ts;
-        val hsetT = Library.foldr (op -->) (Ts, nth bTs (i - 1) --> HOLogic.mk_setT T);
-      in
-        Term.list_comb (Const (nth (nth hset_namess (i - 1)) (j - 1), hsetT), ss)
-      end;
-
-    val hsetssAs = map (fn i => map2 (mk_hset ss i) ls passiveAs) ks;
-
-    val (set_incl_hset_thmss, set_hset_incl_hset_thmsss) =
-      let
-        fun mk_set_incl_hset s x set hset = fold_rev Logic.all (x :: ss)
-          (HOLogic.mk_Trueprop (mk_leq (set $ (s $ x)) (hset $ x)));
-
-        fun mk_set_hset_incl_hset s x y set hset1 hset2 =
-          fold_rev Logic.all (x :: y :: ss)
-            (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (x, set $ (s $ y))),
-            HOLogic.mk_Trueprop (mk_leq (hset1 $ x) (hset2 $ y))));
-
-        val set_incl_hset_goalss =
-          map4 (fn s => fn x => fn sets => fn hsets =>
-            map2 (mk_set_incl_hset s x) (take m sets) hsets)
-          ss zs setssAs hsetssAs;
-
-        (*xk : F(i)set(m+k) (si yi) ==> F(k)_hset(j) s1 ... sn xk <= F(i)_hset(j) s1 ... sn yi*)
-        val set_hset_incl_hset_goalsss =
-          map4 (fn si => fn yi => fn sets => fn hsetsi =>
-            map3 (fn xk => fn set => fn hsetsk =>
-              map2 (mk_set_hset_incl_hset si xk yi set) hsetsk hsetsi)
-            zs_copy (drop m sets) hsetssAs)
-          ss zs setssAs hsetssAs;
-      in
-        (map3 (fn goals => fn defs => fn rec_Sucs =>
-          map3 (fn goal => fn def => fn rec_Suc =>
-            Goal.prove_sorry lthy [] [] goal (K (mk_set_incl_hset_tac def rec_Suc))
-            |> Thm.close_derivation)
-          goals defs rec_Sucs)
-        set_incl_hset_goalss hset_defss hset_rec_Sucss,
-        map3 (fn goalss => fn defsi => fn rec_Sucs =>
-          map3 (fn k => fn goals => fn defsk =>
-            map4 (fn goal => fn defk => fn defi => fn rec_Suc =>
-              Goal.prove_sorry lthy [] [] goal
-                (K (mk_set_hset_incl_hset_tac n [defk, defi] rec_Suc k))
-              |> Thm.close_derivation)
-            goals defsk defsi rec_Sucs)
-          ks goalss hset_defss)
-        set_hset_incl_hset_goalsss hset_defss hset_rec_Sucss)
-      end;
-
-    val set_incl_hset_thmss' = transpose set_incl_hset_thmss;
-    val set_hset_incl_hset_thmsss' = transpose (map transpose set_hset_incl_hset_thmsss);
-    val set_hset_thmss = map (map (fn thm => thm RS @{thm set_mp})) set_incl_hset_thmss;
-    val set_hset_hset_thmsss = map (map (map (fn thm => thm RS @{thm set_mp})))
-      set_hset_incl_hset_thmsss;
-    val set_hset_thmss' = transpose set_hset_thmss;
-    val set_hset_hset_thmsss' = transpose (map transpose set_hset_hset_thmsss);
-
-    val hset_minimal_thms =
-      let
-        fun mk_passive_prem set s x K =
-          Logic.all x (HOLogic.mk_Trueprop (mk_leq (set $ (s $ x)) (K $ x)));
-
-        fun mk_active_prem s x1 K1 set x2 K2 =
-          fold_rev Logic.all [x1, x2]
-            (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (x2, set $ (s $ x1))),
-              HOLogic.mk_Trueprop (mk_leq (K2 $ x2) (K1 $ x1))));
-
-        val premss = map2 (fn j => fn Ks =>
-          map4 mk_passive_prem (map (fn xs => nth xs (j - 1)) setssAs) ss zs Ks @
-            flat (map4 (fn sets => fn s => fn x1 => fn K1 =>
-              map3 (mk_active_prem s x1 K1) (drop m sets) zs_copy Ks) setssAs ss zs Ks))
-          ls Kss;
-
-        val hset_rec_minimal_thms =
-          let
-            fun mk_conjunct j T i K x = mk_leq (mk_hset_rec ss nat i j T $ x) (K $ x);
-            fun mk_concl j T Ks = list_all_free zs
-              (Library.foldr1 HOLogic.mk_conj (map3 (mk_conjunct j T) ks Ks zs));
-            val concls = map3 mk_concl ls passiveAs Kss;
-
-            val goals = map2 (fn prems => fn concl =>
-              Logic.list_implies (prems, HOLogic.mk_Trueprop concl)) premss concls
-
-            val ctss =
-              map (fn phi => map (SOME o certify lthy) [Term.absfree nat' phi, nat]) concls;
-          in
-            map4 (fn goal => fn cts => fn hset_rec_0s => fn hset_rec_Sucs =>
-              singleton (Proof_Context.export names_lthy lthy)
-                (Goal.prove_sorry lthy [] [] goal
-                  (mk_hset_rec_minimal_tac m cts hset_rec_0s hset_rec_Sucs))
-              |> Thm.close_derivation)
-            goals ctss hset_rec_0ss' hset_rec_Sucss'
-          end;
-
-        fun mk_conjunct j T i K x = mk_leq (mk_hset ss i j T $ x) (K $ x);
-        fun mk_concl j T Ks = Library.foldr1 HOLogic.mk_conj (map3 (mk_conjunct j T) ks Ks zs);
-        val concls = map3 mk_concl ls passiveAs Kss;
-
-        val goals = map3 (fn Ks => fn prems => fn concl =>
-          fold_rev Logic.all (Ks @ ss @ zs)
-            (Logic.list_implies (prems, HOLogic.mk_Trueprop concl))) Kss premss concls;
-      in
-        map3 (fn goal => fn hset_defs => fn hset_rec_minimal =>
-          Goal.prove_sorry lthy [] [] goal
-            (mk_hset_minimal_tac n hset_defs hset_rec_minimal)
-          |> Thm.close_derivation)
-        goals hset_defss' hset_rec_minimal_thms
-      end;
-
-    val timer = time (timer "Hereditary sets");
-
-    (* bisimulation *)
-
-    val bis_bind = mk_internal_b bisN;
-    val bis_name = Binding.name_of bis_bind;
-    val bis_def_bind = (Thm.def_binding bis_bind, []);
-
-    fun mk_bis_le_conjunct R B1 B2 = mk_leq R (mk_Times (B1, B2));
-    val bis_le = Library.foldr1 HOLogic.mk_conj (map3 mk_bis_le_conjunct Rs Bs B's)
-
-    val bis_spec =
-      let
-        val bisT = Library.foldr (op -->) (ATs @ BTs @ sTs @ B'Ts @ s'Ts @ setRTs, HOLogic.boolT);
-
-        val fst_args = passive_ids @ fsts;
-        val snd_args = passive_ids @ snds;
-        fun mk_bis R s s' b1 b2 RF map1 map2 sets =
-          list_all_free [b1, b2] (HOLogic.mk_imp
-            (HOLogic.mk_mem (HOLogic.mk_prod (b1, b2), R),
-            mk_Bex (mk_in (As @ Rs) sets (snd (dest_Free RF))) (Term.absfree (dest_Free RF)
-              (HOLogic.mk_conj
-                (HOLogic.mk_eq (Term.list_comb (map1, fst_args) $ RF, s $ b1),
-                HOLogic.mk_eq (Term.list_comb (map2, snd_args) $ RF, s' $ b2))))));
-
-        val lhs = Term.list_comb (Free (bis_name, bisT), As @ Bs @ ss @ B's @ s's @ Rs);
-        val rhs = HOLogic.mk_conj
-          (bis_le, Library.foldr1 HOLogic.mk_conj
-            (map9 mk_bis Rs ss s's zs z's RFs map_fsts map_snds bis_setss))
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((bis_free, (_, bis_def_free)), (lthy, lthy_old)) =
-      lthy
-      |> Specification.definition (SOME (bis_bind, NONE, NoSyn), (bis_def_bind, bis_spec))
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-    val bis = fst (Term.dest_Const (Morphism.term phi bis_free));
-    val bis_def = Morphism.thm phi bis_def_free;
-
-    fun mk_bis As Bs1 ss1 Bs2 ss2 Rs =
-      let
-        val args = As @ Bs1 @ ss1 @ Bs2 @ ss2 @ Rs;
-        val Ts = map fastype_of args;
-        val bisT = Library.foldr (op -->) (Ts, HOLogic.boolT);
-      in
-        Term.list_comb (Const (bis, bisT), args)
-      end;
-
-    val bis_cong_thm =
-      let
-        val prems = map HOLogic.mk_Trueprop
-         (mk_bis As Bs ss B's s's Rs :: map2 (curry HOLogic.mk_eq) Rs_copy Rs)
-        val concl = HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's Rs_copy);
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ Rs @ Rs_copy)
-            (Logic.list_implies (prems, concl)))
-          (K ((hyp_subst_tac lthy THEN' atac) 1))
-        |> Thm.close_derivation
-      end;
-
-    val bis_rel_thm =
-      let
-        fun mk_conjunct R s s' b1 b2 rel =
-          list_all_free [b1, b2] (HOLogic.mk_imp
-            (HOLogic.mk_mem (HOLogic.mk_prod (b1, b2), R),
-            Term.list_comb (rel, map mk_in_rel (passive_Id_ons @ Rs)) $ (s $ b1) $ (s' $ b2)));
-
-        val rhs = HOLogic.mk_conj
-          (bis_le, Library.foldr1 HOLogic.mk_conj
-            (map6 mk_conjunct Rs ss s's zs z's relsAsBs))
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ Rs)
-            (mk_Trueprop_eq (mk_bis As Bs ss B's s's Rs, rhs)))
-          (K (mk_bis_rel_tac lthy m bis_def rel_OO_Grps map_comps map_cong0s set_mapss))
-        |> Thm.close_derivation
-      end;
-
-    val bis_converse_thm =
-      Goal.prove_sorry lthy [] []
-        (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ Rs)
-          (Logic.mk_implies
-            (HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's Rs),
-            HOLogic.mk_Trueprop (mk_bis As B's s's Bs ss (map mk_converse Rs)))))
-        (K (mk_bis_converse_tac m bis_rel_thm rel_congs rel_converseps))
-      |> Thm.close_derivation;
-
-    val bis_O_thm =
-      let
-        val prems =
-          [HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's Rs),
-           HOLogic.mk_Trueprop (mk_bis As B's s's B''s s''s R's)];
-        val concl =
-          HOLogic.mk_Trueprop (mk_bis As Bs ss B''s s''s (map2 (curry mk_rel_comp) Rs R's));
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ B''s @ s''s @ Rs @ R's)
-            (Logic.list_implies (prems, concl)))
-          (K (mk_bis_O_tac lthy m bis_rel_thm rel_congs rel_OOs))
-        |> Thm.close_derivation
-      end;
-
-    val bis_Gr_thm =
-      let
-        val concl =
-          HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's (map2 mk_Gr Bs fs));
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ fs)
-            (Logic.list_implies ([coalg_prem, mor_prem], concl)))
-          (mk_bis_Gr_tac bis_rel_thm rel_Grps mor_image_thms morE_thms coalg_in_thms)
-        |> Thm.close_derivation
-      end;
-
-    val bis_image2_thm = bis_cong_thm OF
-      ((bis_O_thm OF [bis_Gr_thm RS bis_converse_thm, bis_Gr_thm]) ::
-      replicate n @{thm image2_Gr});
-
-    val bis_Id_on_thm = bis_cong_thm OF ((mor_id_thm RSN (2, bis_Gr_thm)) ::
-      replicate n @{thm Id_on_Gr});
-
-    val bis_Union_thm =
-      let
-        val prem =
-          HOLogic.mk_Trueprop (mk_Ball Idx
-            (Term.absfree idx' (mk_bis As Bs ss B's s's (map (fn R => R $ idx) Ris))));
-        val concl =
-          HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's (map (mk_UNION Idx) Ris));
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (Idx :: As @ Bs @ ss @ B's @ s's @ Ris)
-            (Logic.mk_implies (prem, concl)))
-          (mk_bis_Union_tac bis_def in_mono'_thms)
-        |> Thm.close_derivation
-      end;
-
-    (* self-bisimulation *)
-
-    fun mk_sbis As Bs ss Rs = mk_bis As Bs ss Bs ss Rs;
-
-    val sbis_prem = HOLogic.mk_Trueprop (mk_sbis As Bs ss sRs);
-
-    (* largest self-bisimulation *)
-
-    val lsbis_binds = mk_internal_bs lsbisN;
-    fun lsbis_bind i = nth lsbis_binds (i - 1);
-    val lsbis_name = Binding.name_of o lsbis_bind;
-    val lsbis_def_bind = rpair [] o Thm.def_binding o lsbis_bind;
-
-    val all_sbis = HOLogic.mk_Collect (fst Rtuple', snd Rtuple', list_exists_free sRs
-      (HOLogic.mk_conj (HOLogic.mk_eq (Rtuple, HOLogic.mk_tuple sRs), mk_sbis As Bs ss sRs)));
-
-    fun lsbis_spec i RT =
-      let
-        fun mk_lsbisT RT =
-          Library.foldr (op -->) (map fastype_of (As @ Bs @ ss), RT);
-        val lhs = Term.list_comb (Free (lsbis_name i, mk_lsbisT RT), As @ Bs @ ss);
-        val rhs = mk_UNION all_sbis (Term.absfree Rtuple' (mk_nthN n Rtuple i));
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((lsbis_frees, (_, lsbis_def_frees)), (lthy, lthy_old)) =
-      lthy
-      |> fold_map2 (fn i => fn RT => Specification.definition
-        (SOME (lsbis_bind i, NONE, NoSyn), (lsbis_def_bind i, lsbis_spec i RT))) ks setsRTs
-      |>> apsnd split_list o split_list
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-
-    val lsbis_defs = map (Morphism.thm phi) lsbis_def_frees;
-    val lsbiss = map (fst o Term.dest_Const o Morphism.term phi) lsbis_frees;
-
-    fun mk_lsbis As Bs ss i =
-      let
-        val args = As @ Bs @ ss;
-        val Ts = map fastype_of args;
-        val RT = mk_relT (`I (HOLogic.dest_setT (fastype_of (nth Bs (i - 1)))));
-        val lsbisT = Library.foldr (op -->) (Ts, RT);
-      in
-        Term.list_comb (Const (nth lsbiss (i - 1), lsbisT), args)
-      end;
-
-    val sbis_lsbis_thm =
-      Goal.prove_sorry lthy [] []
-        (fold_rev Logic.all (As @ Bs @ ss)
-          (HOLogic.mk_Trueprop (mk_sbis As Bs ss (map (mk_lsbis As Bs ss) ks))))
-        (K (mk_sbis_lsbis_tac lthy lsbis_defs bis_Union_thm bis_cong_thm))
-      |> Thm.close_derivation;
-
-    val lsbis_incl_thms = map (fn i => sbis_lsbis_thm RS
-      (bis_def RS iffD1 RS conjunct1 RS mk_conjunctN n i)) ks;
-    val lsbisE_thms = map (fn i => (mk_specN 2 (sbis_lsbis_thm RS
-      (bis_def RS iffD1 RS conjunct2 RS mk_conjunctN n i))) RS mp) ks;
-
-    val incl_lsbis_thms =
-      let
-        fun mk_concl i R = HOLogic.mk_Trueprop (mk_leq R (mk_lsbis As Bs ss i));
-        val goals = map2 (fn i => fn R => fold_rev Logic.all (As @ Bs @ ss @ sRs)
-          (Logic.mk_implies (sbis_prem, mk_concl i R))) ks sRs;
-      in
-        map3 (fn goal => fn i => fn def => Goal.prove_sorry lthy [] [] goal
-          (K (mk_incl_lsbis_tac n i def)) |> Thm.close_derivation) goals ks lsbis_defs
-      end;
-
-    val equiv_lsbis_thms =
-      let
-        fun mk_concl i B = HOLogic.mk_Trueprop (mk_equiv B (mk_lsbis As Bs ss i));
-        val goals = map2 (fn i => fn B => fold_rev Logic.all (As @ Bs @ ss)
-          (Logic.mk_implies (coalg_prem, mk_concl i B))) ks Bs;
-      in
-        map3 (fn goal => fn l_incl => fn incl_l =>
-          Goal.prove_sorry lthy [] [] goal
-            (K (mk_equiv_lsbis_tac sbis_lsbis_thm l_incl incl_l
-              bis_Id_on_thm bis_converse_thm bis_O_thm))
-          |> Thm.close_derivation)
-        goals lsbis_incl_thms incl_lsbis_thms
-      end;
-
-    val timer = time (timer "Bisimulations");
-
-    (* bounds *)
-
-    val (lthy, sbd, sbdT,
-      sbd_card_order, sbd_Cinfinite, sbd_Card_order, set_sbdss) =
-      if n = 1
-      then (lthy, sum_bd, sum_bdT, bd_card_order, bd_Cinfinite, bd_Card_order, set_bdss)
-      else
-        let
-          val sbdT_bind = mk_internal_b sum_bdTN;
-
-          val ((sbdT_name, (sbdT_glob_info, sbdT_loc_info)), lthy) =
-            typedef (sbdT_bind, dead_params, NoSyn)
-              (HOLogic.mk_UNIV sum_bdT) NONE (EVERY' [rtac exI, rtac UNIV_I] 1) lthy;
-
-          val sbdT = Type (sbdT_name, dead_params');
-          val Abs_sbdT = Const (#Abs_name sbdT_glob_info, sum_bdT --> sbdT);
-
-          val sbd_bind = mk_internal_b sum_bdN;
-          val sbd_name = Binding.name_of sbd_bind;
-          val sbd_def_bind = (Thm.def_binding sbd_bind, []);
-
-          val sbd_spec = HOLogic.mk_Trueprop
-            (HOLogic.mk_eq (Free (sbd_name, mk_relT (`I sbdT)), mk_dir_image sum_bd Abs_sbdT));
-
-          val ((sbd_free, (_, sbd_def_free)), (lthy, lthy_old)) =
-            lthy
-            |> Specification.definition (SOME (sbd_bind, NONE, NoSyn), (sbd_def_bind, sbd_spec))
-            ||> `Local_Theory.restore;
-
-          val phi = Proof_Context.export_morphism lthy_old lthy;
-
-          val sbd_def = Morphism.thm phi sbd_def_free;
-          val sbd = Const (fst (Term.dest_Const (Morphism.term phi sbd_free)), mk_relT (`I sbdT));
-
-          val Abs_sbdT_inj = mk_Abs_inj_thm (#Abs_inject sbdT_loc_info);
-          val Abs_sbdT_bij = mk_Abs_bij_thm lthy Abs_sbdT_inj (#Abs_cases sbdT_loc_info);
-
-          fun mk_sum_Cinfinite [thm] = thm
-            | mk_sum_Cinfinite (thm :: thms) =
-              @{thm Cinfinite_csum_strong} OF [thm, mk_sum_Cinfinite thms];
-
-          val sum_Cinfinite = mk_sum_Cinfinite bd_Cinfinites;
-          val sum_Card_order = sum_Cinfinite RS conjunct2;
-
-          fun mk_sum_card_order [thm] = thm
-            | mk_sum_card_order (thm :: thms) =
-              @{thm card_order_csum} OF [thm, mk_sum_card_order thms];
-
-          val sum_card_order = mk_sum_card_order bd_card_orders;
-
-          val sbd_ordIso = fold_thms lthy [sbd_def]
-            (@{thm dir_image} OF [Abs_sbdT_inj, sum_Card_order]);
-          val sbd_card_order =  fold_thms lthy [sbd_def]
-            (@{thm card_order_dir_image} OF [Abs_sbdT_bij, sum_card_order]);
-          val sbd_Cinfinite = @{thm Cinfinite_cong} OF [sbd_ordIso, sum_Cinfinite];
-          val sbd_Card_order = sbd_Cinfinite RS conjunct2;
-
-          fun mk_set_sbd i bd_Card_order bds =
-            map (fn thm => @{thm ordLeq_ordIso_trans} OF
-              [bd_Card_order RS mk_ordLeq_csum n i thm, sbd_ordIso]) bds;
-          val set_sbdss = map3 mk_set_sbd ks bd_Card_orders set_bdss;
-       in
-         (lthy, sbd, sbdT, sbd_card_order, sbd_Cinfinite, sbd_Card_order, set_sbdss)
-       end;
-
-    val sbdTs = replicate n sbdT;
-    val sum_sbd = Library.foldr1 (uncurry mk_csum) (replicate n sbd);
-    val sum_sbdT = mk_sumTN sbdTs;
-    val sum_sbd_listT = HOLogic.listT sum_sbdT;
-    val sum_sbd_list_setT = HOLogic.mk_setT sum_sbd_listT;
-    val bdTs = passiveAs @ replicate n sbdT;
-    val to_sbd_maps = map4 mk_map_of_bnf Dss Ass (replicate n bdTs) bnfs;
-    val bdFTs = mk_FTs bdTs;
-    val sbdFT = mk_sumTN bdFTs;
-    val treeT = HOLogic.mk_prodT (sum_sbd_list_setT, sum_sbd_listT --> sbdFT);
-    val treeQT = HOLogic.mk_setT treeT;
-    val treeTs = passiveAs @ replicate n treeT;
-    val treeQTs = passiveAs @ replicate n treeQT;
-    val treeFTs = mk_FTs treeTs;
-    val tree_maps = map4 mk_map_of_bnf Dss (replicate n bdTs) (replicate n treeTs) bnfs;
-    val final_maps = map4 mk_map_of_bnf Dss (replicate n treeTs) (replicate n treeQTs) bnfs;
-    val isNode_setss = mk_setss (passiveAs @ replicate n sbdT);
-
-    val root = HOLogic.mk_set sum_sbd_listT [HOLogic.mk_list sum_sbdT []];
-    val Zero = HOLogic.mk_tuple (map (fn U => absdummy U root) activeAs);
-    val Lev_recT = fastype_of Zero;
-    val LevT = Library.foldr (op -->) (sTs, HOLogic.natT --> Lev_recT);
-
-    val Nil = HOLogic.mk_tuple (map3 (fn i => fn z => fn z'=>
-      Term.absfree z' (mk_InN activeAs z i)) ks zs zs');
-    val rv_recT = fastype_of Nil;
-    val rvT = Library.foldr (op -->) (sTs, sum_sbd_listT --> rv_recT);
-
-    val (((((((((((sumx, sumx'), (kks, kks')), (kl, kl')), (kl_copy, kl'_copy)), (Kl, Kl')),
-      (lab, lab')), (Kl_lab, Kl_lab')), xs), (Lev_rec, Lev_rec')), (rv_rec, rv_rec')),
-      names_lthy) = names_lthy
-      |> yield_singleton (apfst (op ~~) oo mk_Frees' "sumx") sum_sbdT
-      ||>> mk_Frees' "k" sbdTs
-      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "kl") sum_sbd_listT
-      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "kl") sum_sbd_listT
-      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "Kl") sum_sbd_list_setT
-      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "lab") (sum_sbd_listT --> sbdFT)
-      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "Kl_lab") treeT
-      ||>> mk_Frees "x" bdFTs
-      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "rec") Lev_recT
-      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "rec") rv_recT;
-
-    val (k, k') = (hd kks, hd kks')
-
-    val timer = time (timer "Bounds");
-
-    (* tree coalgebra *)
-
-    val isNode_binds = mk_internal_bs isNodeN;
-    fun isNode_bind i = nth isNode_binds (i - 1);
-    val isNode_name = Binding.name_of o isNode_bind;
-    val isNode_def_bind = rpair [] o Thm.def_binding o isNode_bind;
-
-    val isNodeT =
-      Library.foldr (op -->) (map fastype_of (As @ [Kl, lab, kl]), HOLogic.boolT);
-
-    val Succs = map3 (fn i => fn k => fn k' =>
-      HOLogic.mk_Collect (fst k', snd k', HOLogic.mk_mem (mk_InN sbdTs k i, mk_Succ Kl kl)))
-      ks kks kks';
-
-    fun isNode_spec sets x i =
-      let
-        val (passive_sets, active_sets) = chop m (map (fn set => set $ x) sets);
-        val lhs = Term.list_comb (Free (isNode_name i, isNodeT), As @ [Kl, lab, kl]);
-        val rhs = list_exists_free [x]
-          (Library.foldr1 HOLogic.mk_conj (HOLogic.mk_eq (lab $ kl, mk_InN bdFTs x i) ::
-          map2 mk_leq passive_sets As @ map2 (curry HOLogic.mk_eq) active_sets Succs));
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((isNode_frees, (_, isNode_def_frees)), (lthy, lthy_old)) =
-      lthy
-      |> fold_map3 (fn i => fn x => fn sets => Specification.definition
-        (SOME (isNode_bind i, NONE, NoSyn), (isNode_def_bind i, isNode_spec sets x i)))
-        ks xs isNode_setss
-      |>> apsnd split_list o split_list
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-
-    val isNode_defs = map (Morphism.thm phi) isNode_def_frees;
-    val isNodes = map (fst o Term.dest_Const o Morphism.term phi) isNode_frees;
-
-    fun mk_isNode As kl i =
-      Term.list_comb (Const (nth isNodes (i - 1), isNodeT), As @ [Kl, lab, kl]);
-
-    val isTree =
-      let
-        val empty = HOLogic.mk_mem (HOLogic.mk_list sum_sbdT [], Kl);
-        val Field = mk_leq Kl (mk_Field (mk_clists sum_sbd));
-        val prefCl = mk_prefCl Kl;
-
-        val tree = mk_Ball Kl (Term.absfree kl'
-          (HOLogic.mk_conj
-            (Library.foldr1 HOLogic.mk_disj (map (mk_isNode As kl) ks),
-            Library.foldr1 HOLogic.mk_conj (map4 (fn Succ => fn i => fn k => fn k' =>
-              mk_Ball Succ (Term.absfree k' (mk_isNode As
-                (mk_append (kl, HOLogic.mk_list sum_sbdT [mk_InN sbdTs k i])) i)))
-            Succs ks kks kks'))));
-
-        val undef = list_all_free [kl] (HOLogic.mk_imp
-          (HOLogic.mk_not (HOLogic.mk_mem (kl, Kl)),
-          HOLogic.mk_eq (lab $ kl, mk_undefined sbdFT)));
-      in
-        Library.foldr1 HOLogic.mk_conj [empty, Field, prefCl, tree, undef]
-      end;
-
-    val carT_binds = mk_internal_bs carTN;
-    fun carT_bind i = nth carT_binds (i - 1);
-    val carT_name = Binding.name_of o carT_bind;
-    val carT_def_bind = rpair [] o Thm.def_binding o carT_bind;
-
-    fun carT_spec i =
-      let
-        val carTT = Library.foldr (op -->) (ATs, HOLogic.mk_setT treeT);
-
-        val lhs = Term.list_comb (Free (carT_name i, carTT), As);
-        val rhs = HOLogic.mk_Collect (fst Kl_lab', snd Kl_lab', list_exists_free [Kl, lab]
-          (HOLogic.mk_conj (HOLogic.mk_eq (Kl_lab, HOLogic.mk_prod (Kl, lab)),
-            HOLogic.mk_conj (isTree, mk_isNode As (HOLogic.mk_list sum_sbdT []) i))));
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((carT_frees, (_, carT_def_frees)), (lthy, lthy_old)) =
-      lthy
-      |> fold_map (fn i => Specification.definition
-        (SOME (carT_bind i, NONE, NoSyn), (carT_def_bind i, carT_spec i))) ks
-      |>> apsnd split_list o split_list
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-
-    val carT_defs = map (Morphism.thm phi) carT_def_frees;
-    val carTs = map (fst o Term.dest_Const o Morphism.term phi) carT_frees;
-
-    fun mk_carT As i = Term.list_comb
-      (Const (nth carTs (i - 1),
-         Library.foldr (op -->) (map fastype_of As, HOLogic.mk_setT treeT)), As);
-
-    val strT_binds = mk_internal_bs strTN;
-    fun strT_bind i = nth strT_binds (i - 1);
-    val strT_name = Binding.name_of o strT_bind;
-    val strT_def_bind = rpair [] o Thm.def_binding o strT_bind;
-
-    fun strT_spec mapFT FT i =
-      let
-        val strTT = treeT --> FT;
-
-        fun mk_f i k k' =
-          let val in_k = mk_InN sbdTs k i;
-          in Term.absfree k' (HOLogic.mk_prod (mk_Shift Kl in_k, mk_shift lab in_k)) end;
-
-        val f = Term.list_comb (mapFT, passive_ids @ map3 mk_f ks kks kks');
-        val (fTs1, fTs2) = apsnd tl (chop (i - 1) (map (fn T => T --> FT) bdFTs));
-        val fs = map mk_undefined fTs1 @ (f :: map mk_undefined fTs2);
-        val lhs = Free (strT_name i, strTT);
-        val rhs = HOLogic.mk_split (Term.absfree Kl' (Term.absfree lab'
-          (mk_sum_caseN fs $ (lab $ HOLogic.mk_list sum_sbdT []))));
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((strT_frees, (_, strT_def_frees)), (lthy, lthy_old)) =
-      lthy
-      |> fold_map3 (fn i => fn mapFT => fn FT => Specification.definition
-        (SOME (strT_bind i, NONE, NoSyn), (strT_def_bind i, strT_spec mapFT FT i)))
-        ks tree_maps treeFTs
-      |>> apsnd split_list o split_list
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-
-    val strT_defs = map ((fn def => trans OF [def RS fun_cong, @{thm prod.cases}]) o
-      Morphism.thm phi) strT_def_frees;
-    val strTs = map (fst o Term.dest_Const o Morphism.term phi) strT_frees;
-
-    fun mk_strT FT i = Const (nth strTs (i - 1), treeT --> FT);
-
-    val carTAs = map (mk_carT As) ks;
-    val strTAs = map2 mk_strT treeFTs ks;
-
-    val coalgT_thm =
-      Goal.prove_sorry lthy [] []
-        (fold_rev Logic.all As (HOLogic.mk_Trueprop (mk_coalg As carTAs strTAs)))
-        (mk_coalgT_tac m (coalg_def :: isNode_defs @ carT_defs) strT_defs set_mapss)
-      |> Thm.close_derivation;
-
-    val timer = time (timer "Tree coalgebra");
-
-    fun mk_to_sbd s x i i' =
-      mk_toCard (nth (nth setssAs (i - 1)) (m + i' - 1) $ (s $ x)) sbd;
-    fun mk_from_sbd s x i i' =
-      mk_fromCard (nth (nth setssAs (i - 1)) (m + i' - 1) $ (s $ x)) sbd;
-
-    fun mk_to_sbd_thmss thm = map (map (fn set_sbd =>
-      thm OF [set_sbd, sbd_Card_order]) o drop m) set_sbdss;
-
-    val to_sbd_inj_thmss = mk_to_sbd_thmss @{thm toCard_inj};
-    val to_sbd_thmss = mk_to_sbd_thmss @{thm toCard};
-    val from_to_sbd_thmss = mk_to_sbd_thmss @{thm fromCard_toCard};
-
-    val Lev_bind = mk_internal_b LevN;
-    val Lev_name = Binding.name_of Lev_bind;
-    val Lev_def_bind = rpair [] (Thm.def_binding Lev_bind);
-
-    val Lev_spec =
-      let
-        fun mk_Suc i s setsAs a a' =
-          let
-            val sets = drop m setsAs;
-            fun mk_set i' set b =
-              let
-                val Cons = HOLogic.mk_eq (kl_copy,
-                  mk_Cons (mk_InN sbdTs (mk_to_sbd s a i i' $ b) i') kl)
-                val b_set = HOLogic.mk_mem (b, set $ (s $ a));
-                val kl_rec = HOLogic.mk_mem (kl, mk_nthN n Lev_rec i' $ b);
-              in
-                HOLogic.mk_Collect (fst kl'_copy, snd kl'_copy, list_exists_free [b, kl]
-                  (HOLogic.mk_conj (Cons, HOLogic.mk_conj (b_set, kl_rec))))
-              end;
-          in
-            Term.absfree a' (Library.foldl1 mk_union (map3 mk_set ks sets zs_copy))
-          end;
-
-        val Suc = Term.absdummy HOLogic.natT (Term.absfree Lev_rec'
-          (HOLogic.mk_tuple (map5 mk_Suc ks ss setssAs zs zs')));
-
-        val lhs = Term.list_comb (Free (Lev_name, LevT), ss);
-        val rhs = mk_nat_rec Zero Suc;
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((Lev_free, (_, Lev_def_free)), (lthy, lthy_old)) =
-      lthy
-      |> Specification.definition (SOME (Lev_bind, NONE, NoSyn), (Lev_def_bind, Lev_spec))
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-
-    val Lev_def = Morphism.thm phi Lev_def_free;
-    val Lev = fst (Term.dest_Const (Morphism.term phi Lev_free));
-
-    fun mk_Lev ss nat i =
-      let
-        val Ts = map fastype_of ss;
-        val LevT = Library.foldr (op -->) (Ts, HOLogic.natT -->
-          HOLogic.mk_tupleT (map (fn U => domain_type U --> sum_sbd_list_setT) Ts));
-      in
-        mk_nthN n (Term.list_comb (Const (Lev, LevT), ss) $ nat) i
-      end;
-
-    val Lev_0s = flat (mk_rec_simps n @{thm nat_rec_0} [Lev_def]);
-    val Lev_Sucs = flat (mk_rec_simps n @{thm nat_rec_Suc} [Lev_def]);
-
-    val rv_bind = mk_internal_b rvN;
-    val rv_name = Binding.name_of rv_bind;
-    val rv_def_bind = rpair [] (Thm.def_binding rv_bind);
-
-    val rv_spec =
-      let
-        fun mk_Cons i s b b' =
-          let
-            fun mk_case i' =
-              Term.absfree k' (mk_nthN n rv_rec i' $ (mk_from_sbd s b i i' $ k));
-          in
-            Term.absfree b' (mk_sum_caseN (map mk_case ks) $ sumx)
-          end;
-
-        val Cons = Term.absfree sumx' (Term.absdummy sum_sbd_listT (Term.absfree rv_rec'
-          (HOLogic.mk_tuple (map4 mk_Cons ks ss zs zs'))));
-
-        val lhs = Term.list_comb (Free (rv_name, rvT), ss);
-        val rhs = mk_list_rec Nil Cons;
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((rv_free, (_, rv_def_free)), (lthy, lthy_old)) =
-      lthy
-      |> Specification.definition (SOME (rv_bind, NONE, NoSyn), (rv_def_bind, rv_spec))
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-
-    val rv_def = Morphism.thm phi rv_def_free;
-    val rv = fst (Term.dest_Const (Morphism.term phi rv_free));
-
-    fun mk_rv ss kl i =
-      let
-        val Ts = map fastype_of ss;
-        val As = map domain_type Ts;
-        val rvT = Library.foldr (op -->) (Ts, fastype_of kl -->
-          HOLogic.mk_tupleT (map (fn U => U --> mk_sumTN As) As));
-      in
-        mk_nthN n (Term.list_comb (Const (rv, rvT), ss) $ kl) i
-      end;
-
-    val rv_Nils = flat (mk_rec_simps n @{thm list_rec_Nil} [rv_def]);
-    val rv_Conss = flat (mk_rec_simps n @{thm list_rec_Cons} [rv_def]);
-
-    val beh_binds = mk_internal_bs behN;
-    fun beh_bind i = nth beh_binds (i - 1);
-    val beh_name = Binding.name_of o beh_bind;
-    val beh_def_bind = rpair [] o Thm.def_binding o beh_bind;
-
-    fun beh_spec i z =
-      let
-        val mk_behT = Library.foldr (op -->) (map fastype_of (ss @ [z]), treeT);
-
-        fun mk_case i to_sbd_map s k k' =
-          Term.absfree k' (mk_InN bdFTs
-            (Term.list_comb (to_sbd_map, passive_ids @ map (mk_to_sbd s k i) ks) $ (s $ k)) i);
-
-        val Lab = Term.absfree kl' (mk_If
-          (HOLogic.mk_mem (kl, mk_Lev ss (mk_size kl) i $ z))
-          (mk_sum_caseN (map5 mk_case ks to_sbd_maps ss zs zs') $ (mk_rv ss kl i $ z))
-          (mk_undefined sbdFT));
-
-        val lhs = Term.list_comb (Free (beh_name i, mk_behT), ss) $ z;
-        val rhs = HOLogic.mk_prod (mk_UNION (HOLogic.mk_UNIV HOLogic.natT)
-          (Term.absfree nat' (mk_Lev ss nat i $ z)), Lab);
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((beh_frees, (_, beh_def_frees)), (lthy, lthy_old)) =
-      lthy
-      |> fold_map2 (fn i => fn z => Specification.definition
-        (SOME (beh_bind i, NONE, NoSyn), (beh_def_bind i, beh_spec i z))) ks zs
-      |>> apsnd split_list o split_list
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-
-    val beh_defs = map (Morphism.thm phi) beh_def_frees;
-    val behs = map (fst o Term.dest_Const o Morphism.term phi) beh_frees;
-
-    fun mk_beh ss i =
-      let
-        val Ts = map fastype_of ss;
-        val behT = Library.foldr (op -->) (Ts, nth activeAs (i - 1) --> treeT);
-      in
-        Term.list_comb (Const (nth behs (i - 1), behT), ss)
-      end;
-
-    val Lev_sbd_thms =
-      let
-        fun mk_conjunct i z = mk_leq (mk_Lev ss nat i $ z) (mk_Field (mk_clists sum_sbd));
-        val goal = list_all_free zs
-          (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs));
-
-        val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
-
-        val Lev_sbd = singleton (Proof_Context.export names_lthy lthy)
-          (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
-            (K (mk_Lev_sbd_tac lthy cts Lev_0s Lev_Sucs to_sbd_thmss))
-          |> Thm.close_derivation);
-
-        val Lev_sbd' = mk_specN n Lev_sbd;
-      in
-        map (fn i => Lev_sbd' RS mk_conjunctN n i) ks
-      end;
-
-    val (length_Lev_thms, length_Lev'_thms) =
-      let
-        fun mk_conjunct i z = HOLogic.mk_imp (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z),
-          HOLogic.mk_eq (mk_size kl, nat));
-        val goal = list_all_free (kl :: zs)
-          (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs));
-
-        val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
-
-        val length_Lev = singleton (Proof_Context.export names_lthy lthy)
-          (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
-            (K (mk_length_Lev_tac lthy cts Lev_0s Lev_Sucs))
-          |> Thm.close_derivation);
-
-        val length_Lev' = mk_specN (n + 1) length_Lev;
-        val length_Levs = map (fn i => length_Lev' RS mk_conjunctN n i RS mp) ks;
-
-        fun mk_goal i z = fold_rev Logic.all (z :: kl :: nat :: ss) (Logic.mk_implies
-            (HOLogic.mk_Trueprop (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z)),
-            HOLogic.mk_Trueprop (HOLogic.mk_mem (kl, mk_Lev ss (mk_size kl) i $ z))));
-        val goals = map2 mk_goal ks zs;
-
-        val length_Levs' = map2 (fn goal => fn length_Lev =>
-          Goal.prove_sorry lthy [] [] goal (K (mk_length_Lev'_tac length_Lev))
-          |> Thm.close_derivation) goals length_Levs;
-      in
-        (length_Levs, length_Levs')
-      end;
-
-    val prefCl_Lev_thms =
-      let
-        fun mk_conjunct i z = HOLogic.mk_imp
-          (HOLogic.mk_conj (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z), mk_prefixeq kl_copy kl),
-          HOLogic.mk_mem (kl_copy, mk_Lev ss (mk_size kl_copy) i $ z));
-        val goal = list_all_free (kl :: kl_copy :: zs)
-          (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs));
-
-        val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
-
-        val prefCl_Lev = singleton (Proof_Context.export names_lthy lthy)
-          (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
-            (K (mk_prefCl_Lev_tac lthy cts Lev_0s Lev_Sucs)))
-          |> Thm.close_derivation;
-
-        val prefCl_Lev' = mk_specN (n + 2) prefCl_Lev;
-      in
-        map (fn i => prefCl_Lev' RS mk_conjunctN n i RS mp) ks
-      end;
-
-    val rv_last_thmss =
-      let
-        fun mk_conjunct i z i' z_copy = list_exists_free [z_copy]
-          (HOLogic.mk_eq
-            (mk_rv ss (mk_append (kl, HOLogic.mk_list sum_sbdT [mk_InN sbdTs k i'])) i $ z,
-            mk_InN activeAs z_copy i'));
-        val goal = list_all_free (k :: zs)
-          (Library.foldr1 HOLogic.mk_conj (map2 (fn i => fn z =>
-            Library.foldr1 HOLogic.mk_conj
-              (map2 (mk_conjunct i z) ks zs_copy)) ks zs));
-
-        val cTs = [SOME (certifyT lthy sum_sbdT)];
-        val cts = map (SOME o certify lthy) [Term.absfree kl' goal, kl];
-
-        val rv_last = singleton (Proof_Context.export names_lthy lthy)
-          (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
-            (K (mk_rv_last_tac cTs cts rv_Nils rv_Conss)))
-          |> Thm.close_derivation;
-
-        val rv_last' = mk_specN (n + 1) rv_last;
-      in
-        map (fn i => map (fn i' => rv_last' RS mk_conjunctN n i RS mk_conjunctN n i') ks) ks
-      end;
-
-    val set_rv_Lev_thmsss = if m = 0 then replicate n (replicate n []) else
-      let
-        fun mk_case s sets z z_free = Term.absfree z_free (Library.foldr1 HOLogic.mk_conj
-          (map2 (fn set => fn A => mk_leq (set $ (s $ z)) A) (take m sets) As));
-
-        fun mk_conjunct i z B = HOLogic.mk_imp
-          (HOLogic.mk_conj (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z), HOLogic.mk_mem (z, B)),
-          mk_sum_caseN (map4 mk_case ss setssAs zs zs') $ (mk_rv ss kl i $ z));
-
-        val goal = list_all_free (kl :: zs)
-          (Library.foldr1 HOLogic.mk_conj (map3 mk_conjunct ks zs Bs));
-
-        val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
-
-        val set_rv_Lev = singleton (Proof_Context.export names_lthy lthy)
-          (Goal.prove_sorry lthy [] []
-            (Logic.mk_implies (coalg_prem, HOLogic.mk_Trueprop goal))
-            (K (mk_set_rv_Lev_tac lthy m cts Lev_0s Lev_Sucs rv_Nils rv_Conss
-              coalg_set_thmss from_to_sbd_thmss)))
-          |> Thm.close_derivation;
-
-        val set_rv_Lev' = mk_specN (n + 1) set_rv_Lev;
-      in
-        map (fn i => map (fn i' =>
-          split_conj_thm (if n = 1 then set_rv_Lev' RS mk_conjunctN n i RS mp
-            else set_rv_Lev' RS mk_conjunctN n i RS mp RSN
-              (2, @{thm sum_case_weak_cong} RS iffD1) RS
-              (mk_sum_casesN n i' RS iffD1))) ks) ks
-      end;
-
-    val set_Lev_thmsss =
-      let
-        fun mk_conjunct i z =
-          let
-            fun mk_conjunct' i' sets s z' =
-              let
-                fun mk_conjunct'' i'' set z'' = HOLogic.mk_imp
-                  (HOLogic.mk_mem (z'', set $ (s $ z')),
-                    HOLogic.mk_mem (mk_append (kl,
-                      HOLogic.mk_list sum_sbdT [mk_InN sbdTs (mk_to_sbd s z' i' i'' $ z'') i'']),
-                      mk_Lev ss (HOLogic.mk_Suc nat) i $ z));
-              in
-                HOLogic.mk_imp (HOLogic.mk_eq (mk_rv ss kl i $ z, mk_InN activeAs z' i'),
-                  (Library.foldr1 HOLogic.mk_conj (map3 mk_conjunct'' ks (drop m sets) zs_copy2)))
-              end;
-          in
-            HOLogic.mk_imp (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z),
-              Library.foldr1 HOLogic.mk_conj (map4 mk_conjunct' ks setssAs ss zs_copy))
-          end;
-
-        val goal = list_all_free (kl :: zs @ zs_copy @ zs_copy2)
-          (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs));
-
-        val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
-
-        val set_Lev = singleton (Proof_Context.export names_lthy lthy)
-          (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
-            (K (mk_set_Lev_tac lthy cts Lev_0s Lev_Sucs rv_Nils rv_Conss from_to_sbd_thmss)))
-          |> Thm.close_derivation;
-
-        val set_Lev' = mk_specN (3 * n + 1) set_Lev;
-      in
-        map (fn i => map (fn i' => map (fn i'' => set_Lev' RS
-          mk_conjunctN n i RS mp RS
-          mk_conjunctN n i' RS mp RS
-          mk_conjunctN n i'' RS mp) ks) ks) ks
-      end;
-
-    val set_image_Lev_thmsss =
-      let
-        fun mk_conjunct i z =
-          let
-            fun mk_conjunct' i' sets =
-              let
-                fun mk_conjunct'' i'' set s z'' = HOLogic.mk_imp
-                  (HOLogic.mk_eq (mk_rv ss kl i $ z, mk_InN activeAs z'' i''),
-                  HOLogic.mk_mem (k, mk_image (mk_to_sbd s z'' i'' i') $ (set $ (s $ z''))));
-              in
-                HOLogic.mk_imp (HOLogic.mk_mem
-                  (mk_append (kl, HOLogic.mk_list sum_sbdT [mk_InN sbdTs k i']),
-                    mk_Lev ss (HOLogic.mk_Suc nat) i $ z),
-                  (Library.foldr1 HOLogic.mk_conj (map4 mk_conjunct'' ks sets ss zs_copy)))
-              end;
-          in
-            HOLogic.mk_imp (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z),
-              Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct' ks (drop m setssAs')))
-          end;
-
-        val goal = list_all_free (kl :: k :: zs @ zs_copy)
-          (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs));
-
-        val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
-
-        val set_image_Lev = singleton (Proof_Context.export names_lthy lthy)
-          (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
-            (K (mk_set_image_Lev_tac lthy cts Lev_0s Lev_Sucs rv_Nils rv_Conss
-              from_to_sbd_thmss to_sbd_inj_thmss)))
-          |> Thm.close_derivation;
-
-        val set_image_Lev' = mk_specN (2 * n + 2) set_image_Lev;
-      in
-        map (fn i => map (fn i' => map (fn i'' => set_image_Lev' RS
-          mk_conjunctN n i RS mp RS
-          mk_conjunctN n i'' RS mp RS
-          mk_conjunctN n i' RS mp) ks) ks) ks
-      end;
-
-    val mor_beh_thm =
-      Goal.prove_sorry lthy [] []
-        (fold_rev Logic.all (As @ Bs @ ss) (Logic.mk_implies (coalg_prem,
-          HOLogic.mk_Trueprop (mk_mor Bs ss carTAs strTAs (map (mk_beh ss) ks)))))
-        (mk_mor_beh_tac m mor_def mor_cong_thm
-          beh_defs carT_defs strT_defs isNode_defs
-          to_sbd_inj_thmss from_to_sbd_thmss Lev_0s Lev_Sucs rv_Nils rv_Conss Lev_sbd_thms
-          length_Lev_thms length_Lev'_thms prefCl_Lev_thms rv_last_thmss
-          set_rv_Lev_thmsss set_Lev_thmsss set_image_Lev_thmsss
-          set_mapss coalg_set_thmss map_comp_id_thms map_cong0s map_arg_cong_thms)
-      |> Thm.close_derivation;
-
-    val timer = time (timer "Behavioral morphism");
-
-    fun mk_LSBIS As i = mk_lsbis As (map (mk_carT As) ks) strTAs i;
-    fun mk_car_final As i =
-      mk_quotient (mk_carT As i) (mk_LSBIS As i);
-    fun mk_str_final As i =
-      mk_univ (HOLogic.mk_comp (Term.list_comb (nth final_maps (i - 1),
-        passive_ids @ map (mk_proj o mk_LSBIS As) ks), nth strTAs (i - 1)));
-
-    val car_finalAs = map (mk_car_final As) ks;
-    val str_finalAs = map (mk_str_final As) ks;
-    val car_finals = map (mk_car_final passive_UNIVs) ks;
-    val str_finals = map (mk_str_final passive_UNIVs) ks;
-
-    val coalgT_set_thmss = map (map (fn thm => coalgT_thm RS thm)) coalg_set_thmss;
-    val equiv_LSBIS_thms = map (fn thm => coalgT_thm RS thm) equiv_lsbis_thms;
-
-    val congruent_str_final_thms =
-      let
-        fun mk_goal R final_map strT =
-          fold_rev Logic.all As (HOLogic.mk_Trueprop
-            (mk_congruent R (HOLogic.mk_comp
-              (Term.list_comb (final_map, passive_ids @ map (mk_proj o mk_LSBIS As) ks), strT))));
-
-        val goals = map3 mk_goal (map (mk_LSBIS As) ks) final_maps strTAs;
-      in
-        map4 (fn goal => fn lsbisE => fn map_comp_id => fn map_cong0 =>
-          Goal.prove_sorry lthy [] [] goal
-            (K (mk_congruent_str_final_tac m lsbisE map_comp_id map_cong0 equiv_LSBIS_thms))
-          |> Thm.close_derivation)
-        goals lsbisE_thms map_comp_id_thms map_cong0s
-      end;
-
-    val coalg_final_thm = Goal.prove_sorry lthy [] [] (fold_rev Logic.all As
-      (HOLogic.mk_Trueprop (mk_coalg As car_finalAs str_finalAs)))
-      (K (mk_coalg_final_tac m coalg_def congruent_str_final_thms equiv_LSBIS_thms
-        set_mapss coalgT_set_thmss))
-      |> Thm.close_derivation;
-
-    val mor_T_final_thm = Goal.prove_sorry lthy [] [] (fold_rev Logic.all As
-      (HOLogic.mk_Trueprop (mk_mor carTAs strTAs car_finalAs str_finalAs
-        (map (mk_proj o mk_LSBIS As) ks))))
-      (K (mk_mor_T_final_tac mor_def congruent_str_final_thms equiv_LSBIS_thms))
-      |> Thm.close_derivation;
-
-    val mor_final_thm = mor_comp_thm OF [mor_beh_thm, mor_T_final_thm];
-    val in_car_final_thms = map (fn mor_image' => mor_image' OF
-      [tcoalg_thm RS mor_final_thm, UNIV_I]) mor_image'_thms;
-
-    val timer = time (timer "Final coalgebra");
-
-    val ((T_names, (T_glob_infos, T_loc_infos)), lthy) =
-      lthy
-      |> fold_map4 (fn b => fn mx => fn car_final => fn in_car_final =>
-        typedef (Binding.conceal b, params, mx) car_final NONE
-          (EVERY' [rtac exI, rtac in_car_final] 1)) bs mixfixes car_finals in_car_final_thms
-      |>> apsnd split_list o split_list;
-
-    val Ts = map (fn name => Type (name, params')) T_names;
-    fun mk_Ts passive = map (Term.typ_subst_atomic (passiveAs ~~ passive)) Ts;
-    val Ts' = mk_Ts passiveBs;
-    val Rep_Ts = map2 (fn info => fn T => Const (#Rep_name info, T --> treeQT)) T_glob_infos Ts;
-    val Abs_Ts = map2 (fn info => fn T => Const (#Abs_name info, treeQT --> T)) T_glob_infos Ts;
-
-    val Reps = map #Rep T_loc_infos;
-    val Rep_injects = map #Rep_inject T_loc_infos;
-    val Abs_inverses = map #Abs_inverse T_loc_infos;
-
-    val timer = time (timer "THE TYPEDEFs & Rep/Abs thms");
-
-    val UNIVs = map HOLogic.mk_UNIV Ts;
-    val FTs = mk_FTs (passiveAs @ Ts);
-    val FTs' = mk_FTs (passiveBs @ Ts);
-    val prodTs = map (HOLogic.mk_prodT o `I) Ts;
-    val prodFTs = mk_FTs (passiveAs @ prodTs);
-    val FTs_setss = mk_setss (passiveAs @ Ts);
-    val prodFT_setss = mk_setss (passiveAs @ prodTs);
-    val map_FTs = map2 (fn Ds => mk_map_of_bnf Ds treeQTs (passiveAs @ Ts)) Dss bnfs;
-    val map_FT_nths = map2 (fn Ds =>
-      mk_map_of_bnf Ds (passiveAs @ prodTs) (passiveAs @ Ts)) Dss bnfs;
-    val fstsTs = map fst_const prodTs;
-    val sndsTs = map snd_const prodTs;
-    val dtorTs = map2 (curry op -->) Ts FTs;
-    val ctorTs = map2 (curry op -->) FTs Ts;
-    val unfold_fTs = map2 (curry op -->) activeAs Ts;
-    val corec_sTs = map (Term.typ_subst_atomic (activeBs ~~ Ts)) sum_sTs;
-    val corec_maps = map (Term.subst_atomic_types (activeBs ~~ Ts)) map_Inls;
-    val corec_maps_rev = map (Term.subst_atomic_types (activeBs ~~ Ts)) map_Inls_rev;
-    val corec_Inls = map (Term.subst_atomic_types (activeBs ~~ Ts)) Inls;
-    val corec_UNIVs = map2 (HOLogic.mk_UNIV oo curry mk_sumT) Ts activeAs;
-
-    val (((((((((((((Jzs, Jzs'), Jz's), Jzs_copy), Jz's_copy), Jzs1), Jzs2),
-      FJzs), TRs), unfold_fs), corec_ss), phis), dtor_set_induct_phiss),
-      names_lthy) = names_lthy
-      |> mk_Frees' "z" Ts
-      ||>> mk_Frees "y" Ts'
-      ||>> mk_Frees "z'" Ts
-      ||>> mk_Frees "y'" Ts'
-      ||>> mk_Frees "z1" Ts
-      ||>> mk_Frees "z2" Ts
-      ||>> mk_Frees "x" prodFTs
-      ||>> mk_Frees "r" (map (mk_relT o `I) Ts)
-      ||>> mk_Frees "f" unfold_fTs
-      ||>> mk_Frees "s" corec_sTs
-      ||>> mk_Frees "P" (map2 mk_pred2T Ts Ts)
-      ||>> mk_Freess "P" (map (fn A => map (mk_pred2T A) Ts) passiveAs);
-
-    fun dtor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtorN ^ "_");
-    val dtor_name = Binding.name_of o dtor_bind;
-    val dtor_def_bind = rpair [] o Binding.conceal o Thm.def_binding o dtor_bind;
-
-    fun dtor_spec i rep str map_FT dtorT Jz Jz' =
-      let
-        val lhs = Free (dtor_name i, dtorT);
-        val rhs = Term.absfree Jz'
-          (Term.list_comb (map_FT, map HOLogic.id_const passiveAs @ Abs_Ts) $
-            (str $ (rep $ Jz)));
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((dtor_frees, (_, dtor_def_frees)), (lthy, lthy_old)) =
-      lthy
-      |> fold_map7 (fn i => fn rep => fn str => fn mapx => fn dtorT => fn Jz => fn Jz' =>
-        Specification.definition (SOME (dtor_bind i, NONE, NoSyn),
-          (dtor_def_bind i, dtor_spec i rep str mapx dtorT Jz Jz')))
-        ks Rep_Ts str_finals map_FTs dtorTs Jzs Jzs'
-      |>> apsnd split_list o split_list
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-    fun mk_dtors passive =
-      map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ (mk_params passive)) o
-        Morphism.term phi) dtor_frees;
-    val dtors = mk_dtors passiveAs;
-    val dtor's = mk_dtors passiveBs;
-    val dtor_defs = map ((fn thm => thm RS fun_cong) o Morphism.thm phi) dtor_def_frees;
-
-    val coalg_final_set_thmss = map (map (fn thm => coalg_final_thm RS thm)) coalg_set_thmss;
-    val (mor_Rep_thm, mor_Abs_thm) =
-      let
-        val mor_Rep =
-          Goal.prove_sorry lthy [] []
-            (HOLogic.mk_Trueprop (mk_mor UNIVs dtors car_finals str_finals Rep_Ts))
-            (mk_mor_Rep_tac m (mor_def :: dtor_defs) Reps Abs_inverses coalg_final_set_thmss
-              map_comp_id_thms map_cong0L_thms)
-          |> Thm.close_derivation;
-
-        val mor_Abs =
-          Goal.prove_sorry lthy [] []
-            (HOLogic.mk_Trueprop (mk_mor car_finals str_finals UNIVs dtors Abs_Ts))
-            (mk_mor_Abs_tac (mor_def :: dtor_defs) Abs_inverses)
-          |> Thm.close_derivation;
-      in
-        (mor_Rep, mor_Abs)
-      end;
-
-    val timer = time (timer "dtor definitions & thms");
-
-    fun unfold_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtor_unfoldN ^ "_");
-    val unfold_name = Binding.name_of o unfold_bind;
-    val unfold_def_bind = rpair [] o Binding.conceal o Thm.def_binding o unfold_bind;
-
-    fun unfold_spec i T AT abs f z z' =
-      let
-        val unfoldT = Library.foldr (op -->) (sTs, AT --> T);
-
-        val lhs = Term.list_comb (Free (unfold_name i, unfoldT), ss);
-        val rhs = Term.absfree z' (abs $ (f $ z));
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((unfold_frees, (_, unfold_def_frees)), (lthy, lthy_old)) =
-      lthy
-      |> fold_map7 (fn i => fn T => fn AT => fn abs => fn f => fn z => fn z' =>
-        Specification.definition
-          (SOME (unfold_bind i, NONE, NoSyn), (unfold_def_bind i, unfold_spec i T AT abs f z z')))
-          ks Ts activeAs Abs_Ts (map (fn i => HOLogic.mk_comp
-            (mk_proj (mk_LSBIS passive_UNIVs i), mk_beh ss i)) ks) zs zs'
-      |>> apsnd split_list o split_list
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-    val unfolds = map (Morphism.term phi) unfold_frees;
-    val unfold_names = map (fst o dest_Const) unfolds;
-    fun mk_unfolds passives actives =
-      map3 (fn name => fn T => fn active =>
-        Const (name, Library.foldr (op -->)
-          (map2 (curry op -->) actives (mk_FTs (passives @ actives)), active --> T)))
-      unfold_names (mk_Ts passives) actives;
-    fun mk_unfold Ts ss i = Term.list_comb (Const (nth unfold_names (i - 1), Library.foldr (op -->)
-      (map fastype_of ss, domain_type (fastype_of (nth ss (i - 1))) --> nth Ts (i - 1))), ss);
-    val unfold_defs = map ((fn thm => thm RS fun_cong) o Morphism.thm phi) unfold_def_frees;
-
-    val mor_unfold_thm =
-      let
-        val Abs_inverses' = map2 (curry op RS) in_car_final_thms Abs_inverses;
-        val morEs' = map (fn thm =>
-          (thm OF [tcoalg_thm RS mor_final_thm, UNIV_I]) RS sym) morE_thms;
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all ss
-            (HOLogic.mk_Trueprop (mk_mor active_UNIVs ss UNIVs dtors (map (mk_unfold Ts ss) ks))))
-          (K (mk_mor_unfold_tac m mor_UNIV_thm dtor_defs unfold_defs Abs_inverses' morEs'
-            map_comp_id_thms map_cong0s))
-        |> Thm.close_derivation
-      end;
-    val dtor_unfold_thms = map (fn thm => (thm OF [mor_unfold_thm, UNIV_I]) RS sym) morE_thms;
-
-    val (raw_coind_thms, raw_coind_thm) =
-      let
-        val prem = HOLogic.mk_Trueprop (mk_sbis passive_UNIVs UNIVs dtors TRs);
-        val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
-          (map2 (fn R => fn T => mk_leq R (Id_const T)) TRs Ts));
-        val goal = fold_rev Logic.all TRs (Logic.mk_implies (prem, concl));
-      in
-        `split_conj_thm (Goal.prove_sorry lthy [] [] goal
-          (K (mk_raw_coind_tac bis_def bis_cong_thm bis_O_thm bis_converse_thm bis_Gr_thm
-            tcoalg_thm coalgT_thm mor_T_final_thm sbis_lsbis_thm
-            lsbis_incl_thms incl_lsbis_thms equiv_LSBIS_thms mor_Rep_thm Rep_injects))
-          |> Thm.close_derivation)
-      end;
-
-    val (unfold_unique_mor_thms, unfold_unique_mor_thm) =
-      let
-        val prem = HOLogic.mk_Trueprop (mk_mor active_UNIVs ss UNIVs dtors unfold_fs);
-        fun mk_fun_eq f i = HOLogic.mk_eq (f, mk_unfold Ts ss i);
-        val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
-          (map2 mk_fun_eq unfold_fs ks));
-
-        val bis_thm = tcoalg_thm RSN (2, tcoalg_thm RS bis_image2_thm);
-        val mor_thm = mor_comp_thm OF [tcoalg_thm RS mor_final_thm, mor_Abs_thm];
-
-        val unique_mor = Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (ss @ unfold_fs) (Logic.mk_implies (prem, unique)))
-          (K (mk_unfold_unique_mor_tac raw_coind_thms bis_thm mor_thm unfold_defs))
-          |> Thm.close_derivation;
-      in
-        `split_conj_thm unique_mor
-      end;
-
-    val (dtor_unfold_unique_thms, dtor_unfold_unique_thm) = `split_conj_thm (split_conj_prems n
-      (mor_UNIV_thm RS iffD2 RS unfold_unique_mor_thm));
-
-    val unfold_dtor_thms = map (fn thm => mor_id_thm RS thm RS sym) unfold_unique_mor_thms;
-
-    val unfold_o_dtor_thms =
-      let
-        val mor = mor_comp_thm OF [mor_str_thm, mor_unfold_thm];
-      in
-        map2 (fn unique => fn unfold_ctor =>
-          trans OF [mor RS unique, unfold_ctor]) unfold_unique_mor_thms unfold_dtor_thms
-      end;
-
-    val timer = time (timer "unfold definitions & thms");
-
-    val map_dtors = map2 (fn Ds => fn bnf =>
-      Term.list_comb (mk_map_of_bnf Ds (passiveAs @ Ts) (passiveAs @ FTs) bnf,
-        map HOLogic.id_const passiveAs @ dtors)) Dss bnfs;
-
-    fun ctor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctorN ^ "_");
-    val ctor_name = Binding.name_of o ctor_bind;
-    val ctor_def_bind = rpair [] o Binding.conceal o Thm.def_binding o ctor_bind;
-
-    fun ctor_spec i ctorT =
-      let
-        val lhs = Free (ctor_name i, ctorT);
-        val rhs = mk_unfold Ts map_dtors i;
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((ctor_frees, (_, ctor_def_frees)), (lthy, lthy_old)) =
-      lthy
-      |> fold_map2 (fn i => fn ctorT =>
-        Specification.definition
-          (SOME (ctor_bind i, NONE, NoSyn), (ctor_def_bind i, ctor_spec i ctorT))) ks ctorTs
-      |>> apsnd split_list o split_list
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-    fun mk_ctors params =
-      map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ params) o Morphism.term phi)
-        ctor_frees;
-    val ctors = mk_ctors params';
-    val ctor_defs = map (Morphism.thm phi) ctor_def_frees;
-
-    val ctor_o_dtor_thms = map2 (fold_thms lthy o single) ctor_defs unfold_o_dtor_thms;
-
-    val dtor_o_ctor_thms =
-      let
-        fun mk_goal dtor ctor FT =
-         mk_Trueprop_eq (HOLogic.mk_comp (dtor, ctor), HOLogic.id_const FT);
-        val goals = map3 mk_goal dtors ctors FTs;
-      in
-        map5 (fn goal => fn ctor_def => fn unfold => fn map_comp_id => fn map_cong0L =>
-          Goal.prove_sorry lthy [] [] goal
-            (mk_dtor_o_ctor_tac ctor_def unfold map_comp_id map_cong0L unfold_o_dtor_thms)
-          |> Thm.close_derivation)
-          goals ctor_defs dtor_unfold_thms map_comp_id_thms map_cong0L_thms
-      end;
-
-    val dtor_ctor_thms = map (fn thm => thm RS @{thm pointfree_idE}) dtor_o_ctor_thms;
-    val ctor_dtor_thms = map (fn thm => thm RS @{thm pointfree_idE}) ctor_o_dtor_thms;
-
-    val bij_dtor_thms =
-      map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) ctor_o_dtor_thms dtor_o_ctor_thms;
-    val inj_dtor_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_dtor_thms;
-    val surj_dtor_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_dtor_thms;
-    val dtor_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_dtor_thms;
-    val dtor_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_dtor_thms;
-    val dtor_exhaust_thms = map (fn thm => thm RS exE) dtor_nchotomy_thms;
-
-    val bij_ctor_thms =
-      map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) dtor_o_ctor_thms ctor_o_dtor_thms;
-    val inj_ctor_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_ctor_thms;
-    val surj_ctor_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_ctor_thms;
-    val ctor_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_ctor_thms;
-    val ctor_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_ctor_thms;
-    val ctor_exhaust_thms = map (fn thm => thm RS exE) ctor_nchotomy_thms;
-
-    val timer = time (timer "ctor definitions & thms");
-
-    val corec_Inl_sum_thms =
-      let
-        val mor = mor_comp_thm OF [mor_sum_case_thm, mor_unfold_thm];
-      in
-        map2 (fn unique => fn unfold_dtor =>
-          trans OF [mor RS unique, unfold_dtor]) unfold_unique_mor_thms unfold_dtor_thms
-      end;
-
-    fun corec_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtor_corecN ^ "_");
-    val corec_name = Binding.name_of o corec_bind;
-    val corec_def_bind = rpair [] o Binding.conceal o Thm.def_binding o corec_bind;
-
-    val corec_strs =
-      map3 (fn dtor => fn sum_s => fn mapx =>
-        mk_sum_case
-          (HOLogic.mk_comp (Term.list_comb (mapx, passive_ids @ corec_Inls), dtor), sum_s))
-      dtors corec_ss corec_maps;
-
-    fun corec_spec i T AT =
-      let
-        val corecT = Library.foldr (op -->) (corec_sTs, AT --> T);
-
-        val lhs = Term.list_comb (Free (corec_name i, corecT), corec_ss);
-        val rhs = HOLogic.mk_comp (mk_unfold Ts corec_strs i, Inr_const T AT);
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((corec_frees, (_, corec_def_frees)), (lthy, lthy_old)) =
-      lthy
-      |> fold_map3 (fn i => fn T => fn AT =>
-        Specification.definition
-          (SOME (corec_bind i, NONE, NoSyn), (corec_def_bind i, corec_spec i T AT)))
-          ks Ts activeAs
-      |>> apsnd split_list o split_list
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-    val corecs = map (Morphism.term phi) corec_frees;
-    val corec_names = map (fst o dest_Const) corecs;
-    fun mk_corec ss i = Term.list_comb (Const (nth corec_names (i - 1), Library.foldr (op -->)
-      (map fastype_of ss, domain_type (fastype_of (nth ss (i - 1))) --> nth Ts (i - 1))), ss);
-    val corec_defs = map (Morphism.thm phi) corec_def_frees;
-
-    val sum_cases =
-      map2 (fn T => fn i => mk_sum_case (HOLogic.id_const T, mk_corec corec_ss i)) Ts ks;
-    val dtor_corec_thms =
-      let
-        fun mk_goal i corec_s corec_map dtor z =
-          let
-            val lhs = dtor $ (mk_corec corec_ss i $ z);
-            val rhs = Term.list_comb (corec_map, passive_ids @ sum_cases) $ (corec_s $ z);
-          in
-            fold_rev Logic.all (z :: corec_ss) (mk_Trueprop_eq (lhs, rhs))
-          end;
-        val goals = map5 mk_goal ks corec_ss corec_maps_rev dtors zs;
-      in
-        map3 (fn goal => fn unfold => fn map_cong0 =>
-          Goal.prove_sorry lthy [] [] goal
-            (mk_corec_tac m corec_defs unfold map_cong0 corec_Inl_sum_thms)
-          |> Thm.close_derivation)
-        goals dtor_unfold_thms map_cong0s
-      end;
-
-    val corec_unique_mor_thm =
-      let
-        val id_fs = map2 (fn T => fn f => mk_sum_case (HOLogic.id_const T, f)) Ts unfold_fs;
-        val prem = HOLogic.mk_Trueprop (mk_mor corec_UNIVs corec_strs UNIVs dtors id_fs);
-        fun mk_fun_eq f i = HOLogic.mk_eq (f, mk_corec corec_ss i);
-        val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
-          (map2 mk_fun_eq unfold_fs ks));
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (corec_ss @ unfold_fs) (Logic.mk_implies (prem, unique)))
-          (mk_corec_unique_mor_tac corec_defs corec_Inl_sum_thms unfold_unique_mor_thm)
-          |> Thm.close_derivation
-      end;
-
-    val map_id0s_o_id =
-      map (fn thm =>
-        mk_trans (thm RS @{thm arg_cong2[of _ _ _ _ "op o", OF _ refl]}) @{thm id_o})
-      map_id0s;
-
-    val (dtor_corec_unique_thms, dtor_corec_unique_thm) =
-      `split_conj_thm (split_conj_prems n
-        (mor_UNIV_thm RS iffD2 RS corec_unique_mor_thm)
-        |> Local_Defs.unfold lthy (@{thms o_sum_case o_id id_o o_assoc sum_case_o_inj(1)} @
-           map_id0s_o_id @ sym_map_comps)
-        OF replicate n @{thm arg_cong2[of _ _ _ _ sum_case, OF refl]});
-
-    val timer = time (timer "corec definitions & thms");
-
-    val (dtor_map_coinduct_thm, coinduct_params, dtor_coinduct_thm) =
-      let
-        val zs = Jzs1 @ Jzs2;
-        val frees = phis @ zs;
-
-        val rels = map (Term.subst_atomic_types ((activeAs ~~ Ts) @ (activeBs ~~ Ts))) relsAsBs;
-
-        fun mk_concl phi z1 z2 = HOLogic.mk_imp (phi $ z1 $ z2, HOLogic.mk_eq (z1, z2));
-        val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
-          (map3 mk_concl phis Jzs1 Jzs2));
-
-        fun mk_rel_prem phi dtor rel Jz Jz_copy =
-          let
-            val concl = Term.list_comb (rel, map HOLogic.eq_const passiveAs @ phis) $
-              (dtor $ Jz) $ (dtor $ Jz_copy);
-          in
-            HOLogic.mk_Trueprop
-              (list_all_free [Jz, Jz_copy] (HOLogic.mk_imp (phi $ Jz $ Jz_copy, concl)))
-          end;
-
-        val rel_prems = map5 mk_rel_prem phis dtors rels Jzs Jzs_copy;
-        val dtor_coinduct_goal =
-          fold_rev Logic.all frees (Logic.list_implies (rel_prems, concl));
-
-        val dtor_coinduct =
-          Goal.prove_sorry lthy [] [] dtor_coinduct_goal
-            (K (mk_dtor_coinduct_tac m raw_coind_thm bis_rel_thm rel_congs))
-          |> Thm.close_derivation;
-
-        fun mk_prem phi dtor map_nth sets Jz Jz_copy FJz =
-          let
-            val xs = [Jz, Jz_copy];
-
-            fun mk_map_conjunct nths x =
-              HOLogic.mk_eq (Term.list_comb (map_nth, passive_ids @ nths) $ FJz, dtor $ x);
-
-            fun mk_set_conjunct set phi z1 z2 =
-              list_all_free [z1, z2]
-                (HOLogic.mk_imp (HOLogic.mk_mem (HOLogic.mk_prod (z1, z2), set $ FJz),
-                  phi $ z1 $ z2));
-
-            val concl = list_exists_free [FJz] (HOLogic.mk_conj
-              (Library.foldr1 HOLogic.mk_conj (map2 mk_map_conjunct [fstsTs, sndsTs] xs),
-              Library.foldr1 HOLogic.mk_conj
-                (map4 mk_set_conjunct (drop m sets) phis Jzs1 Jzs2)));
-          in
-            fold_rev Logic.all xs (Logic.mk_implies
-              (HOLogic.mk_Trueprop (Term.list_comb (phi, xs)), HOLogic.mk_Trueprop concl))
-          end;
-
-        val prems = map7 mk_prem phis dtors map_FT_nths prodFT_setss Jzs Jzs_copy FJzs;
-
-        val dtor_map_coinduct_goal = fold_rev Logic.all frees (Logic.list_implies (prems, concl));
-        val dtor_map_coinduct =
-          Goal.prove_sorry lthy [] [] dtor_map_coinduct_goal
-            (K (mk_dtor_map_coinduct_tac m ks raw_coind_thm bis_def))
-          |> Thm.close_derivation;
-      in
-        (dtor_map_coinduct, rev (Term.add_tfrees dtor_map_coinduct_goal []), dtor_coinduct)
-      end;
-
-    val timer = time (timer "coinduction");
-
-    val setss_by_bnf = map (fn i => map2 (mk_hset dtors i) ls passiveAs) ks;
-    val setss_by_range = transpose setss_by_bnf;
-
-    val (hset_dtor_incl_thmss, hset_hset_dtor_incl_thmsss, dtor_hset_induct_thms) =
-      let
-        fun tinst_of dtor =
-          map (SOME o certify lthy) (dtor :: remove (op =) dtor dtors);
-        fun tinst_of' dtor = case tinst_of dtor of t :: ts => t :: NONE :: ts;
-        val Tinst = map (pairself (certifyT lthy))
-          (map Logic.varifyT_global (deads @ allAs) ~~ (deads @ passiveAs @ Ts));
-        val set_incl_thmss =
-          map2 (fn dtor => map (singleton (Proof_Context.export names_lthy lthy) o
-            Drule.instantiate' [] (tinst_of' dtor) o
-            Thm.instantiate (Tinst, []) o Drule.zero_var_indexes))
-          dtors set_incl_hset_thmss;
-
-        val tinst = splice (map (SOME o certify lthy) dtors) (replicate n NONE)
-        val set_minimal_thms =
-          map (Drule.instantiate' [] tinst o Thm.instantiate (Tinst, []) o
-            Drule.zero_var_indexes)
-          hset_minimal_thms;
-
-        val set_set_incl_thmsss =
-          map2 (fn dtor => map (map (singleton (Proof_Context.export names_lthy lthy) o
-            Drule.instantiate' [] (NONE :: tinst_of' dtor) o
-            Thm.instantiate (Tinst, []) o Drule.zero_var_indexes)))
-          dtors set_hset_incl_hset_thmsss;
-
-        val set_set_incl_thmsss' = transpose (map transpose set_set_incl_thmsss);
-
-        val incls =
-          maps (map (fn thm => thm RS @{thm subset_Collect_iff})) set_incl_thmss @
-            @{thms subset_Collect_iff[OF subset_refl]};
-
-        fun mk_induct_tinst phis jsets y y' =
-          map4 (fn phi => fn jset => fn Jz => fn Jz' =>
-            SOME (certify lthy (Term.absfree Jz' (HOLogic.mk_Collect (fst y', snd y',
-              HOLogic.mk_conj (HOLogic.mk_mem (y, jset $ Jz), phi $ y $ Jz))))))
-          phis jsets Jzs Jzs';
-        val dtor_set_induct_thms =
-          map6 (fn set_minimal => fn set_set_inclss => fn jsets => fn y => fn y' => fn phis =>
-            ((set_minimal
-              |> Drule.instantiate' [] (mk_induct_tinst phis jsets y y')
-              |> unfold_thms lthy incls) OF
-              (replicate n ballI @
-                maps (map (fn thm => thm RS @{thm subset_CollectI})) set_set_inclss))
-            |> singleton (Proof_Context.export names_lthy lthy)
-            |> rule_by_tactic lthy (ALLGOALS (TRY o etac asm_rl)))
-          set_minimal_thms set_set_incl_thmsss' setss_by_range ys ys' dtor_set_induct_phiss
-      in
-        (set_incl_thmss, set_set_incl_thmsss, dtor_set_induct_thms)
-      end;
-
-    fun mk_dtor_map_DEADID_thm dtor_inject map_id0 =
-      trans OF [iffD2 OF [dtor_inject, id_apply], map_id0 RS sym];
-
-    fun mk_dtor_Jrel_DEADID_thm dtor_inject bnf =
-      trans OF [rel_eq_of_bnf bnf RS @{thm predicate2_eqD}, dtor_inject] RS sym;
-
-    val JphiTs = map2 mk_pred2T passiveAs passiveBs;
-    val Jpsi1Ts = map2 mk_pred2T passiveAs passiveCs;
-    val Jpsi2Ts = map2 mk_pred2T passiveCs passiveBs;
-    val prodTsTs' = map2 (curry HOLogic.mk_prodT) Ts Ts';
-    val fstsTsTs' = map fst_const prodTsTs';
-    val sndsTsTs' = map snd_const prodTsTs';
-    val activephiTs = map2 mk_pred2T activeAs activeBs;
-    val activeJphiTs = map2 mk_pred2T Ts Ts';
-    val (((((Jphis, Jpsi1s), Jpsi2s), activephis), activeJphis), names_lthy) = names_lthy
-      |> mk_Frees "R" JphiTs
-      ||>> mk_Frees "R" Jpsi1Ts
-      ||>> mk_Frees "Q" Jpsi2Ts
-      ||>> mk_Frees "S" activephiTs
-      ||>> mk_Frees "JR" activeJphiTs;
-    val rels = map2 (fn Ds => mk_rel_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
-    val in_rels = map in_rel_of_bnf bnfs;
-
-    fun mk_Jrel_DEADID_coinduct_thm () = 
-      mk_rel_xtor_co_induct_thm Greatest_FP rels activeJphis (map HOLogic.eq_const Ts) Jphis
-        Jzs Jz's dtors dtor's (fn {context = ctxt, prems} =>
-          (unfold_thms_tac ctxt @{thms le_fun_def le_bool_def all_simps(1,2)[symmetric]} THEN
-          REPEAT_DETERM (rtac allI 1) THEN rtac (dtor_coinduct_thm OF prems) 1)) lthy;
-
-    (*register new codatatypes as BNFs*)
-    val (timer, Jbnfs, (dtor_Jmap_o_thms, dtor_Jmap_thms), dtor_Jset_thmss',
-      dtor_Jrel_thms, Jrel_coinduct_thm, Jbnf_notes, lthy) =
-      if m = 0 then
-        (timer, replicate n DEADID_bnf,
-        map_split (`(mk_pointfree lthy)) (map2 mk_dtor_map_DEADID_thm dtor_inject_thms map_ids),
-        replicate n [], map2 mk_dtor_Jrel_DEADID_thm dtor_inject_thms bnfs,
-        mk_Jrel_DEADID_coinduct_thm (), [], lthy)
-      else let
-        val fTs = map2 (curry op -->) passiveAs passiveBs;
-        val gTs = map2 (curry op -->) passiveBs passiveCs;
-        val uTs = map2 (curry op -->) Ts Ts';
-
-        val ((((((((fs, fs'), fs_copy), gs), us), (Jys, Jys')), (Jys_copy, Jys'_copy)),
-          (ys_copy, ys'_copy)), names_lthy) = names_lthy
-          |> mk_Frees' "f" fTs
-          ||>> mk_Frees "f" fTs
-          ||>> mk_Frees "g" gTs
-          ||>> mk_Frees "u" uTs
-          ||>> mk_Frees' "b" Ts'
-          ||>> mk_Frees' "b" Ts'
-          ||>> mk_Frees' "y" passiveAs;
-
-        val map_FTFT's = map2 (fn Ds =>
-          mk_map_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
-
-        fun mk_maps ATs BTs Ts mk_T =
-          map2 (fn Ds => mk_map_of_bnf Ds (ATs @ Ts) (BTs @ map mk_T Ts)) Dss bnfs;
-        fun mk_Fmap mk_const fs Ts Fmap = Term.list_comb (Fmap, fs @ map mk_const Ts);
-        fun mk_map mk_const mk_T Ts fs Ts' dtors mk_maps =
-          mk_unfold Ts' (map2 (fn dtor => fn Fmap =>
-            HOLogic.mk_comp (mk_Fmap mk_const fs Ts Fmap, dtor)) dtors (mk_maps Ts mk_T));
-        val mk_map_id = mk_map HOLogic.id_const I;
-        val mk_mapsAB = mk_maps passiveAs passiveBs;
-        val fs_maps = map (mk_map_id Ts fs Ts' dtors mk_mapsAB) ks;
-
-        val set_bss =
-          map (flat o map2 (fn B => fn b =>
-            if member (op =) resDs (TFree B) then [] else [b]) resBs) set_bss0;
-
-        fun close_wit I wit = (I, fold_rev Term.absfree (map (nth ys') I) wit);
-
-        val all_unitTs = replicate live HOLogic.unitT;
-        val unitTs = replicate n HOLogic.unitT;
-        val unit_funs = replicate n (Term.absdummy HOLogic.unitT HOLogic.unit);
-        fun mk_map_args I =
-          map (fn i =>
-            if member (op =) I i then Term.absdummy HOLogic.unitT (nth ys i)
-            else mk_undefined (HOLogic.unitT --> nth passiveAs i))
-          (0 upto (m - 1));
-
-        fun mk_nat_wit Ds bnf (I, wit) () =
-          let
-            val passiveI = filter (fn i => i < m) I;
-            val map_args = mk_map_args passiveI;
-          in
-            Term.absdummy HOLogic.unitT (Term.list_comb
-              (mk_map_of_bnf Ds all_unitTs (passiveAs @ unitTs) bnf, map_args @ unit_funs) $ wit)
-          end;
-
-        fun mk_dummy_wit Ds bnf I =
-          let
-            val map_args = mk_map_args I;
-          in
-            Term.absdummy HOLogic.unitT (Term.list_comb
-              (mk_map_of_bnf Ds all_unitTs (passiveAs @ unitTs) bnf, map_args @ unit_funs) $
-              mk_undefined (mk_T_of_bnf Ds all_unitTs bnf))
-          end;
-
-        val nat_witss =
-          map2 (fn Ds => fn bnf => mk_wits_of_bnf (replicate (nwits_of_bnf bnf) Ds)
-            (replicate (nwits_of_bnf bnf) (replicate live HOLogic.unitT)) bnf
-            |> map (fn (I, wit) =>
-              (I, Lazy.lazy (mk_nat_wit Ds bnf (I, Term.list_comb (wit, map (K HOLogic.unit) I))))))
-          Dss bnfs;
-
-        val nat_wit_thmss = map2 (curry op ~~) nat_witss (map wit_thmss_of_bnf bnfs)
-
-        val Iss = map (map fst) nat_witss;
-
-        fun filter_wits (I, wit) =
-          let val J = filter (fn i => i < m) I;
-          in (J, (length J < length I, wit)) end;
-
-        val wit_treess = map_index (fn (i, Is) =>
-          map_index (finish Iss m [i+m] (i+m)) Is) Iss
-          |> map (minimize_wits o map filter_wits o minimize_wits o flat);
-
-        val coind_wit_argsss =
-          map (map (tree_to_coind_wits nat_wit_thmss o snd o snd) o filter (fst o snd)) wit_treess;
-
-        val nonredundant_coind_wit_argsss =
-          fold (fn i => fn argsss =>
-            nth_map (i - 1) (filter_out (fn xs =>
-              exists (fn ys =>
-                let
-                  val xs' = (map (fst o fst) xs, snd (fst (hd xs)));
-                  val ys' = (map (fst o fst) ys, snd (fst (hd ys)));
-                in
-                  eq_pair (subset (op =)) (eq_set (op =)) (xs', ys') andalso not (fst xs' = fst ys')
-                end)
-              (flat argsss)))
-            argsss)
-          ks coind_wit_argsss;
-
-        fun prepare_args args =
-          let
-            val I = snd (fst (hd args));
-            val (dummys, args') =
-              map_split (fn i =>
-                (case find_first (fn arg => fst (fst arg) = i - 1) args of
-                  SOME (_, ((_, wit), thms)) => (NONE, (Lazy.force wit, thms))
-                | NONE =>
-                  (SOME (i - 1), (mk_dummy_wit (nth Dss (i - 1)) (nth bnfs (i - 1)) I, []))))
-              ks;
-          in
-            ((I, dummys), apsnd flat (split_list args'))
-          end;
-
-        fun mk_coind_wits ((I, dummys), (args, thms)) =
-          ((I, dummys), (map (fn i => mk_unfold Ts args i $ HOLogic.unit) ks, thms));
-
-        val coind_witss =
-          maps (map (mk_coind_wits o prepare_args)) nonredundant_coind_wit_argsss;
-
-        val witss = map2 (fn Ds => fn bnf => mk_wits_of_bnf
-          (replicate (nwits_of_bnf bnf) Ds)
-          (replicate (nwits_of_bnf bnf) (passiveAs @ Ts)) bnf) Dss bnfs;
-
-        val ctor_witss =
-          map (map (uncurry close_wit o tree_to_ctor_wit ys ctors witss o snd o snd) o
-            filter_out (fst o snd)) wit_treess;
-
-        fun mk_coind_wit_thms ((I, dummys), (wits, wit_thms)) =
-          let
-            fun mk_goal sets y y_copy y'_copy j =
-              let
-                fun mk_conjunct set z dummy wit =
-                  mk_Ball (set $ z) (Term.absfree y'_copy
-                    (if dummy = NONE orelse member (op =) I (j - 1) then
-                      HOLogic.mk_imp (HOLogic.mk_eq (z, wit),
-                        if member (op =) I (j - 1) then HOLogic.mk_eq (y_copy, y)
-                        else @{term False})
-                    else @{term True}));
-              in
-                fold_rev Logic.all (map (nth ys) I @ Jzs) (HOLogic.mk_Trueprop
-                  (Library.foldr1 HOLogic.mk_conj (map4 mk_conjunct sets Jzs dummys wits)))
-              end;
-            val goals = map5 mk_goal setss_by_range ys ys_copy ys'_copy ls;
-          in
-            map2 (fn goal => fn induct =>
-              Goal.prove_sorry lthy [] [] goal
-                (mk_coind_wit_tac induct dtor_unfold_thms (flat set_mapss) wit_thms)
-              |> Thm.close_derivation)
-            goals dtor_hset_induct_thms
-            |> map split_conj_thm
-            |> transpose
-            |> map (map_filter (try (fn thm => thm RS bspec RS mp)))
-            |> curry op ~~ (map_index Library.I (map (close_wit I) wits))
-            |> filter (fn (_, thms) => length thms = m)
-          end;
-
-        val coind_wit_thms = maps mk_coind_wit_thms coind_witss;
-
-        val (wit_thmss, all_witss) =
-          fold (fn ((i, wit), thms) => fn witss =>
-            nth_map i (fn (thms', wits) => (thms @ thms', wit :: wits)) witss)
-          coind_wit_thms (map (pair []) ctor_witss)
-          |> map (apsnd (map snd o minimize_wits))
-          |> split_list;
-
-        val (Jbnf_consts, lthy) =
-          fold_map8 (fn b => fn map_b => fn rel_b => fn set_bs => fn mapx => fn sets => fn wits =>
-              fn T => fn lthy =>
-            define_bnf_consts Dont_Inline (user_policy Note_Some lthy) (SOME deads)
-              map_b rel_b set_bs
-              ((((((b, T), fold_rev Term.absfree fs' mapx), sets), sbd), wits), NONE) lthy)
-          bs map_bs rel_bs set_bss fs_maps setss_by_bnf all_witss Ts lthy;
-
-        val (_, Jconsts, Jconst_defs, mk_Jconsts) = split_list4 Jbnf_consts;
-        val (_, Jsetss, Jbds_Ds, Jwitss_Ds, _) = split_list5 Jconsts;
-        val (Jmap_defs, Jset_defss, Jbd_defs, Jwit_defss, Jrel_defs) = split_list5 Jconst_defs;
-        val (mk_Jmaps_Ds, mk_Jt_Ds, _, mk_Jrels_Ds, _) = split_list5 mk_Jconsts;
-
-        val Jrel_unabs_defs = map (fn def => mk_unabs_def m (def RS meta_eq_to_obj_eq)) Jrel_defs;
-        val Jset_defs = flat Jset_defss;
-        val Jset_unabs_defs = map (fn def => def RS meta_eq_to_obj_eq RS fun_cong) Jset_defs;
-
-        fun mk_Jmaps As Bs = map (fn mk => mk deads As Bs) mk_Jmaps_Ds;
-        fun mk_Jsetss As = map2 (fn mk => fn Jsets => map (mk deads As) Jsets) mk_Jt_Ds Jsetss;
-        val Jbds = map2 (fn mk => mk deads passiveAs) mk_Jt_Ds Jbds_Ds;
-        val Jwitss =
-          map2 (fn mk => fn Jwits => map (mk deads passiveAs o snd) Jwits) mk_Jt_Ds Jwitss_Ds;
-        fun mk_Jrels As Bs = map (fn mk => mk deads As Bs) mk_Jrels_Ds;
-
-        val Jmaps = mk_Jmaps passiveAs passiveBs;
-        val fs_Jmaps = map (fn m => Term.list_comb (m, fs)) Jmaps;
-        val fs_copy_Jmaps = map (fn m => Term.list_comb (m, fs_copy)) Jmaps;
-        val gs_Jmaps = map (fn m => Term.list_comb (m, gs)) (mk_Jmaps passiveBs passiveCs);
-        val fgs_Jmaps = map (fn m => Term.list_comb (m, map2 (curry HOLogic.mk_comp) gs fs))
-          (mk_Jmaps passiveAs passiveCs);
-        val (Jsetss_by_range, Jsetss_by_bnf) = `transpose (mk_Jsetss passiveAs);
-
-        val timer = time (timer "bnf constants for the new datatypes");
-
-        val (dtor_Jmap_thms, Jmap_thms) =
-          let
-            fun mk_goal fs_Jmap map dtor dtor' = fold_rev Logic.all fs
-              (mk_Trueprop_eq (HOLogic.mk_comp (dtor', fs_Jmap),
-                HOLogic.mk_comp (Term.list_comb (map, fs @ fs_Jmaps), dtor)));
-            val goals = map4 mk_goal fs_Jmaps map_FTFT's dtors dtor's;
-            val cTs = map (SOME o certifyT lthy) FTs';
-            val maps =
-              map5 (fn goal => fn cT => fn unfold => fn map_comp => fn map_cong0 =>
-                Goal.prove_sorry lthy [] [] goal
-                  (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jmap_defs THEN
-                     mk_map_tac m n cT unfold map_comp map_cong0)
-                |> Thm.close_derivation)
-              goals cTs dtor_unfold_thms map_comps map_cong0s;
-          in
-            map_split (fn thm => (thm RS @{thm comp_eq_dest}, thm)) maps
-          end;
-
-        val dtor_Jmap_unique_thm =
-          let
-            fun mk_prem u map dtor dtor' =
-              mk_Trueprop_eq (HOLogic.mk_comp (dtor', u),
-                HOLogic.mk_comp (Term.list_comb (map, fs @ us), dtor));
-            val prems = map4 mk_prem us map_FTFT's dtors dtor's;
-            val goal =
-              HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
-                (map2 (curry HOLogic.mk_eq) us fs_Jmaps));
-          in
-            Goal.prove_sorry lthy [] []
-              (fold_rev Logic.all (us @ fs) (Logic.list_implies (prems, goal)))
-                (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jmap_defs THEN
-                  mk_dtor_map_unique_tac dtor_unfold_unique_thm sym_map_comps ctxt)
-              |> Thm.close_derivation
-          end;
-
-        val Jmap_comp0_thms =
-          let
-            val goal = fold_rev Logic.all (fs @ gs)
-              (HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
-                (map3 (fn fmap => fn gmap => fn fgmap =>
-                   HOLogic.mk_eq (HOLogic.mk_comp (gmap, fmap), fgmap))
-                fs_Jmaps gs_Jmaps fgs_Jmaps)))
-          in
-            split_conj_thm (Goal.prove_sorry lthy [] [] goal
-              (K (mk_map_comp0_tac Jmap_thms map_comp0s dtor_Jmap_unique_thm))
-              |> Thm.close_derivation)
-          end;
-
-        val timer = time (timer "map functions for the new codatatypes");
-
-        val (dtor_Jset_thmss', dtor_Jset_thmss) =
-          let
-            fun mk_simp_goal relate pas_set act_sets sets dtor z set =
-              relate (set $ z, mk_union (pas_set $ (dtor $ z),
-                 Library.foldl1 mk_union
-                   (map2 (fn X => mk_UNION (X $ (dtor $ z))) act_sets sets)));
-            fun mk_goals eq =
-              map2 (fn i => fn sets =>
-                map4 (fn Fsets =>
-                  mk_simp_goal eq (nth Fsets (i - 1)) (drop m Fsets) sets)
-                FTs_setss dtors Jzs sets)
-              ls Jsetss_by_range;
-
-            val le_goals = map
-              (fold_rev Logic.all Jzs o HOLogic.mk_Trueprop o Library.foldr1 HOLogic.mk_conj)
-              (mk_goals (uncurry mk_leq));
-            val set_le_thmss = map split_conj_thm
-              (map4 (fn goal => fn hset_minimal => fn set_hsets => fn set_hset_hsetss =>
-                Goal.prove_sorry lthy [] [] goal
-                  (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jset_defs THEN
-                    mk_set_le_tac n hset_minimal set_hsets set_hset_hsetss)
-                |> Thm.close_derivation)
-              le_goals hset_minimal_thms set_hset_thmss' set_hset_hset_thmsss');
-
-            val ge_goalss = map (map2 (fn z => fn goal =>
-                Logic.all z (HOLogic.mk_Trueprop goal)) Jzs)
-              (mk_goals (uncurry mk_leq o swap));
-            val set_ge_thmss = 
-              map3 (map3 (fn goal => fn set_incl_hset => fn set_hset_incl_hsets =>
-                Goal.prove_sorry lthy [] [] goal
-                  (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jset_defs THEN
-                    mk_set_ge_tac n set_incl_hset set_hset_incl_hsets)
-                |> Thm.close_derivation))
-              ge_goalss set_incl_hset_thmss' set_hset_incl_hset_thmsss'
-          in
-            map2 (map2 (fn le => fn ge => equalityI OF [le, ge])) set_le_thmss set_ge_thmss
-            |> `transpose
-          end;
-
-        val timer = time (timer "set functions for the new codatatypes");
-
-        val colss = map2 (fn j => fn T =>
-          map (fn i => mk_hset_rec dtors nat i j T) ks) ls passiveAs;
-        val colss' = map2 (fn j => fn T =>
-          map (fn i => mk_hset_rec dtor's nat i j T) ks) ls passiveBs;
-
-        val col_natural_thmss =
-          let
-            fun mk_col_natural f map z col col' =
-              HOLogic.mk_eq (mk_image f $ (col $ z), col' $ (map $ z));
-
-            fun mk_goal f cols cols' = list_all_free Jzs (Library.foldr1 HOLogic.mk_conj
-              (map4 (mk_col_natural f) fs_Jmaps Jzs cols cols'));
-
-            val goals = map3 mk_goal fs colss colss';
-
-            val ctss =
-              map (fn phi => map (SOME o certify lthy) [Term.absfree nat' phi, nat]) goals;
-
-            val thms =
-              map4 (fn goal => fn cts => fn rec_0s => fn rec_Sucs =>
-                singleton (Proof_Context.export names_lthy lthy)
-                  (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
-                    (mk_col_natural_tac cts rec_0s rec_Sucs dtor_Jmap_thms set_mapss))
-                |> Thm.close_derivation)
-              goals ctss hset_rec_0ss' hset_rec_Sucss';
-          in
-            map (split_conj_thm o mk_specN n) thms
-          end;
-
-        val col_bd_thmss =
-          let
-            fun mk_col_bd z col bd = mk_ordLeq (mk_card_of (col $ z)) bd;
-
-            fun mk_goal bds cols = list_all_free Jzs (Library.foldr1 HOLogic.mk_conj
-              (map3 mk_col_bd Jzs cols bds));
-
-            val goals = map (mk_goal Jbds) colss;
-
-            val ctss = map (fn phi => map (SOME o certify lthy) [Term.absfree nat' phi, nat])
-              (map (mk_goal (replicate n sbd)) colss);
-
-            val thms =
-              map5 (fn j => fn goal => fn cts => fn rec_0s => fn rec_Sucs =>
-                singleton (Proof_Context.export names_lthy lthy)
-                  (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
-                    (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jbd_defs THEN
-                      mk_col_bd_tac m j cts rec_0s rec_Sucs sbd_Card_order sbd_Cinfinite set_sbdss))
-                |> Thm.close_derivation)
-              ls goals ctss hset_rec_0ss' hset_rec_Sucss';
-          in
-            map (split_conj_thm o mk_specN n) thms
-          end;
-
-        val map_cong0_thms =
-          let
-            val cTs = map (SOME o certifyT lthy o
-              Term.typ_subst_atomic (passiveAs ~~ passiveBs) o TFree) coinduct_params;
-
-            fun mk_prem z set f g y y' =
-              mk_Ball (set $ z) (Term.absfree y' (HOLogic.mk_eq (f $ y, g $ y)));
-
-            fun mk_prems sets z =
-              Library.foldr1 HOLogic.mk_conj (map5 (mk_prem z) sets fs fs_copy ys ys')
-
-            fun mk_map_cong0 sets z fmap gmap =
-              HOLogic.mk_imp (mk_prems sets z, HOLogic.mk_eq (fmap $ z, gmap $ z));
-
-            fun mk_coind_body sets (x, T) z fmap gmap y y_copy =
-              HOLogic.mk_conj
-                (HOLogic.mk_mem (z, HOLogic.mk_Collect (x, T, mk_prems sets z)),
-                  HOLogic.mk_conj (HOLogic.mk_eq (y, fmap $ z),
-                    HOLogic.mk_eq (y_copy, gmap $ z)))
-
-            fun mk_cphi sets (z' as (x, T)) z fmap gmap y' y y'_copy y_copy =
-              HOLogic.mk_exists (x, T, mk_coind_body sets z' z fmap gmap y y_copy)
-              |> Term.absfree y'_copy
-              |> Term.absfree y'
-              |> certify lthy;
-
-            val cphis = map9 mk_cphi
-              Jsetss_by_bnf Jzs' Jzs fs_Jmaps fs_copy_Jmaps Jys' Jys Jys'_copy Jys_copy;
-
-            val coinduct = unfold_thms lthy Jset_defs
-              (Drule.instantiate' cTs (map SOME cphis) dtor_map_coinduct_thm);
-
-            val goal =
-              HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
-                (map4 mk_map_cong0 Jsetss_by_bnf Jzs fs_Jmaps fs_copy_Jmaps));
-
-            val thm = singleton (Proof_Context.export names_lthy lthy)
-              (Goal.prove_sorry lthy [] [] goal
-                (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jset_defs THEN
-                  mk_mcong_tac lthy m (rtac coinduct) map_comps dtor_Jmap_thms map_cong0s
-                    set_mapss set_hset_thmss set_hset_hset_thmsss))
-              |> Thm.close_derivation
-          in
-            split_conj_thm thm
-          end;
-
-        val in_Jrels = map (fn def => trans OF [def, @{thm OO_Grp_alt}] RS @{thm predicate2_eqD})
-            Jrel_unabs_defs;
-
-        val fold_Jsets = fold_thms lthy Jset_unabs_defs;
-        val dtor_Jset_incl_thmss = map (map fold_Jsets) hset_dtor_incl_thmss;
-        val dtor_set_Jset_incl_thmsss = map (map (map fold_Jsets)) hset_hset_dtor_incl_thmsss;
-        val dtor_Jset_induct_thms = map fold_Jsets dtor_hset_induct_thms;
-        val Jwit_thmss = map (map fold_Jsets) wit_thmss;
-
-        val Jrels = mk_Jrels passiveAs passiveBs;
-        val Jrelphis = map (fn rel => Term.list_comb (rel, Jphis)) Jrels;
-        val relphis = map (fn rel => Term.list_comb (rel, Jphis @ Jrelphis)) rels;
-        val Jrelpsi1s = map (fn rel => Term.list_comb (rel, Jpsi1s)) (mk_Jrels passiveAs passiveCs);
-        val Jrelpsi2s = map (fn rel => Term.list_comb (rel, Jpsi2s)) (mk_Jrels passiveCs passiveBs);
-        val Jrelpsi12s = map (fn rel =>
-            Term.list_comb (rel, map2 (curry mk_rel_compp) Jpsi1s Jpsi2s)) Jrels;
-
-        val dtor_Jrel_thms =
-          let
-            fun mk_goal Jz Jz' dtor dtor' Jrelphi relphi = fold_rev Logic.all (Jz :: Jz' :: Jphis)
-              (mk_Trueprop_eq (Jrelphi $ Jz $ Jz', relphi $ (dtor $ Jz) $ (dtor' $ Jz')));
-            val goals = map6 mk_goal Jzs Jz's dtors dtor's Jrelphis relphis;
-          in
-            map12 (fn i => fn goal => fn in_rel => fn map_comp0 => fn map_cong0 =>
-              fn dtor_map => fn dtor_sets => fn dtor_inject => fn dtor_ctor =>
-              fn set_map0s => fn dtor_set_incls => fn dtor_set_set_inclss =>
-              Goal.prove_sorry lthy [] [] goal
-                (K (mk_dtor_rel_tac lthy in_Jrels i in_rel map_comp0 map_cong0 dtor_map dtor_sets
-                  dtor_inject dtor_ctor set_map0s dtor_set_incls dtor_set_set_inclss))
-              |> Thm.close_derivation)
-            ks goals in_rels map_comps map_cong0s dtor_Jmap_thms dtor_Jset_thmss'
-              dtor_inject_thms dtor_ctor_thms set_mapss dtor_Jset_incl_thmss
-              dtor_set_Jset_incl_thmsss
-          end;
-
-      val passiveABs = map2 (curry HOLogic.mk_prodT) passiveAs passiveBs;
-      val zip_ranTs = passiveABs @ prodTsTs';
-      val allJphis = Jphis @ activeJphis;
-      val zipFTs = mk_FTs zip_ranTs;
-      val zipTs = map3 (fn T => fn T' => fn FT => T --> T' --> FT) Ts Ts' zipFTs;
-      val zip_zTs = mk_Ts passiveABs;
-      val (((zips, (abs, abs')), zip_zs), names_lthy) = names_lthy
-        |> mk_Frees "zip" zipTs
-        ||>> mk_Frees' "ab" passiveABs
-        ||>> mk_Frees "z" zip_zTs;
-
-      val Iphi_sets =
-        map2 (fn phi => fn T => HOLogic.Collect_const T $ HOLogic.mk_split phi) allJphis zip_ranTs;
-      val in_phis = map2 (mk_in Iphi_sets) (mk_setss zip_ranTs) zipFTs;
-      val fstABs = map fst_const passiveABs;
-      val all_fsts = fstABs @ fstsTsTs';
-      val map_all_fsts = map2 (fn Ds => fn bnf =>
-        Term.list_comb (mk_map_of_bnf Ds zip_ranTs (passiveAs @ Ts) bnf, all_fsts)) Dss bnfs;
-      val Jmap_fsts = map2 (fn map => fn T => if m = 0 then HOLogic.id_const T
-        else Term.list_comb (map, fstABs)) (mk_Jmaps passiveABs passiveAs) Ts;
-
-      val sndABs = map snd_const passiveABs;
-      val all_snds = sndABs @ sndsTsTs';
-      val map_all_snds = map2 (fn Ds => fn bnf =>
-        Term.list_comb (mk_map_of_bnf Ds zip_ranTs (passiveBs @ Ts') bnf, all_snds)) Dss bnfs;
-      val Jmap_snds = map2 (fn map => fn T => if m = 0 then HOLogic.id_const T
-        else Term.list_comb (map, sndABs)) (mk_Jmaps passiveABs passiveBs) Ts;
-      val zip_unfolds = map (mk_unfold zip_zTs (map HOLogic.mk_split zips)) ks;
-      val zip_setss = mk_Jsetss passiveABs |> transpose;
-
-      val Jrel_coinduct_tac =
-        let
-          fun mk_helper_prem phi in_phi zip x y map map' dtor dtor' =
-            let
-              val zipxy = zip $ x $ y;
-            in
-              HOLogic.mk_Trueprop (list_all_free [x, y]
-                (HOLogic.mk_imp (phi $ x $ y, HOLogic.mk_conj (HOLogic.mk_mem (zipxy, in_phi),
-                  HOLogic.mk_conj (HOLogic.mk_eq (map $ zipxy, dtor $ x),
-                    HOLogic.mk_eq (map' $ zipxy, dtor' $ y))))))
-            end;
-          val helper_prems = map9 mk_helper_prem
-            activeJphis in_phis zips Jzs Jz's map_all_fsts map_all_snds dtors dtor's;
-          fun mk_helper_coind_concl fst phi x alt y map zip_unfold =
-            HOLogic.mk_imp (list_exists_free [if fst then y else x] (HOLogic.mk_conj (phi $ x $ y,
-              HOLogic.mk_eq (alt, map $ (zip_unfold $ HOLogic.mk_prod (x, y))))),
-            HOLogic.mk_eq (alt, if fst then x else y));
-          val helper_coind1_concl =
-            HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
-              (map6 (mk_helper_coind_concl true)
-              activeJphis Jzs Jzs_copy Jz's Jmap_fsts zip_unfolds));
-          val helper_coind2_concl =
-            HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
-              (map6 (mk_helper_coind_concl false)
-              activeJphis Jzs Jz's_copy Jz's Jmap_snds zip_unfolds));
-          val helper_coind_tac = mk_rel_coinduct_coind_tac m dtor_map_coinduct_thm ks map_comps
-            map_cong0s map_arg_cong_thms set_mapss dtor_unfold_thms dtor_Jmap_thms;
-          fun mk_helper_coind_thms vars concl =
-            Goal.prove_sorry lthy [] []
-              (fold_rev Logic.all (Jphis @ activeJphis @ vars @ zips)
-                (Logic.list_implies (helper_prems, concl)))
-              helper_coind_tac
-            |> Thm.close_derivation
-            |> split_conj_thm;
-          val helper_coind1_thms = mk_helper_coind_thms (Jzs @ Jzs_copy) helper_coind1_concl;
-          val helper_coind2_thms = mk_helper_coind_thms (Jz's @ Jz's_copy) helper_coind2_concl;
-  
-          fun mk_helper_ind_concl phi ab' ab fst snd z active_phi x y zip_unfold set =
-            mk_Ball (set $ z) (Term.absfree ab' (list_all_free [x, y] (HOLogic.mk_imp
-              (HOLogic.mk_conj (active_phi $ x $ y,
-                 HOLogic.mk_eq (z, zip_unfold $ HOLogic.mk_prod (x, y))),
-              phi $ (fst $ ab) $ (snd $ ab)))));
-  
-          val mk_helper_ind_concls =
-            map6 (fn Jphi => fn ab' => fn ab => fn fst => fn snd => fn zip_sets =>
-              map6 (mk_helper_ind_concl Jphi ab' ab fst snd)
-              zip_zs activeJphis Jzs Jz's zip_unfolds zip_sets)
-            Jphis abs' abs fstABs sndABs zip_setss
-            |> map (HOLogic.mk_Trueprop o Library.foldr1 HOLogic.mk_conj);
-  
-          val helper_ind_thmss = if m = 0 then replicate n [] else
-            map3 (fn concl => fn j => fn set_induct =>
-              Goal.prove_sorry lthy [] []
-                (fold_rev Logic.all (Jphis @ activeJphis @ zip_zs @ zips)
-                  (Logic.list_implies (helper_prems, concl)))
-                (mk_rel_coinduct_ind_tac m ks dtor_unfold_thms set_mapss j set_induct)
-              |> Thm.close_derivation
-              |> split_conj_thm)
-            mk_helper_ind_concls ls dtor_Jset_induct_thms
-            |> transpose;
-        in
-          mk_rel_coinduct_tac in_rels in_Jrels
-            helper_ind_thmss helper_coind1_thms helper_coind2_thms
-        end;
-
-      val Jrel_coinduct_thm =
-        mk_rel_xtor_co_induct_thm Greatest_FP rels activeJphis Jrels Jphis Jzs Jz's dtors dtor's
-          Jrel_coinduct_tac lthy;
-
-        val le_Jrel_OO_thm =
-          let
-            fun mk_le_Jrel_OO Jrelpsi1 Jrelpsi2 Jrelpsi12 =
-              mk_leq (mk_rel_compp (Jrelpsi1, Jrelpsi2)) Jrelpsi12;
-            val goals = map3 mk_le_Jrel_OO Jrelpsi1s Jrelpsi2s Jrelpsi12s;
-
-            val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj goals);
-          in
-            singleton (Proof_Context.export names_lthy lthy)
-              (Goal.prove_sorry lthy [] [] goal
-                (K (mk_le_rel_OO_tac Jrel_coinduct_thm dtor_Jrel_thms rel_OOs)))
-              |> Thm.close_derivation
-          end;
-
-        val timer = time (timer "helpers for BNF properties");
-
-        val map_id0_tacs =
-          map2 (K oo mk_map_id0_tac Jmap_thms) dtor_unfold_unique_thms unfold_dtor_thms;
-        val map_comp0_tacs = map (fn thm => K (rtac (thm RS sym) 1)) Jmap_comp0_thms;
-        val map_cong0_tacs = map (mk_map_cong0_tac m) map_cong0_thms;
-        val set_nat_tacss =
-          map2 (map2 (fn def => fn col => fn {context = ctxt, prems = _} =>
-              unfold_thms_tac ctxt Jset_defs THEN mk_set_map0_tac def col))
-            hset_defss (transpose col_natural_thmss);
-
-        val Jbd_card_orders = map (fn def => fold_thms lthy [def] sbd_card_order) Jbd_defs;
-        val Jbd_Cinfinites = map (fn def => fold_thms lthy [def] sbd_Cinfinite) Jbd_defs;
-
-        val bd_co_tacs = map (fn thm => K (rtac thm 1)) Jbd_card_orders;
-        val bd_cinf_tacs = map (fn thm => K (rtac (thm RS conjunct1) 1)) Jbd_Cinfinites;
-
-        val set_bd_tacss =
-          map3 (fn Cinf => map2 (fn def => fn col => fn {context = ctxt, prems = _} =>
-            unfold_thms_tac ctxt Jset_defs THEN mk_set_bd_tac Cinf def col))
-          Jbd_Cinfinites hset_defss (transpose col_bd_thmss);
-
-        val le_rel_OO_tacs = map (fn i => K (rtac (le_Jrel_OO_thm RS mk_conjunctN n i) 1)) ks;
-
-        val rel_OO_Grp_tacs = map (fn def => K (rtac def 1)) Jrel_unabs_defs;
-
-        val tacss = map9 zip_axioms map_id0_tacs map_comp0_tacs map_cong0_tacs set_nat_tacss
-          bd_co_tacs bd_cinf_tacs set_bd_tacss le_rel_OO_tacs rel_OO_Grp_tacs;
-
-        fun wit_tac thms {context = ctxt, prems = _} = unfold_thms_tac ctxt (flat Jwit_defss) THEN
-          mk_wit_tac n dtor_ctor_thms (flat dtor_Jset_thmss) (maps wit_thms_of_bnf bnfs) thms ctxt;
-
-        val (Jbnfs, lthy) =
-          fold_map6 (fn tacs => fn map_b => fn rel_b => fn set_bs => fn Jwit_thms => fn consts =>
-              fn lthy =>
-            bnf_def Do_Inline (user_policy Note_Some) I tacs (wit_tac Jwit_thms) (SOME deads)
-              map_b rel_b set_bs consts lthy
-            |> register_bnf (Local_Theory.full_name lthy b))
-          tacss map_bs rel_bs set_bss Jwit_thmss
-          ((((((bs ~~ Ts) ~~ Jmaps) ~~ Jsetss_by_bnf) ~~ Jbds) ~~ Jwitss) ~~ map SOME Jrels)
-          lthy;
-
-        val timer = time (timer "registered new codatatypes as BNFs");
-
-        val ls' = if m = 1 then [0] else ls;
-
-        val Jbnf_common_notes =
-          [(dtor_map_uniqueN, [dtor_Jmap_unique_thm])] @
-          map2 (fn i => fn thm => (mk_dtor_set_inductN i, [thm])) ls' dtor_Jset_induct_thms
-          |> map (fn (thmN, thms) =>
-            ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]));
-
-        val Jbnf_notes =
-          [(dtor_mapN, map single dtor_Jmap_thms),
-          (dtor_relN, map single dtor_Jrel_thms),
-          (dtor_set_inclN, dtor_Jset_incl_thmss),
-          (dtor_set_set_inclN, map flat dtor_set_Jset_incl_thmsss)] @
-          map2 (fn i => fn thms => (mk_dtor_setN i, map single thms)) ls' dtor_Jset_thmss
-          |> maps (fn (thmN, thmss) =>
-            map2 (fn b => fn thms =>
-              ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]))
-            bs thmss)
-      in
-        (timer, Jbnfs, (Jmap_thms, dtor_Jmap_thms), dtor_Jset_thmss',
-          dtor_Jrel_thms, Jrel_coinduct_thm, Jbnf_common_notes @ Jbnf_notes, lthy)
-      end;
-
-      val dtor_unfold_o_Jmap_thms = mk_xtor_un_fold_o_map_thms Greatest_FP false m
-        dtor_unfold_unique_thm dtor_Jmap_o_thms (map (mk_pointfree lthy) dtor_unfold_thms)
-        sym_map_comps map_cong0s;
-      val dtor_corec_o_Jmap_thms = mk_xtor_un_fold_o_map_thms Greatest_FP true m
-        dtor_corec_unique_thm dtor_Jmap_o_thms (map (mk_pointfree lthy) dtor_corec_thms)
-        sym_map_comps map_cong0s;
-
-      val rels = map2 (fn Ds => mk_rel_of_bnf Ds allAs allBs') Dss bnfs;
-
-      val dtor_unfold_transfer_thms =
-        mk_un_fold_transfer_thms Greatest_FP rels activephis
-          (if m = 0 then map HOLogic.eq_const Ts
-            else map (mk_rel_of_bnf deads passiveAs passiveBs) Jbnfs) Jphis
-          (mk_unfolds passiveAs activeAs) (mk_unfolds passiveBs activeBs)
-          (mk_unfold_transfer_tac m Jrel_coinduct_thm (map map_transfer_of_bnf bnfs)
-            dtor_unfold_thms)
-          lthy;
-
-      val timer = time (timer "relator coinduction");
-
-      val common_notes =
-        [(dtor_coinductN, [dtor_coinduct_thm]),
-        (dtor_map_coinductN, [dtor_map_coinduct_thm]),
-        (rel_coinductN, [Jrel_coinduct_thm]),
-        (dtor_unfold_transferN, dtor_unfold_transfer_thms)]
-        |> map (fn (thmN, thms) =>
-          ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]));
-
-      val notes =
-        [(ctor_dtorN, ctor_dtor_thms),
-        (ctor_exhaustN, ctor_exhaust_thms),
-        (ctor_injectN, ctor_inject_thms),
-        (dtor_corecN, dtor_corec_thms),
-        (dtor_ctorN, dtor_ctor_thms),
-        (dtor_exhaustN, dtor_exhaust_thms),
-        (dtor_injectN, dtor_inject_thms),
-        (dtor_unfoldN, dtor_unfold_thms),
-        (dtor_unfold_uniqueN, dtor_unfold_unique_thms),
-        (dtor_corec_uniqueN, dtor_corec_unique_thms),
-        (dtor_unfold_o_mapN, dtor_unfold_o_Jmap_thms),
-        (dtor_corec_o_mapN, dtor_corec_o_Jmap_thms)]
-        |> map (apsnd (map single))
-        |> maps (fn (thmN, thmss) =>
-          map2 (fn b => fn thms =>
-            ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]))
-          bs thmss);
-
-    (*FIXME: once the package exports all the necessary high-level characteristic theorems,
-       those should not only be concealed but rather not noted at all*)
-    val maybe_conceal_notes = note_all = false ? map (apfst (apfst Binding.conceal));
-  in
-    timer;
-    ({Ts = Ts, bnfs = Jbnfs, ctors = ctors, dtors = dtors,
-      xtor_co_iterss = transpose [unfolds, corecs],
-      xtor_co_induct = dtor_coinduct_thm, dtor_ctors = dtor_ctor_thms,
-      ctor_dtors = ctor_dtor_thms,
-      ctor_injects = ctor_inject_thms, dtor_injects = dtor_inject_thms,
-      xtor_map_thms = dtor_Jmap_thms, xtor_set_thmss = dtor_Jset_thmss',
-      xtor_rel_thms = dtor_Jrel_thms,
-      xtor_co_iter_thmss = transpose [dtor_unfold_thms, dtor_corec_thms],
-      xtor_co_iter_o_map_thmss = transpose [dtor_unfold_o_Jmap_thms, dtor_corec_o_Jmap_thms],
-      rel_xtor_co_induct_thm = Jrel_coinduct_thm},
-     lthy |> Local_Theory.notes (maybe_conceal_notes (common_notes @ notes @ Jbnf_notes)) |> snd)
-  end;
-
-val _ =
-  Outer_Syntax.local_theory @{command_spec "codatatype"} "define coinductive datatypes"
-    (parse_co_datatype_cmd Greatest_FP construct_gfp);
-
-val option_parser = Parse.group (fn () => "option")
-  ((Parse.reserved "sequential" >> K Sequential_Option)
-  || (Parse.reserved "exhaustive" >> K Exhaustive_Option))
-
-val where_alt_specs_of_parser = Parse.where_ |-- Parse.!!! (Parse.enum1 "|"
-  (Parse_Spec.spec -- Scan.option (Parse.reserved "of" |-- Parse.const)));
-
-val _ = Outer_Syntax.local_theory_to_proof @{command_spec "primcorecursive"}
-  "define primitive corecursive functions"
-  ((Scan.optional (@{keyword "("} |--
-      Parse.!!! (Parse.list1 option_parser) --| @{keyword ")"}) []) --
-    (Parse.fixes -- where_alt_specs_of_parser) >> uncurry add_primcorecursive_cmd);
-
-val _ = Outer_Syntax.local_theory @{command_spec "primcorec"}
-  "define primitive corecursive functions"
-  ((Scan.optional (@{keyword "("} |--
-      Parse.!!! (Parse.list1 option_parser) --| @{keyword ")"}) []) --
-    (Parse.fixes -- where_alt_specs_of_parser) >> uncurry add_primcorec_cmd);
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_gfp_rec_sugar.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1359 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_gfp_rec_sugar.ML
-    Author:     Lorenz Panny, TU Muenchen
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2013
-
-Corecursor sugar.
-*)
-
-signature BNF_GFP_REC_SUGAR =
-sig
-  datatype primcorec_option = Sequential_Option | Exhaustive_Option
-
-  val add_primcorecursive_cmd: primcorec_option list ->
-    (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
-    Proof.context -> Proof.state
-  val add_primcorec_cmd: primcorec_option list ->
-    (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
-    local_theory -> local_theory
-end;
-
-structure BNF_GFP_Rec_Sugar : BNF_GFP_REC_SUGAR =
-struct
-
-open Ctr_Sugar_General_Tactics
-open Ctr_Sugar
-open BNF_Util
-open BNF_Def
-open BNF_FP_Util
-open BNF_FP_Def_Sugar
-open BNF_FP_N2M_Sugar
-open BNF_FP_Rec_Sugar_Util
-open BNF_GFP_Rec_Sugar_Tactics
-
-val codeN = "code"
-val ctrN = "ctr"
-val discN = "disc"
-val disc_iffN = "disc_iff"
-val excludeN = "exclude"
-val selN = "sel"
-
-val nitpicksimp_attrs = @{attributes [nitpick_simp]};
-val simp_attrs = @{attributes [simp]};
-val code_nitpicksimp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs;
-
-exception Primcorec_Error of string * term list;
-
-fun primcorec_error str = raise Primcorec_Error (str, []);
-fun primcorec_error_eqn str eqn = raise Primcorec_Error (str, [eqn]);
-fun primcorec_error_eqns str eqns = raise Primcorec_Error (str, eqns);
-
-datatype primcorec_option = Sequential_Option | Exhaustive_Option;
-
-datatype corec_call =
-  Dummy_No_Corec of int |
-  No_Corec of int |
-  Mutual_Corec of int * int * int |
-  Nested_Corec of int;
-
-type basic_corec_ctr_spec =
-  {ctr: term,
-   disc: term,
-   sels: term list};
-
-type corec_ctr_spec =
-  {ctr: term,
-   disc: term,
-   sels: term list,
-   pred: int option,
-   calls: corec_call list,
-   discI: thm,
-   sel_thms: thm list,
-   disc_excludess: thm list list,
-   collapse: thm,
-   corec_thm: thm,
-   disc_corec: thm,
-   sel_corecs: thm list};
-
-type corec_spec =
-  {corec: term,
-   disc_exhausts: thm list,
-   nested_maps: thm list,
-   nested_map_idents: thm list,
-   nested_map_comps: thm list,
-   ctr_specs: corec_ctr_spec list};
-
-exception AINT_NO_MAP of term;
-
-fun not_codatatype ctxt T =
-  error ("Not a codatatype: " ^ Syntax.string_of_typ ctxt T);
-fun ill_formed_corec_call ctxt t =
-  error ("Ill-formed corecursive call: " ^ quote (Syntax.string_of_term ctxt t));
-fun invalid_map ctxt t =
-  error ("Invalid map function in " ^ quote (Syntax.string_of_term ctxt t));
-fun unexpected_corec_call ctxt t =
-  error ("Unexpected corecursive call: " ^ quote (Syntax.string_of_term ctxt t));
-
-fun order_list_duplicates xs = map snd (sort (int_ord o pairself fst) xs)
-
-val mk_conjs = try (foldr1 HOLogic.mk_conj) #> the_default @{const True};
-val mk_disjs = try (foldr1 HOLogic.mk_disj) #> the_default @{const False};
-val mk_dnf = mk_disjs o map mk_conjs;
-
-val conjuncts_s = filter_out (curry (op aconv) @{const True}) o HOLogic.conjuncts;
-
-fun s_not @{const True} = @{const False}
-  | s_not @{const False} = @{const True}
-  | s_not (@{const Not} $ t) = t
-  | s_not (@{const conj} $ t $ u) = @{const disj} $ s_not t $ s_not u
-  | s_not (@{const disj} $ t $ u) = @{const conj} $ s_not t $ s_not u
-  | s_not t = @{const Not} $ t;
-
-val s_not_conj = conjuncts_s o s_not o mk_conjs;
-
-fun propagate_unit_pos u cs = if member (op aconv) cs u then [@{const False}] else cs;
-
-fun propagate_unit_neg not_u cs = remove (op aconv) not_u cs;
-
-fun propagate_units css =
-  (case List.partition (can the_single) css of
-     ([], _) => css
-   | ([u] :: uss, css') =>
-     [u] :: propagate_units (map (propagate_unit_neg (s_not u))
-       (map (propagate_unit_pos u) (uss @ css'))));
-
-fun s_conjs cs =
-  if member (op aconv) cs @{const False} then @{const False}
-  else mk_conjs (remove (op aconv) @{const True} cs);
-
-fun s_disjs ds =
-  if member (op aconv) ds @{const True} then @{const True}
-  else mk_disjs (remove (op aconv) @{const False} ds);
-
-fun s_dnf css0 =
-  let val css = propagate_units css0 in
-    if null css then
-      [@{const False}]
-    else if exists null css then
-      []
-    else
-      map (fn c :: cs => (c, cs)) css
-      |> AList.coalesce (op =)
-      |> map (fn (c, css) => c :: s_dnf css)
-      |> (fn [cs] => cs | css => [s_disjs (map s_conjs css)])
-  end;
-
-fun fold_rev_let_if_case ctxt f bound_Ts t =
-  let
-    val thy = Proof_Context.theory_of ctxt;
-
-    fun fld conds t =
-      (case Term.strip_comb t of
-        (Const (@{const_name Let}, _), [_, _]) => fld conds (unfold_let t)
-      | (Const (@{const_name If}, _), [cond, then_branch, else_branch]) =>
-        fld (conds @ conjuncts_s cond) then_branch o fld (conds @ s_not_conj [cond]) else_branch
-      | (Const (c, _), args as _ :: _ :: _) =>
-        let val n = num_binder_types (Sign.the_const_type thy c) - 1 in
-          if n >= 0 andalso n < length args then
-            (case fastype_of1 (bound_Ts, nth args n) of
-              Type (s, Ts) =>
-              (case dest_case ctxt s Ts t of
-                SOME (ctr_sugar as {sel_splits = _ :: _, ...}, conds', branches) =>
-                apfst (cons ctr_sugar) o fold_rev (uncurry fld)
-                  (map (append conds o conjuncts_s) conds' ~~ branches)
-              | _ => apsnd (f conds t))
-            | _ => apsnd (f conds t))
-          else
-            apsnd (f conds t)
-        end
-      | _ => apsnd (f conds t))
-  in
-    fld [] t o pair []
-  end;
-
-fun case_of ctxt s =
-  (case ctr_sugar_of ctxt s of
-    SOME {casex = Const (s', _), sel_splits = _ :: _, ...} => SOME s'
-  | _ => NONE);
-
-fun massage_let_if_case ctxt has_call massage_leaf =
-  let
-    val thy = Proof_Context.theory_of ctxt;
-
-    fun check_no_call t = if has_call t then unexpected_corec_call ctxt t else ();
-
-    fun massage_abs bound_Ts 0 t = massage_rec bound_Ts t
-      | massage_abs bound_Ts m (Abs (s, T, t)) = Abs (s, T, massage_abs (T :: bound_Ts) (m - 1) t)
-      | massage_abs bound_Ts m t =
-        let val T = domain_type (fastype_of1 (bound_Ts, t)) in
-          Abs (Name.uu, T, massage_abs (T :: bound_Ts) (m - 1) (incr_boundvars 1 t $ Bound 0))
-        end
-    and massage_rec bound_Ts t =
-      let val typof = curry fastype_of1 bound_Ts in
-        (case Term.strip_comb t of
-          (Const (@{const_name Let}, _), [_, _]) => massage_rec bound_Ts (unfold_let t)
-        | (Const (@{const_name If}, _), obj :: (branches as [_, _])) =>
-          let val branches' = map (massage_rec bound_Ts) branches in
-            Term.list_comb (If_const (typof (hd branches')) $ tap check_no_call obj, branches')
-          end
-        | (Const (c, _), args as _ :: _ :: _) =>
-          (case try strip_fun_type (Sign.the_const_type thy c) of
-            SOME (gen_branch_Ts, gen_body_fun_T) =>
-            let
-              val gen_branch_ms = map num_binder_types gen_branch_Ts;
-              val n = length gen_branch_ms;
-            in
-              if n < length args then
-                (case gen_body_fun_T of
-                  Type (_, [Type (T_name, _), _]) =>
-                  if case_of ctxt T_name = SOME c then
-                    let
-                      val (branches, obj_leftovers) = chop n args;
-                      val branches' = map2 (massage_abs bound_Ts) gen_branch_ms branches;
-                      val branch_Ts' = map typof branches';
-                      val body_T' = snd (strip_typeN (hd gen_branch_ms) (hd branch_Ts'));
-                      val casex' = Const (c, branch_Ts' ---> map typof obj_leftovers ---> body_T');
-                    in
-                      Term.list_comb (casex',
-                        branches' @ tap (List.app check_no_call) obj_leftovers)
-                    end
-                  else
-                    massage_leaf bound_Ts t
-                | _ => massage_leaf bound_Ts t)
-              else
-                massage_leaf bound_Ts t
-            end
-          | NONE => massage_leaf bound_Ts t)
-        | _ => massage_leaf bound_Ts t)
-      end
-  in
-    massage_rec
-  end;
-
-val massage_mutual_corec_call = massage_let_if_case;
-
-fun curried_type (Type (@{type_name fun}, [Type (@{type_name prod}, Ts), T])) = Ts ---> T;
-
-fun massage_nested_corec_call ctxt has_call raw_massage_call bound_Ts U t =
-  let
-    fun check_no_call t = if has_call t then unexpected_corec_call ctxt t else ();
-
-    val build_map_Inl = build_map ctxt (uncurry Inl_const o dest_sumT o snd);
-
-    fun massage_mutual_call bound_Ts U T t =
-      if has_call t then
-        (case try dest_sumT U of
-          SOME (U1, U2) => if U1 = T then raw_massage_call bound_Ts T U2 t else invalid_map ctxt t
-        | NONE => invalid_map ctxt t)
-      else
-        build_map_Inl (T, U) $ t;
-
-    fun massage_mutual_fun bound_Ts U T t =
-      (case t of
-        Const (@{const_name comp}, _) $ t1 $ t2 =>
-        mk_comp bound_Ts (massage_mutual_fun bound_Ts U T t1, tap check_no_call t2)
-      | _ =>
-        let
-          val var = Var ((Name.uu, Term.maxidx_of_term t + 1),
-            domain_type (fastype_of1 (bound_Ts, t)));
-        in
-          Term.lambda var (massage_mutual_call bound_Ts U T (t $ var))
-        end);
-
-    fun massage_map bound_Ts (Type (_, Us)) (Type (s, Ts)) t =
-        (case try (dest_map ctxt s) t of
-          SOME (map0, fs) =>
-          let
-            val Type (_, dom_Ts) = domain_type (fastype_of1 (bound_Ts, t));
-            val map' = mk_map (length fs) dom_Ts Us map0;
-            val fs' =
-              map_flattened_map_args ctxt s (map3 (massage_map_or_map_arg bound_Ts) Us Ts) fs;
-          in
-            Term.list_comb (map', fs')
-          end
-        | NONE => raise AINT_NO_MAP t)
-      | massage_map _ _ _ t = raise AINT_NO_MAP t
-    and massage_map_or_map_arg bound_Ts U T t =
-      if T = U then
-        tap check_no_call t
-      else
-        massage_map bound_Ts U T t
-        handle AINT_NO_MAP _ => massage_mutual_fun bound_Ts U T t;
-
-    fun massage_call bound_Ts U T =
-      massage_let_if_case ctxt has_call (fn bound_Ts => fn t =>
-        if has_call t then
-          (case U of
-            Type (s, Us) =>
-            (case try (dest_ctr ctxt s) t of
-              SOME (f, args) =>
-              let
-                val typof = curry fastype_of1 bound_Ts;
-                val f' = mk_ctr Us f
-                val f'_T = typof f';
-                val arg_Ts = map typof args;
-              in
-                Term.list_comb (f', map3 (massage_call bound_Ts) (binder_types f'_T) arg_Ts args)
-              end
-            | NONE =>
-              (case t of
-                Const (@{const_name prod_case}, _) $ t' =>
-                let
-                  val U' = curried_type U;
-                  val T' = curried_type T;
-                in
-                  Const (@{const_name prod_case}, U' --> U) $ massage_call bound_Ts U' T' t'
-                end
-              | t1 $ t2 =>
-                (if has_call t2 then
-                  massage_mutual_call bound_Ts U T t
-                else
-                  massage_map bound_Ts U T t1 $ t2
-                  handle AINT_NO_MAP _ => massage_mutual_call bound_Ts U T t)
-              | Abs (s, T', t') =>
-                Abs (s, T', massage_call (T' :: bound_Ts) (range_type U) (range_type T) t')
-              | _ => massage_mutual_call bound_Ts U T t))
-          | _ => ill_formed_corec_call ctxt t)
-        else
-          build_map_Inl (T, U) $ t) bound_Ts;
-
-    val T = fastype_of1 (bound_Ts, t);
-  in
-    if has_call t then massage_call bound_Ts U T t else build_map_Inl (T, U) $ t
-  end;
-
-fun expand_to_ctr_term ctxt s Ts t =
-  (case ctr_sugar_of ctxt s of
-    SOME {ctrs, casex, ...} =>
-    Term.list_comb (mk_case Ts (Type (s, Ts)) casex, map (mk_ctr Ts) ctrs) $ t
-  | NONE => raise Fail "expand_to_ctr_term");
-
-fun expand_corec_code_rhs ctxt has_call bound_Ts t =
-  (case fastype_of1 (bound_Ts, t) of
-    Type (s, Ts) =>
-    massage_let_if_case ctxt has_call (fn _ => fn t =>
-      if can (dest_ctr ctxt s) t then t else expand_to_ctr_term ctxt s Ts t) bound_Ts t
-  | _ => raise Fail "expand_corec_code_rhs");
-
-fun massage_corec_code_rhs ctxt massage_ctr =
-  massage_let_if_case ctxt (K false)
-    (fn bound_Ts => uncurry (massage_ctr bound_Ts) o Term.strip_comb);
-
-fun fold_rev_corec_code_rhs ctxt f =
-  snd ooo fold_rev_let_if_case ctxt (fn conds => uncurry (f conds) o Term.strip_comb);
-
-fun case_thms_of_term ctxt bound_Ts t =
-  let val (ctr_sugars, _) = fold_rev_let_if_case ctxt (K (K I)) bound_Ts t () in
-    (maps #distincts ctr_sugars, maps #discIs ctr_sugars, maps #disc_exhausts ctr_sugars,
-     maps #sel_splits ctr_sugars, maps #sel_split_asms ctr_sugars)
-  end;
-
-fun basic_corec_specs_of ctxt res_T =
-  (case res_T of
-    Type (T_name, _) =>
-    (case Ctr_Sugar.ctr_sugar_of ctxt T_name of
-      NONE => not_codatatype ctxt res_T
-    | SOME {ctrs, discs, selss, ...} =>
-      let
-        val thy = Proof_Context.theory_of ctxt;
-
-        val gfpT = body_type (fastype_of (hd ctrs));
-        val As_rho = tvar_subst thy [gfpT] [res_T];
-        val substA = Term.subst_TVars As_rho;
-
-        fun mk_spec ctr disc sels = {ctr = substA ctr, disc = substA disc, sels = map substA sels};
-      in
-        map3 mk_spec ctrs discs selss
-        handle ListPair.UnequalLengths => not_codatatype ctxt res_T
-      end)
-  | _ => not_codatatype ctxt res_T);
-
-fun map_thms_of_typ ctxt (Type (s, _)) =
-    (case fp_sugar_of ctxt s of
-      SOME {index, mapss, ...} => nth mapss index
-    | NONE => [])
-  | map_thms_of_typ _ _ = [];
-
-fun corec_specs_of bs arg_Ts res_Ts get_indices callssss0 lthy0 =
-  let
-    val thy = Proof_Context.theory_of lthy0;
-
-    val ((missing_res_Ts, perm0_kks,
-          fp_sugars as {nested_bnfs, fp_res = {xtor_co_iterss = dtor_coiters1 :: _, ...},
-            co_inducts = coinduct_thms, ...} :: _, (_, gfp_sugar_thms)), lthy) =
-      nested_to_mutual_fps Greatest_FP bs res_Ts get_indices callssss0 lthy0;
-
-    val perm_fp_sugars = sort (int_ord o pairself #index) fp_sugars;
-
-    val indices = map #index fp_sugars;
-    val perm_indices = map #index perm_fp_sugars;
-
-    val perm_ctrss = map (#ctrs o of_fp_sugar #ctr_sugars) perm_fp_sugars;
-    val perm_ctr_Tsss = map (map (binder_types o fastype_of)) perm_ctrss;
-    val perm_gfpTs = map (body_type o fastype_of o hd) perm_ctrss;
-
-    val nn0 = length res_Ts;
-    val nn = length perm_gfpTs;
-    val kks = 0 upto nn - 1;
-    val perm_ns = map length perm_ctr_Tsss;
-
-    val perm_Cs = map (domain_type o body_fun_type o fastype_of o co_rec_of o
-      of_fp_sugar (#xtor_co_iterss o #fp_res)) perm_fp_sugars;
-    val (perm_p_Tss, (perm_q_Tssss, _, perm_f_Tssss, _)) =
-      mk_coiter_fun_arg_types perm_ctr_Tsss perm_Cs perm_ns (co_rec_of dtor_coiters1);
-
-    val (perm_p_hss, h) = indexedd perm_p_Tss 0;
-    val (perm_q_hssss, h') = indexedddd perm_q_Tssss h;
-    val (perm_f_hssss, _) = indexedddd perm_f_Tssss h';
-
-    val fun_arg_hs =
-      flat (map3 flat_corec_preds_predsss_gettersss perm_p_hss perm_q_hssss perm_f_hssss);
-
-    fun unpermute0 perm0_xs = permute_like (op =) perm0_kks kks perm0_xs;
-    fun unpermute perm_xs = permute_like (op =) perm_indices indices perm_xs;
-
-    val coinduct_thmss = map (unpermute0 o conj_dests nn) coinduct_thms;
-
-    val p_iss = map (map (find_index_eq fun_arg_hs)) (unpermute perm_p_hss);
-    val q_issss = map (map (map (map (find_index_eq fun_arg_hs)))) (unpermute perm_q_hssss);
-    val f_issss = map (map (map (map (find_index_eq fun_arg_hs)))) (unpermute perm_f_hssss);
-
-    val f_Tssss = unpermute perm_f_Tssss;
-    val gfpTs = unpermute perm_gfpTs;
-    val Cs = unpermute perm_Cs;
-
-    val As_rho = tvar_subst thy (take nn0 gfpTs) res_Ts;
-    val Cs_rho = map (fst o dest_TVar) Cs ~~ pad_list HOLogic.unitT nn arg_Ts;
-
-    val substA = Term.subst_TVars As_rho;
-    val substAT = Term.typ_subst_TVars As_rho;
-    val substCT = Term.typ_subst_TVars Cs_rho;
-
-    val perm_Cs' = map substCT perm_Cs;
-
-    fun call_of nullary [] [g_i] [Type (@{type_name fun}, [_, T])] =
-        (if exists_subtype_in Cs T then Nested_Corec
-         else if nullary then Dummy_No_Corec
-         else No_Corec) g_i
-      | call_of _ [q_i] [g_i, g_i'] _ = Mutual_Corec (q_i, g_i, g_i');
-
-    fun mk_ctr_spec ctr disc sels p_io q_iss f_iss f_Tss discI sel_thms disc_excludess collapse
-        corec_thm disc_corec sel_corecs =
-      let val nullary = not (can dest_funT (fastype_of ctr)) in
-        {ctr = substA ctr, disc = substA disc, sels = map substA sels, pred = p_io,
-         calls = map3 (call_of nullary) q_iss f_iss f_Tss, discI = discI, sel_thms = sel_thms,
-         disc_excludess = disc_excludess, collapse = collapse, corec_thm = corec_thm,
-         disc_corec = disc_corec, sel_corecs = sel_corecs}
-      end;
-
-    fun mk_ctr_specs index ctr_sugars p_is q_isss f_isss f_Tsss coiter_thmsss disc_coitersss
-        sel_coiterssss =
-      let
-        val {ctrs, discs, selss, discIs, sel_thmss, disc_excludesss, collapses, ...} : ctr_sugar =
-          nth ctr_sugars index;
-        val p_ios = map SOME p_is @ [NONE];
-        val discIs = #discIs (nth ctr_sugars index);
-        val corec_thms = co_rec_of (nth coiter_thmsss index);
-        val disc_corecs = co_rec_of (nth disc_coitersss index);
-        val sel_corecss = co_rec_of (nth sel_coiterssss index);
-      in
-        map14 mk_ctr_spec ctrs discs selss p_ios q_isss f_isss f_Tsss discIs sel_thmss
-          disc_excludesss collapses corec_thms disc_corecs sel_corecss
-      end;
-
-    fun mk_spec ({T, index, ctr_sugars, co_iterss = coiterss, co_iter_thmsss = coiter_thmsss,
-          disc_co_itersss = disc_coitersss, sel_co_iterssss = sel_coiterssss, ...} : fp_sugar)
-        p_is q_isss f_isss f_Tsss =
-      {corec = mk_co_iter thy Greatest_FP (substAT T) perm_Cs' (co_rec_of (nth coiterss index)),
-       disc_exhausts = #disc_exhausts (nth ctr_sugars index),
-       nested_maps = maps (map_thms_of_typ lthy o T_of_bnf) nested_bnfs,
-       nested_map_idents = map (unfold_thms lthy @{thms id_def} o map_id0_of_bnf) nested_bnfs,
-       nested_map_comps = map map_comp_of_bnf nested_bnfs,
-       ctr_specs = mk_ctr_specs index ctr_sugars p_is q_isss f_isss f_Tsss coiter_thmsss
-         disc_coitersss sel_coiterssss};
-  in
-    ((is_some gfp_sugar_thms, map5 mk_spec fp_sugars p_iss q_issss f_issss f_Tssss, missing_res_Ts,
-      co_induct_of coinduct_thms, strong_co_induct_of coinduct_thms, co_induct_of coinduct_thmss,
-      strong_co_induct_of coinduct_thmss), lthy)
-  end;
-
-val undef_const = Const (@{const_name undefined}, dummyT);
-
-val abs_tuple = HOLogic.tupled_lambda o HOLogic.mk_tuple;
-
-fun abstract vs =
-  let fun a n (t $ u) = a n t $ a n u
-        | a n (Abs (v, T, b)) = Abs (v, T, a (n + 1) b)
-        | a n t = let val idx = find_index (curry (op =) t) vs in
-            if idx < 0 then t else Bound (n + idx) end
-  in a 0 end;
-
-fun mk_prod1 bound_Ts (t, u) =
-  HOLogic.pair_const (fastype_of1 (bound_Ts, t)) (fastype_of1 (bound_Ts, u)) $ t $ u;
-fun mk_tuple1 bound_Ts = the_default HOLogic.unit o try (foldr1 (mk_prod1 bound_Ts));
-
-type coeqn_data_disc = {
-  fun_name: string,
-  fun_T: typ,
-  fun_args: term list,
-  ctr: term,
-  ctr_no: int,
-  disc: term,
-  prems: term list,
-  auto_gen: bool,
-  ctr_rhs_opt: term option,
-  code_rhs_opt: term option,
-  eqn_pos: int,
-  user_eqn: term
-};
-
-type coeqn_data_sel = {
-  fun_name: string,
-  fun_T: typ,
-  fun_args: term list,
-  ctr: term,
-  sel: term,
-  rhs_term: term,
-  ctr_rhs_opt: term option,
-  code_rhs_opt: term option,
-  eqn_pos: int,
-  user_eqn: term
-};
-
-datatype coeqn_data =
-  Disc of coeqn_data_disc |
-  Sel of coeqn_data_sel;
-
-fun dissect_coeqn_disc fun_names sequentials (basic_ctr_specss : basic_corec_ctr_spec list list)
-    eqn_pos ctr_rhs_opt code_rhs_opt prems' concl matchedsss =
-  let
-    fun find_subterm p =
-      let (* FIXME \<exists>? *)
-        fun find (t as u $ v) = if p t then SOME t else merge_options (find u, find v)
-          | find t = if p t then SOME t else NONE;
-      in find end;
-
-    val applied_fun = concl
-      |> find_subterm (member (op = o apsnd SOME) fun_names o try (fst o dest_Free o head_of))
-      |> the
-      handle Option.Option => primcorec_error_eqn "malformed discriminator formula" concl;
-    val ((fun_name, fun_T), fun_args) = strip_comb applied_fun |>> dest_Free;
-    val SOME (sequential, basic_ctr_specs) =
-      AList.lookup (op =) (fun_names ~~ (sequentials ~~ basic_ctr_specss)) fun_name;
-
-    val discs = map #disc basic_ctr_specs;
-    val ctrs = map #ctr basic_ctr_specs;
-    val not_disc = head_of concl = @{term Not};
-    val _ = not_disc andalso length ctrs <> 2 andalso
-      primcorec_error_eqn "negated discriminator for a type with \<noteq> 2 constructors" concl;
-    val disc' = find_subterm (member (op =) discs o head_of) concl;
-    val eq_ctr0 = concl |> perhaps (try HOLogic.dest_not) |> try (HOLogic.dest_eq #> snd)
-        |> (fn SOME t => let val n = find_index (curry (op =) t) ctrs in
-          if n >= 0 then SOME n else NONE end | _ => NONE);
-    val _ = is_some disc' orelse is_some eq_ctr0 orelse
-      primcorec_error_eqn "no discriminator in equation" concl;
-    val ctr_no' =
-      if is_none disc' then the eq_ctr0 else find_index (curry (op =) (head_of (the disc'))) discs;
-    val ctr_no = if not_disc then 1 - ctr_no' else ctr_no';
-    val {ctr, disc, ...} = nth basic_ctr_specs ctr_no;
-
-    val catch_all = try (fst o dest_Free o the_single) prems' = SOME Name.uu_;
-    val matchedss = AList.lookup (op =) matchedsss fun_name |> the_default [];
-    val prems = map (abstract (List.rev fun_args)) prems';
-    val actual_prems =
-      (if catch_all orelse sequential then maps s_not_conj matchedss else []) @
-      (if catch_all then [] else prems);
-
-    val matchedsss' = AList.delete (op =) fun_name matchedsss
-      |> cons (fun_name, if sequential then matchedss @ [prems] else matchedss @ [actual_prems]);
-
-    val user_eqn =
-      (actual_prems, concl)
-      |>> map HOLogic.mk_Trueprop ||> HOLogic.mk_Trueprop o abstract (List.rev fun_args)
-      |> curry Logic.list_all (map dest_Free fun_args) o Logic.list_implies;
-  in
-    (Disc {
-      fun_name = fun_name,
-      fun_T = fun_T,
-      fun_args = fun_args,
-      ctr = ctr,
-      ctr_no = ctr_no,
-      disc = disc,
-      prems = actual_prems,
-      auto_gen = catch_all,
-      ctr_rhs_opt = ctr_rhs_opt,
-      code_rhs_opt = code_rhs_opt,
-      eqn_pos = eqn_pos,
-      user_eqn = user_eqn
-    }, matchedsss')
-  end;
-
-fun dissect_coeqn_sel fun_names (basic_ctr_specss : basic_corec_ctr_spec list list) eqn_pos
-    ctr_rhs_opt code_rhs_opt eqn0 of_spec_opt eqn =
-  let
-    val (lhs, rhs) = HOLogic.dest_eq eqn
-      handle TERM _ =>
-        primcorec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn;
-    val sel = head_of lhs;
-    val ((fun_name, fun_T), fun_args) = dest_comb lhs |> snd |> strip_comb |> apfst dest_Free
-      handle TERM _ =>
-        primcorec_error_eqn "malformed selector argument in left-hand side" eqn;
-    val basic_ctr_specs = the (AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name)
-      handle Option.Option =>
-        primcorec_error_eqn "malformed selector argument in left-hand side" eqn;
-    val {ctr, ...} =
-      (case of_spec_opt of
-        SOME of_spec => the (find_first (curry (op =) of_spec o #ctr) basic_ctr_specs)
-      | NONE => filter (exists (curry (op =) sel) o #sels) basic_ctr_specs |> the_single
-          handle List.Empty => primcorec_error_eqn "ambiguous selector - use \"of\"" eqn);
-    val user_eqn = drop_all eqn0;
-  in
-    Sel {
-      fun_name = fun_name,
-      fun_T = fun_T,
-      fun_args = fun_args,
-      ctr = ctr,
-      sel = sel,
-      rhs_term = rhs,
-      ctr_rhs_opt = ctr_rhs_opt,
-      code_rhs_opt = code_rhs_opt,
-      eqn_pos = eqn_pos,
-      user_eqn = user_eqn
-    }
-  end;
-
-fun dissect_coeqn_ctr fun_names sequentials (basic_ctr_specss : basic_corec_ctr_spec list list)
-    eqn_pos eqn0 code_rhs_opt prems concl matchedsss =
-  let
-    val (lhs, rhs) = HOLogic.dest_eq concl;
-    val (fun_name, fun_args) = strip_comb lhs |>> fst o dest_Free;
-    val SOME basic_ctr_specs = AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name;
-    val (ctr, ctr_args) = strip_comb (unfold_let rhs);
-    val {disc, sels, ...} = the (find_first (curry (op =) ctr o #ctr) basic_ctr_specs)
-      handle Option.Option => primcorec_error_eqn "not a constructor" ctr;
-
-    val disc_concl = betapply (disc, lhs);
-    val (eqn_data_disc_opt, matchedsss') =
-      if null (tl basic_ctr_specs) then
-        (NONE, matchedsss)
-      else
-        apfst SOME (dissect_coeqn_disc fun_names sequentials basic_ctr_specss eqn_pos
-          (SOME (abstract (List.rev fun_args) rhs)) code_rhs_opt prems disc_concl matchedsss);
-
-    val sel_concls = sels ~~ ctr_args
-      |> map (fn (sel, ctr_arg) => HOLogic.mk_eq (betapply (sel, lhs), ctr_arg));
-
-(*
-val _ = tracing ("reduced\n    " ^ Syntax.string_of_term @{context} concl ^ "\nto\n    \<cdot> " ^
- (is_some eqn_data_disc_opt ? K (Syntax.string_of_term @{context} disc_concl ^ "\n    \<cdot> ")) "" ^
- space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) sel_concls) ^
- "\nfor premise(s)\n    \<cdot> " ^
- space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) prems));
-*)
-
-    val eqns_data_sel =
-      map (dissect_coeqn_sel fun_names basic_ctr_specss eqn_pos
-        (SOME (abstract (List.rev fun_args) rhs)) code_rhs_opt eqn0 (SOME ctr)) sel_concls;
-  in
-    (the_list eqn_data_disc_opt @ eqns_data_sel, matchedsss')
-  end;
-
-fun dissect_coeqn_code lthy has_call fun_names basic_ctr_specss eqn_pos eqn0 concl matchedsss =
-  let
-    val (lhs, (rhs', rhs)) = HOLogic.dest_eq concl ||> `(expand_corec_code_rhs lthy has_call []);
-    val (fun_name, fun_args) = strip_comb lhs |>> fst o dest_Free;
-    val SOME basic_ctr_specs = AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name;
-
-    val cond_ctrs = fold_rev_corec_code_rhs lthy (fn cs => fn ctr => fn _ =>
-        if member (op = o apsnd #ctr) basic_ctr_specs ctr then cons (ctr, cs)
-        else primcorec_error_eqn "not a constructor" ctr) [] rhs' []
-      |> AList.group (op =);
-
-    val ctr_premss = (case cond_ctrs of [_] => [[]] | _ => map (s_dnf o snd) cond_ctrs);
-    val ctr_concls = cond_ctrs |> map (fn (ctr, _) =>
-        binder_types (fastype_of ctr)
-        |> map_index (fn (n, T) => massage_corec_code_rhs lthy (fn _ => fn ctr' => fn args =>
-          if ctr' = ctr then nth args n else Const (@{const_name undefined}, T)) [] rhs')
-        |> curry list_comb ctr
-        |> curry HOLogic.mk_eq lhs);
-
-    val sequentials = replicate (length fun_names) false;
-  in
-    fold_map2 (dissect_coeqn_ctr fun_names sequentials basic_ctr_specss eqn_pos eqn0
-        (SOME (abstract (List.rev fun_args) rhs)))
-      ctr_premss ctr_concls matchedsss
-  end;
-
-fun dissect_coeqn lthy has_call fun_names sequentials
-    (basic_ctr_specss : basic_corec_ctr_spec list list) (eqn_pos, eqn0) of_spec_opt matchedsss =
-  let
-    val eqn = drop_all eqn0
-      handle TERM _ => primcorec_error_eqn "malformed function equation" eqn0;
-    val (prems, concl) = Logic.strip_horn eqn
-      |> apfst (map HOLogic.dest_Trueprop) o apsnd HOLogic.dest_Trueprop;
-
-    val head = concl
-      |> perhaps (try HOLogic.dest_not) |> perhaps (try (fst o HOLogic.dest_eq))
-      |> head_of;
-
-    val rhs_opt = concl |> perhaps (try HOLogic.dest_not) |> try (snd o HOLogic.dest_eq);
-
-    val discs = maps (map #disc) basic_ctr_specss;
-    val sels = maps (maps #sels) basic_ctr_specss;
-    val ctrs = maps (map #ctr) basic_ctr_specss;
-  in
-    if member (op =) discs head orelse
-        is_some rhs_opt andalso
-          member (op =) (filter (null o binder_types o fastype_of) ctrs) (the rhs_opt) then
-      dissect_coeqn_disc fun_names sequentials basic_ctr_specss eqn_pos NONE NONE prems concl
-        matchedsss
-      |>> single
-    else if member (op =) sels head then
-      ([dissect_coeqn_sel fun_names basic_ctr_specss eqn_pos NONE NONE eqn0 of_spec_opt concl],
-       matchedsss)
-    else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) then
-      if member (op =) ctrs (head_of (unfold_let (the rhs_opt))) then
-        dissect_coeqn_ctr fun_names sequentials basic_ctr_specss eqn_pos eqn0
-          (if null prems then
-             SOME (snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (Logic.strip_assums_concl eqn0))))
-           else
-             NONE)
-          prems concl matchedsss
-      else if null prems then
-        dissect_coeqn_code lthy has_call fun_names basic_ctr_specss eqn_pos eqn0 concl matchedsss
-        |>> flat
-      else
-        primcorec_error_eqn "cannot mix constructor and code views (see manual for details)" eqn
-    else
-      primcorec_error_eqn "malformed function equation" eqn
-  end;
-
-fun build_corec_arg_disc (ctr_specs : corec_ctr_spec list)
-    ({fun_args, ctr_no, prems, ...} : coeqn_data_disc) =
-  if is_none (#pred (nth ctr_specs ctr_no)) then I else
-    s_conjs prems
-    |> curry subst_bounds (List.rev fun_args)
-    |> HOLogic.tupled_lambda (HOLogic.mk_tuple fun_args)
-    |> K |> nth_map (the (#pred (nth ctr_specs ctr_no)));
-
-fun build_corec_arg_no_call (sel_eqns : coeqn_data_sel list) sel =
-  find_first (curry (op =) sel o #sel) sel_eqns
-  |> try (fn SOME {fun_args, rhs_term, ...} => abs_tuple fun_args rhs_term)
-  |> the_default undef_const
-  |> K;
-
-fun build_corec_args_mutual_call lthy has_call (sel_eqns : coeqn_data_sel list) sel =
-  (case find_first (curry (op =) sel o #sel) sel_eqns of
-    NONE => (I, I, I)
-  | SOME {fun_args, rhs_term, ... } =>
-    let
-      val bound_Ts = List.rev (map fastype_of fun_args);
-      fun rewrite_stop _ t = if has_call t then @{term False} else @{term True};
-      fun rewrite_end _ t = if has_call t then undef_const else t;
-      fun rewrite_cont bound_Ts t =
-        if has_call t then mk_tuple1 bound_Ts (snd (strip_comb t)) else undef_const;
-      fun massage f _ = massage_mutual_corec_call lthy has_call f bound_Ts rhs_term
-        |> abs_tuple fun_args;
-    in
-      (massage rewrite_stop, massage rewrite_end, massage rewrite_cont)
-    end);
-
-fun build_corec_arg_nested_call lthy has_call (sel_eqns : coeqn_data_sel list) sel =
-  (case find_first (curry (op =) sel o #sel) sel_eqns of
-    NONE => I
-  | SOME {fun_args, rhs_term, ...} =>
-    let
-      val bound_Ts = List.rev (map fastype_of fun_args);
-      fun rewrite bound_Ts U T (Abs (v, V, b)) = Abs (v, V, rewrite (V :: bound_Ts) U T b)
-        | rewrite bound_Ts U T (t as _ $ _) =
-          let val (u, vs) = strip_comb t in
-            if is_Free u andalso has_call u then
-              Inr_const U T $ mk_tuple1 bound_Ts vs
-            else if try (fst o dest_Const) u = SOME @{const_name prod_case} then
-              map (rewrite bound_Ts U T) vs |> chop 1 |>> HOLogic.mk_split o the_single |> list_comb
-            else
-              list_comb (rewrite bound_Ts U T u, map (rewrite bound_Ts U T) vs)
-          end
-        | rewrite _ U T t =
-          if is_Free t andalso has_call t then Inr_const U T $ HOLogic.unit else t;
-      fun massage t =
-        rhs_term
-        |> massage_nested_corec_call lthy has_call rewrite bound_Ts (range_type (fastype_of t))
-        |> abs_tuple fun_args;
-    in
-      massage
-    end);
-
-fun build_corec_args_sel lthy has_call (all_sel_eqns : coeqn_data_sel list)
-    (ctr_spec : corec_ctr_spec) =
-  (case filter (curry (op =) (#ctr ctr_spec) o #ctr) all_sel_eqns of
-    [] => I
-  | sel_eqns =>
-    let
-      val sel_call_list = #sels ctr_spec ~~ #calls ctr_spec;
-      val no_calls' = map_filter (try (apsnd (fn No_Corec n => n))) sel_call_list;
-      val mutual_calls' = map_filter (try (apsnd (fn Mutual_Corec n => n))) sel_call_list;
-      val nested_calls' = map_filter (try (apsnd (fn Nested_Corec n => n))) sel_call_list;
-    in
-      I
-      #> fold (fn (sel, n) => nth_map n (build_corec_arg_no_call sel_eqns sel)) no_calls'
-      #> fold (fn (sel, (q, g, h)) =>
-        let val (fq, fg, fh) = build_corec_args_mutual_call lthy has_call sel_eqns sel in
-          nth_map q fq o nth_map g fg o nth_map h fh end) mutual_calls'
-      #> fold (fn (sel, n) => nth_map n
-        (build_corec_arg_nested_call lthy has_call sel_eqns sel)) nested_calls'
-    end);
-
-fun build_codefs lthy bs mxs has_call arg_Tss (corec_specs : corec_spec list)
-    (disc_eqnss : coeqn_data_disc list list) (sel_eqnss : coeqn_data_sel list list) =
-  let
-    val corecs = map #corec corec_specs;
-    val ctr_specss = map #ctr_specs corec_specs;
-    val corec_args = hd corecs
-      |> fst o split_last o binder_types o fastype_of
-      |> map (fn T => if range_type T = @{typ bool}
-          then Abs (Name.uu_, domain_type T, @{term False})
-          else Const (@{const_name undefined}, T))
-      |> fold2 (fold o build_corec_arg_disc) ctr_specss disc_eqnss
-      |> fold2 (fold o build_corec_args_sel lthy has_call) sel_eqnss ctr_specss;
-    fun currys [] t = t
-      | currys Ts t = t $ mk_tuple1 (List.rev Ts) (map Bound (length Ts - 1 downto 0))
-          |> fold_rev (Term.abs o pair Name.uu) Ts;
-
-(*
-val _ = tracing ("corecursor arguments:\n    \<cdot> " ^
- space_implode "\n    \<cdot> " (map (Syntax.string_of_term lthy) corec_args));
-*)
-
-    val excludess' =
-      disc_eqnss
-      |> map (map (fn x => (#fun_args x, #ctr_no x, #prems x, #auto_gen x))
-        #> fst o (fn xs => fold_map (fn x => fn ys => ((x, ys), ys @ [x])) xs [])
-        #> maps (uncurry (map o pair)
-          #> map (fn ((fun_args, c, x, a), (_, c', y, a')) =>
-              ((c, c', a orelse a'), (x, s_not (s_conjs y)))
-            ||> apfst (map HOLogic.mk_Trueprop) o apsnd HOLogic.mk_Trueprop
-            ||> Logic.list_implies
-            ||> curry Logic.list_all (map dest_Free fun_args))))
-  in
-    map (list_comb o rpair corec_args) corecs
-    |> map2 (fn Ts => fn t => if length Ts = 0 then t $ HOLogic.unit else t) arg_Tss
-    |> map2 currys arg_Tss
-    |> Syntax.check_terms lthy
-    |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.conceal (Thm.def_binding b), []), t)))
-      bs mxs
-    |> rpair excludess'
-  end;
-
-fun mk_actual_disc_eqns fun_binding arg_Ts exhaustive ({ctr_specs, ...} : corec_spec)
-    (sel_eqns : coeqn_data_sel list) (disc_eqns : coeqn_data_disc list) =
-  let val num_disc_eqns = length disc_eqns in
-    if (exhaustive andalso num_disc_eqns <> 0) orelse num_disc_eqns <> length ctr_specs - 1 then
-      disc_eqns
-    else
-      let
-        val n = 0 upto length ctr_specs
-          |> the o find_first (fn idx => not (exists (curry (op =) idx o #ctr_no) disc_eqns));
-        val {ctr, disc, ...} = nth ctr_specs n;
-        val fun_args = (try (#fun_args o hd) disc_eqns, try (#fun_args o hd) sel_eqns)
-          |> the_default (map (curry Free Name.uu) arg_Ts) o merge_options;
-        val sel_eqn_opt = find_first (equal ctr o #ctr) sel_eqns;
-        val extra_disc_eqn = {
-          fun_name = Binding.name_of fun_binding,
-          fun_T = arg_Ts ---> body_type (fastype_of (#ctr (hd ctr_specs))),
-          fun_args = fun_args,
-          ctr = ctr,
-          ctr_no = n,
-          disc = disc,
-          prems = maps (s_not_conj o #prems) disc_eqns,
-          auto_gen = true,
-          ctr_rhs_opt = Option.map #ctr_rhs_opt sel_eqn_opt |> the_default NONE,
-          code_rhs_opt = Option.map #ctr_rhs_opt sel_eqn_opt |> the_default NONE,
-          eqn_pos = Option.map (curry (op +) 1 o #eqn_pos) sel_eqn_opt |> the_default 100000 (* FIXME *),
-          user_eqn = undef_const};
-      in
-        chop n disc_eqns ||> cons extra_disc_eqn |> (op @)
-      end
-  end;
-
-fun find_corec_calls ctxt has_call basic_ctr_specs ({ctr, sel, rhs_term, ...} : coeqn_data_sel) =
-  let
-    val sel_no = find_first (curry (op =) ctr o #ctr) basic_ctr_specs
-      |> find_index (curry (op =) sel) o #sels o the;
-    fun find t = if has_call t then snd (fold_rev_let_if_case ctxt (K cons) [] t []) else [];
-  in
-    find rhs_term
-    |> K |> nth_map sel_no |> AList.map_entry (op =) ctr
-  end;
-
-fun applied_fun_of fun_name fun_T fun_args =
-  list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0));
-
-fun is_trivial_implies thm =
-  uncurry (member (op aconv)) (Logic.strip_horn (Thm.prop_of thm));
-
-fun add_primcorec_ursive auto opts fixes specs of_specs_opt lthy =
-  let
-    val thy = Proof_Context.theory_of lthy;
-
-    val (bs, mxs) = map_split (apfst fst) fixes;
-    val (arg_Ts, res_Ts) = map (strip_type o snd o fst #>> HOLogic.mk_tupleT) fixes |> split_list;
-
-    val _ = (case filter_out (fn (_, T) => Sign.of_sort thy (T, HOLogic.typeS)) (bs ~~ arg_Ts) of
-        [] => ()
-      | (b, _) :: _ => primcorec_error ("type of " ^ Binding.print b ^ " contains top sort"));
-
-    val actual_nn = length bs;
-
-    val sequentials = replicate actual_nn (member (op =) opts Sequential_Option);
-    val exhaustives = replicate actual_nn (member (op =) opts Exhaustive_Option);
-
-    val fun_names = map Binding.name_of bs;
-    val basic_ctr_specss = map (basic_corec_specs_of lthy) res_Ts;
-    val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
-    val eqns_data =
-      fold_map2 (dissect_coeqn lthy has_call fun_names sequentials basic_ctr_specss) (tag_list 0 (map snd specs))
-        of_specs_opt []
-      |> flat o fst;
-
-    val callssss =
-      map_filter (try (fn Sel x => x)) eqns_data
-      |> partition_eq (op = o pairself #fun_name)
-      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
-      |> map (flat o snd)
-      |> map2 (fold o find_corec_calls lthy has_call) basic_ctr_specss
-      |> map2 (curry (op |>)) (map (map (fn {ctr, sels, ...} =>
-        (ctr, map (K []) sels))) basic_ctr_specss);
-
-(*
-val _ = tracing ("callssss = " ^ @{make_string} callssss);
-*)
-
-    val ((n2m, corec_specs', _, coinduct_thm, strong_coinduct_thm, coinduct_thms,
-          strong_coinduct_thms), lthy') =
-      corec_specs_of bs arg_Ts res_Ts (get_indices fixes) callssss lthy;
-    val corec_specs = take actual_nn corec_specs';
-    val ctr_specss = map #ctr_specs corec_specs;
-
-    val disc_eqnss' = map_filter (try (fn Disc x => x)) eqns_data
-      |> partition_eq (op = o pairself #fun_name)
-      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
-      |> map (sort (op < o pairself #ctr_no |> make_ord) o flat o snd);
-    val _ = disc_eqnss' |> map (fn x =>
-      let val d = duplicates (op = o pairself #ctr_no) x in null d orelse
-        primcorec_error_eqns "excess discriminator formula in definition"
-          (maps (fn t => filter (curry (op =) (#ctr_no t) o #ctr_no) x) d |> map #user_eqn) end);
-
-    val sel_eqnss = map_filter (try (fn Sel x => x)) eqns_data
-      |> partition_eq (op = o pairself #fun_name)
-      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
-      |> map (flat o snd);
-
-    val arg_Tss = map (binder_types o snd o fst) fixes;
-    val disc_eqnss = map6 mk_actual_disc_eqns bs arg_Tss exhaustives corec_specs sel_eqnss
-      disc_eqnss';
-    val (defs, excludess') =
-      build_codefs lthy' bs mxs has_call arg_Tss corec_specs disc_eqnss sel_eqnss;
-
-    val tac_opts =
-      map (fn {code_rhs_opt, ...} :: _ =>
-        if auto orelse is_some code_rhs_opt then SOME (auto_tac o #context) else NONE) disc_eqnss;
-
-    fun exclude_tac tac_opt sequential (c, c', a) =
-      if a orelse c = c' orelse sequential then
-        SOME (K (HEADGOAL (mk_primcorec_assumption_tac lthy [])))
-      else
-        tac_opt;
-
-(*
-val _ = tracing ("exclusiveness properties:\n    \<cdot> " ^
- space_implode "\n    \<cdot> " (maps (map (Syntax.string_of_term lthy o snd)) excludess'));
-*)
-
-    val excludess'' = map3 (fn tac_opt => fn sequential => map (fn (idx, goal) =>
-        (idx, (Option.map (Goal.prove_sorry lthy [] [] goal #> Thm.close_derivation)
-           (exclude_tac tac_opt sequential idx), goal))))
-      tac_opts sequentials excludess';
-    val taut_thmss = map (map (apsnd (the o fst)) o filter (is_some o fst o snd)) excludess'';
-    val (goal_idxss, exclude_goalss) = excludess''
-      |> map (map (apsnd (rpair [] o snd)) o filter (is_none o fst o snd))
-      |> split_list o map split_list;
-
-    fun list_all_fun_args extras =
-      map2 (fn [] => I
-          | {fun_args, ...} :: _ => map (curry Logic.list_all (extras @ map dest_Free fun_args)))
-        disc_eqnss;
-
-    val syntactic_exhaustives =
-      map (fn disc_eqns => forall (null o #prems orf is_some o #code_rhs_opt) disc_eqns
-          orelse exists #auto_gen disc_eqns)
-        disc_eqnss;
-    val de_facto_exhaustives =
-      map2 (fn b => fn b' => b orelse b') exhaustives syntactic_exhaustives;
-
-    val nchotomy_goalss =
-      map2 (fn false => K [] | true => single o HOLogic.mk_Trueprop o mk_dnf o map #prems)
-        de_facto_exhaustives disc_eqnss
-      |> list_all_fun_args []
-    val nchotomy_taut_thmss =
-      map6 (fn tac_opt => fn {disc_exhausts = res_disc_exhausts, ...} => fn arg_Ts =>
-          fn {code_rhs_opt, ...} :: _ => fn [] => K []
-            | [goal] => fn true =>
-              let
-                val (_, _, arg_disc_exhausts, _, _) =
-                  case_thms_of_term lthy arg_Ts (the_default Term.dummy code_rhs_opt);
-              in
-                [Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
-                   mk_primcorec_nchotomy_tac ctxt (res_disc_exhausts @ arg_disc_exhausts))
-                 |> Thm.close_derivation]
-              end
-            | false =>
-              (case tac_opt of
-                SOME tac => [Goal.prove_sorry lthy [] [] goal tac |> Thm.close_derivation]
-              | NONE => []))
-        tac_opts corec_specs arg_Tss disc_eqnss nchotomy_goalss syntactic_exhaustives;
-
-    val syntactic_exhaustives =
-      map (fn disc_eqns => forall (null o #prems orf is_some o #code_rhs_opt) disc_eqns
-          orelse exists #auto_gen disc_eqns)
-        disc_eqnss;
-
-    val nchotomy_goalss =
-      map2 (fn (NONE, false) => map (rpair []) | _ => K []) (tac_opts ~~ syntactic_exhaustives)
-        nchotomy_goalss;
-
-    val goalss = nchotomy_goalss @ exclude_goalss;
-
-    fun prove thmss'' def_thms' lthy =
-      let
-        val def_thms = map (snd o snd) def_thms';
-
-        val (nchotomy_thmss, exclude_thmss) =
-          (map2 append (take actual_nn thmss'') nchotomy_taut_thmss, drop actual_nn thmss'');
-
-        val ps =
-          Variable.variant_frees lthy (maps (maps #fun_args) disc_eqnss) [("P", HOLogic.boolT)];
-
-        val exhaust_thmss =
-          map2 (fn false => K []
-              | true => fn disc_eqns as {fun_args, ...} :: _ =>
-                let
-                  val p = Bound (length fun_args);
-                  fun mk_imp_p Qs = Logic.list_implies (Qs, HOLogic.mk_Trueprop p);
-                in
-                  [mk_imp_p (map (mk_imp_p o map HOLogic.mk_Trueprop o #prems) disc_eqns)]
-                end)
-            de_facto_exhaustives disc_eqnss
-          |> list_all_fun_args ps
-          |> map3 (fn disc_eqns as {fun_args, ...} :: _ => fn [] => K []
-              | [nchotomy_thm] => fn [goal] =>
-                [mk_primcorec_exhaust_tac lthy ("" (* for "P" *) :: map (fst o dest_Free) fun_args)
-                   (length disc_eqns) nchotomy_thm
-                 |> K |> Goal.prove_sorry lthy [] [] goal
-                 |> Thm.close_derivation])
-            disc_eqnss nchotomy_thmss;
-        val nontriv_exhaust_thmss = map (filter_out is_trivial_implies) exhaust_thmss;
-
-        val excludess' = map (op ~~) (goal_idxss ~~ exclude_thmss);
-        fun mk_excludesss excludes n =
-          fold (fn ((c, c', _), thm) => nth_map c (nth_map c' (K [thm])))
-            excludes (map (fn k => replicate k [asm_rl] @ replicate (n - k) []) (0 upto n - 1));
-        val excludessss =
-          map2 (fn excludes => mk_excludesss excludes o length o #ctr_specs)
-            (map2 append excludess' taut_thmss) corec_specs;
-
-        fun prove_disc ({ctr_specs, ...} : corec_spec) excludesss
-            ({fun_name, fun_T, fun_args, ctr_no, prems, eqn_pos, ...} : coeqn_data_disc) =
-          if Term.aconv_untyped (#disc (nth ctr_specs ctr_no), @{term "\<lambda>x. x = x"}) then
-            []
-          else
-            let
-              val {disc, disc_corec, ...} = nth ctr_specs ctr_no;
-              val k = 1 + ctr_no;
-              val m = length prems;
-              val goal =
-                applied_fun_of fun_name fun_T fun_args
-                |> curry betapply disc
-                |> HOLogic.mk_Trueprop
-                |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
-                |> curry Logic.list_all (map dest_Free fun_args);
-            in
-              if prems = [@{term False}] then
-                []
-              else
-                mk_primcorec_disc_tac lthy def_thms disc_corec k m excludesss
-                |> K |> Goal.prove_sorry lthy [] [] goal
-                |> Thm.close_derivation
-                |> pair (#disc (nth ctr_specs ctr_no))
-                |> pair eqn_pos
-                |> single
-            end;
-
-        fun prove_sel ({nested_maps, nested_map_idents, nested_map_comps, ctr_specs, ...}
-              : corec_spec) (disc_eqns : coeqn_data_disc list) excludesss
-            ({fun_name, fun_T, fun_args, ctr, sel, rhs_term, eqn_pos, ...} : coeqn_data_sel) =
-          let
-            val SOME ctr_spec = find_first (curry (op =) ctr o #ctr) ctr_specs;
-            val ctr_no = find_index (curry (op =) ctr o #ctr) ctr_specs;
-            val prems = the_default (maps (s_not_conj o #prems) disc_eqns)
-              (find_first (curry (op =) ctr_no o #ctr_no) disc_eqns |> Option.map #prems);
-            val sel_corec = find_index (curry (op =) sel) (#sels ctr_spec)
-              |> nth (#sel_corecs ctr_spec);
-            val k = 1 + ctr_no;
-            val m = length prems;
-            val goal =
-              applied_fun_of fun_name fun_T fun_args
-              |> curry betapply sel
-              |> rpair (abstract (List.rev fun_args) rhs_term)
-              |> HOLogic.mk_Trueprop o HOLogic.mk_eq
-              |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
-              |> curry Logic.list_all (map dest_Free fun_args);
-            val (distincts, _, _, sel_splits, sel_split_asms) = case_thms_of_term lthy [] rhs_term;
-          in
-            mk_primcorec_sel_tac lthy def_thms distincts sel_splits sel_split_asms nested_maps
-              nested_map_idents nested_map_comps sel_corec k m excludesss
-            |> K |> Goal.prove_sorry lthy [] [] goal
-            |> Thm.close_derivation
-            |> pair sel
-            |> pair eqn_pos
-          end;
-
-        fun prove_ctr disc_alist sel_alist (disc_eqns : coeqn_data_disc list)
-            (sel_eqns : coeqn_data_sel list) ({ctr, disc, sels, collapse, ...} : corec_ctr_spec) =
-          (* don't try to prove theorems when some sel_eqns are missing *)
-          if not (exists (curry (op =) ctr o #ctr) disc_eqns)
-              andalso not (exists (curry (op =) ctr o #ctr) sel_eqns)
-            orelse
-              filter (curry (op =) ctr o #ctr) sel_eqns
-              |> fst o finds (op = o apsnd #sel) sels
-              |> exists (null o snd) then
-            []
-          else
-            let
-              val (fun_name, fun_T, fun_args, prems, rhs_opt, eqn_pos) =
-                (find_first (curry (op =) ctr o #ctr) disc_eqns,
-                 find_first (curry (op =) ctr o #ctr) sel_eqns)
-                |>> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #prems x,
-                  #ctr_rhs_opt x, #eqn_pos x))
-                ||> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, [], #ctr_rhs_opt x,
-                  #eqn_pos x))
-                |> the o merge_options;
-              val m = length prems;
-              val goal =
-                (case rhs_opt of
-                  SOME rhs => rhs
-                | NONE =>
-                  filter (curry (op =) ctr o #ctr) sel_eqns
-                  |> fst o finds (op = o apsnd #sel) sels
-                  |> map (snd #> (fn [x] => (List.rev (#fun_args x), #rhs_term x)) #-> abstract)
-                  |> curry list_comb ctr)
-                |> curry mk_Trueprop_eq (applied_fun_of fun_name fun_T fun_args)
-                |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
-                |> curry Logic.list_all (map dest_Free fun_args);
-              val disc_thm_opt = AList.lookup (op =) disc_alist disc;
-              val sel_thms = map snd (filter (member (op =) sels o fst) sel_alist);
-            in
-              if prems = [@{term False}] then [] else
-                mk_primcorec_ctr_tac lthy m collapse disc_thm_opt sel_thms
-                |> K |> Goal.prove_sorry lthy [] [] goal
-                |> Thm.close_derivation
-                |> pair ctr
-                |> pair eqn_pos
-                |> single
-            end;
-
-        fun prove_code exhaustive disc_eqns sel_eqns nchotomys ctr_alist ctr_specs =
-          let
-            val fun_data_opt =
-              (find_first (member (op =) (map #ctr ctr_specs) o #ctr) disc_eqns,
-               find_first (member (op =) (map #ctr ctr_specs) o #ctr) sel_eqns)
-              |>> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #code_rhs_opt x))
-              ||> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #code_rhs_opt x))
-              |> merge_options;
-          in
-            (case fun_data_opt of
-              NONE => []
-            | SOME (fun_name, fun_T, fun_args, rhs_opt) =>
-              let
-                val bound_Ts = List.rev (map fastype_of fun_args);
-
-                val lhs = applied_fun_of fun_name fun_T fun_args;
-                val rhs_info_opt =
-                  (case rhs_opt of
-                    SOME rhs =>
-                    let
-                      val raw_rhs = expand_corec_code_rhs lthy has_call bound_Ts rhs;
-                      val cond_ctrs =
-                        fold_rev_corec_code_rhs lthy (K oo (cons oo pair)) bound_Ts raw_rhs [];
-                      val ctr_thms =
-                        map (the_default FalseE o AList.lookup (op =) ctr_alist o snd) cond_ctrs;
-                    in SOME (false, rhs, raw_rhs, ctr_thms) end
-                  | NONE =>
-                    let
-                      fun prove_code_ctr {ctr, sels, ...} =
-                        if not (exists (curry (op =) ctr o fst) ctr_alist) then NONE else
-                          let
-                            val prems = find_first (curry (op =) ctr o #ctr) disc_eqns
-                              |> Option.map #prems |> the_default [];
-                            val t =
-                              filter (curry (op =) ctr o #ctr) sel_eqns
-                              |> fst o finds (op = o apsnd #sel) sels
-                              |> map (snd #> (fn [x] => (List.rev (#fun_args x), #rhs_term x))
-                                #-> abstract)
-                              |> curry list_comb ctr;
-                          in
-                            SOME (prems, t)
-                          end;
-                      val ctr_conds_argss_opt = map prove_code_ctr ctr_specs;
-                      val exhaustive_code =
-                        exhaustive
-                        orelse exists (is_some andf (null o fst o the)) ctr_conds_argss_opt
-                        orelse forall is_some ctr_conds_argss_opt
-                          andalso exists #auto_gen disc_eqns;
-                      val rhs =
-                        (if exhaustive_code then
-                           split_last (map_filter I ctr_conds_argss_opt) ||> snd
-                         else
-                           Const (@{const_name Code.abort}, @{typ String.literal} -->
-                               (@{typ unit} --> body_type fun_T) --> body_type fun_T) $
-                             HOLogic.mk_literal fun_name $
-                             absdummy @{typ unit} (incr_boundvars 1 lhs)
-                           |> pair (map_filter I ctr_conds_argss_opt))
-                         |-> fold_rev (fn (prems, u) => mk_If (s_conjs prems) u)
-                    in
-                      SOME (exhaustive_code, rhs, rhs, map snd ctr_alist)
-                    end);
-              in
-                (case rhs_info_opt of
-                  NONE => []
-                | SOME (exhaustive_code, rhs, raw_rhs, ctr_thms) =>
-                  let
-                    val ms = map (Logic.count_prems o prop_of) ctr_thms;
-                    val (raw_goal, goal) = (raw_rhs, rhs)
-                      |> pairself (curry mk_Trueprop_eq (applied_fun_of fun_name fun_T fun_args)
-                        #> curry Logic.list_all (map dest_Free fun_args));
-                    val (distincts, discIs, _, sel_splits, sel_split_asms) =
-                      case_thms_of_term lthy bound_Ts raw_rhs;
-
-                    val raw_code_thm = mk_primcorec_raw_code_tac lthy distincts discIs sel_splits
-                        sel_split_asms ms ctr_thms
-                        (if exhaustive_code then try the_single nchotomys else NONE)
-                      |> K |> Goal.prove_sorry lthy [] [] raw_goal
-                      |> Thm.close_derivation;
-                  in
-                    mk_primcorec_code_tac lthy distincts sel_splits raw_code_thm
-                    |> K |> Goal.prove_sorry lthy [] [] goal
-                    |> Thm.close_derivation
-                    |> single
-                  end)
-              end)
-          end;
-
-        val disc_alistss = map3 (map oo prove_disc) corec_specs excludessss disc_eqnss;
-        val disc_alists = map (map snd o flat) disc_alistss;
-        val sel_alists = map4 (map ooo prove_sel) corec_specs disc_eqnss excludessss sel_eqnss;
-        val disc_thmss = map (map snd o order_list_duplicates o flat) disc_alistss;
-        val disc_thmsss' = map (map (map (snd o snd))) disc_alistss;
-        val sel_thmss = map (map snd o order_list_duplicates) sel_alists;
-
-        fun prove_disc_iff ({ctr_specs, ...} : corec_spec) exhaust_thms disc_thmss'
-            (({fun_args = exhaust_fun_args, ...} : coeqn_data_disc) :: _) disc_thms
-            ({fun_name, fun_T, fun_args, ctr_no, prems, eqn_pos, ...} : coeqn_data_disc) =
-          if null exhaust_thms orelse null (tl ctr_specs) then
-            []
-          else
-            let
-              val {disc, disc_excludess, ...} = nth ctr_specs ctr_no;
-              val goal =
-                mk_Trueprop_eq (applied_fun_of fun_name fun_T fun_args |> curry betapply disc,
-                  mk_conjs prems)
-                |> curry Logic.list_all (map dest_Free fun_args);
-            in
-              mk_primcorec_disc_iff_tac lthy (map (fst o dest_Free) exhaust_fun_args)
-                (the_single exhaust_thms) disc_thms disc_thmss' (flat disc_excludess)
-              |> K |> Goal.prove_sorry lthy [] [] goal
-              |> Thm.close_derivation
-              |> fold (fn rule => perhaps (try (fn thm => thm RS rule)))
-                @{thms eqTrueE eq_False[THEN iffD1] notnotD}
-              |> pair eqn_pos
-              |> single
-            end;
-
-        val disc_iff_thmss = map6 (flat ooo map2 oooo prove_disc_iff) corec_specs exhaust_thmss
-          disc_thmsss' disc_eqnss disc_thmsss' disc_eqnss
-          |> map order_list_duplicates;
-
-        val ctr_alists = map5 (maps oooo prove_ctr) disc_alists (map (map snd) sel_alists) disc_eqnss
-          sel_eqnss ctr_specss;
-        val ctr_thmss' = map (map snd) ctr_alists;
-        val ctr_thmss = map (map snd o order_list) ctr_alists;
-
-        val code_thmss = map6 prove_code exhaustives disc_eqnss sel_eqnss nchotomy_thmss ctr_thmss'
-          ctr_specss;
-
-        val disc_iff_or_disc_thmss =
-          map2 (fn [] => I | disc_iffs => K disc_iffs) disc_iff_thmss disc_thmss;
-        val simp_thmss = map2 append disc_iff_or_disc_thmss sel_thmss;
-
-        val common_name = mk_common_name fun_names;
-
-        val notes =
-          [(coinductN, map (if n2m then single else K []) coinduct_thms, []),
-           (codeN, code_thmss, code_nitpicksimp_attrs),
-           (ctrN, ctr_thmss, []),
-           (discN, disc_thmss, simp_attrs),
-           (disc_iffN, disc_iff_thmss, []),
-           (excludeN, exclude_thmss, []),
-           (exhaustN, nontriv_exhaust_thmss, []),
-           (selN, sel_thmss, simp_attrs),
-           (simpsN, simp_thmss, []),
-           (strong_coinductN, map (if n2m then single else K []) strong_coinduct_thms, [])]
-          |> maps (fn (thmN, thmss, attrs) =>
-            map2 (fn fun_name => fn thms =>
-                ((Binding.qualify true fun_name (Binding.name thmN), attrs), [(thms, [])]))
-              fun_names (take actual_nn thmss))
-          |> filter_out (null o fst o hd o snd);
-
-        val common_notes =
-          [(coinductN, if n2m then [coinduct_thm] else [], []),
-           (strong_coinductN, if n2m then [strong_coinduct_thm] else [], [])]
-          |> filter_out (null o #2)
-          |> map (fn (thmN, thms, attrs) =>
-            ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
-      in
-        lthy |> Local_Theory.notes (notes @ common_notes) |> snd
-      end;
-
-    fun after_qed thmss' = fold_map Local_Theory.define defs #-> prove thmss';
-  in
-    (goalss, after_qed, lthy')
-  end;
-
-fun add_primcorec_ursive_cmd auto opts (raw_fixes, raw_specs') lthy =
-  let
-    val (raw_specs, of_specs_opt) =
-      split_list raw_specs' ||> map (Option.map (Syntax.read_term lthy));
-    val ((fixes, specs), _) = Specification.read_spec raw_fixes raw_specs lthy;
-  in
-    add_primcorec_ursive auto opts fixes specs of_specs_opt lthy
-    handle ERROR str => primcorec_error str
-  end
-  handle Primcorec_Error (str, eqns) =>
-    if null eqns
-    then error ("primcorec error:\n  " ^ str)
-    else error ("primcorec error:\n  " ^ str ^ "\nin\n  " ^
-      space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns));
-
-val add_primcorecursive_cmd = (fn (goalss, after_qed, lthy) =>
-  lthy
-  |> Proof.theorem NONE after_qed goalss
-  |> Proof.refine (Method.primitive_text (K I))
-  |> Seq.hd) ooo add_primcorec_ursive_cmd false;
-
-val add_primcorec_cmd = (fn (goalss, after_qed, lthy) =>
-  lthy
-  |> after_qed (map (fn [] => []
-      | _ => error "\"auto\" failed -- use \"primcorecursive\" instead of \"primcorec\"")
-    goalss)) ooo add_primcorec_ursive_cmd true;
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_gfp_rec_sugar_tactics.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,209 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_gfp_rec_sugar_tactics.ML
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2013
-
-Tactics for corecursor sugar.
-*)
-
-signature BNF_GFP_REC_SUGAR_TACTICS =
-sig
-  val mk_primcorec_assumption_tac: Proof.context -> thm list -> int -> tactic
-  val mk_primcorec_code_tac: Proof.context -> thm list -> thm list -> thm -> tactic
-  val mk_primcorec_ctr_tac: Proof.context -> int -> thm -> thm option -> thm list -> tactic
-  val mk_primcorec_disc_tac: Proof.context -> thm list -> thm -> int -> int -> thm list list list ->
-    tactic
-  val mk_primcorec_disc_iff_tac: Proof.context -> string list -> thm -> thm list -> thm list list ->
-    thm list -> tactic
-  val mk_primcorec_exhaust_tac: Proof.context -> string list -> int -> thm -> tactic
-  val mk_primcorec_nchotomy_tac: Proof.context -> thm list -> tactic
-  val mk_primcorec_raw_code_tac: Proof.context -> thm list -> thm list -> thm list -> thm list ->
-    int list -> thm list -> thm option -> tactic
-  val mk_primcorec_sel_tac: Proof.context -> thm list -> thm list -> thm list -> thm list ->
-    thm list -> thm list -> thm list -> thm -> int -> int -> thm list list list -> tactic
-end;
-
-structure BNF_GFP_Rec_Sugar_Tactics : BNF_GFP_REC_SUGAR_TACTICS =
-struct
-
-open BNF_Util
-open BNF_Tactics
-open BNF_FP_Util
-
-val atomize_conjL = @{thm atomize_conjL};
-val falseEs = @{thms not_TrueE FalseE};
-val neq_eq_eq_contradict = @{thm neq_eq_eq_contradict};
-val split_if = @{thm split_if};
-val split_if_asm = @{thm split_if_asm};
-val split_connectI = @{thms allI impI conjI};
-val unfold_lets = @{thms Let_def[abs_def] split_beta}
-
-fun exhaust_inst_as_projs ctxt frees thm =
-  let
-    val num_frees = length frees;
-    val fs = Term.add_vars (prop_of thm) [] |> filter (can dest_funT o snd);
-    fun find s = find_index (curry (op =) s) frees;
-    fun mk_cfp (f as ((s, _), T)) =
-      (certify ctxt (Var f), certify ctxt (mk_proj T num_frees (find s)));
-    val cfps = map mk_cfp fs;
-  in
-    Drule.cterm_instantiate cfps thm
-  end;
-
-val exhaust_inst_as_projs_tac = PRIMITIVE oo exhaust_inst_as_projs;
-
-fun distinct_in_prems_tac distincts =
-  eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac;
-
-fun mk_primcorec_nchotomy_tac ctxt disc_exhausts =
-  HEADGOAL (Method.insert_tac disc_exhausts THEN' clean_blast_tac ctxt);
-
-fun mk_primcorec_exhaust_tac ctxt frees n nchotomy =
-  let val ks = 1 upto n in
-    HEADGOAL (atac ORELSE'
-      cut_tac nchotomy THEN'
-      K (exhaust_inst_as_projs_tac ctxt frees) THEN'
-      EVERY' (map (fn k =>
-          (if k < n then etac disjE else K all_tac) THEN'
-          REPEAT o (dtac meta_mp THEN' atac ORELSE'
-            etac conjE THEN' dtac meta_mp THEN' atac ORELSE'
-            atac))
-        ks))
-  end;
-
-fun mk_primcorec_assumption_tac ctxt discIs =
-  SELECT_GOAL (unfold_thms_tac ctxt @{thms fst_conv snd_conv not_not not_False_eq_True
-      not_True_eq_False de_Morgan_conj de_Morgan_disj} THEN
-    SOLVE (HEADGOAL (REPEAT o (rtac refl ORELSE' atac ORELSE' etac conjE ORELSE'
-    eresolve_tac falseEs ORELSE'
-    resolve_tac @{thms TrueI conjI disjI1 disjI2} ORELSE'
-    dresolve_tac discIs THEN' atac ORELSE'
-    etac notE THEN' atac ORELSE'
-    etac disjE))));
-
-val ss_fst_snd_conv = simpset_of (ss_only @{thms fst_conv snd_conv} @{context});
-
-fun case_atac ctxt = simp_tac (put_simpset ss_fst_snd_conv ctxt);
-
-fun same_case_tac ctxt m =
-  HEADGOAL (if m = 0 then rtac TrueI
-    else REPEAT_DETERM_N (m - 1) o (rtac conjI THEN' case_atac ctxt) THEN' case_atac ctxt);
-
-fun different_case_tac ctxt m exclude =
-  HEADGOAL (if m = 0 then
-      mk_primcorec_assumption_tac ctxt []
-    else
-      dtac exclude THEN' (REPEAT_DETERM_N (m - 1) o case_atac ctxt) THEN'
-      mk_primcorec_assumption_tac ctxt []);
-
-fun cases_tac ctxt k m excludesss =
-  let val n = length excludesss in
-    EVERY (map (fn [] => if k = n then all_tac else same_case_tac ctxt m
-        | [exclude] => different_case_tac ctxt m exclude)
-      (take k (nth excludesss (k - 1))))
-  end;
-
-fun prelude_tac ctxt defs thm =
-  unfold_thms_tac ctxt defs THEN HEADGOAL (rtac thm) THEN unfold_thms_tac ctxt unfold_lets;
-
-fun mk_primcorec_disc_tac ctxt defs disc_corec k m excludesss =
-  prelude_tac ctxt defs disc_corec THEN cases_tac ctxt k m excludesss;
-
-fun mk_primcorec_disc_iff_tac ctxt fun_exhaust_frees fun_exhaust fun_discs fun_discss
-    disc_excludes =
-  HEADGOAL (rtac iffI THEN'
-    rtac fun_exhaust THEN'
-    K (exhaust_inst_as_projs_tac ctxt fun_exhaust_frees) THEN'
-    EVERY' (map (fn [] => etac FalseE
-        | fun_discs' as [fun_disc'] =>
-          if eq_list Thm.eq_thm (fun_discs', fun_discs) then
-            REPEAT_DETERM o etac conjI THEN' (atac ORELSE' rtac TrueI)
-          else
-            rtac FalseE THEN'
-            (rotate_tac 1 THEN' dtac fun_disc' THEN' REPEAT_DETERM o atac ORELSE'
-             cut_tac fun_disc') THEN'
-            dresolve_tac disc_excludes THEN' etac notE THEN' atac)
-      fun_discss) THEN'
-    (etac FalseE ORELSE'
-     resolve_tac (map (unfold_thms ctxt [atomize_conjL]) fun_discs) THEN_MAYBE' atac));
-
-fun mk_primcorec_sel_tac ctxt defs distincts splits split_asms mapsx map_idents map_comps fun_sel k
-    m excludesss =
-  prelude_tac ctxt defs (fun_sel RS trans) THEN
-  cases_tac ctxt k m excludesss THEN
-  HEADGOAL (REPEAT_DETERM o (rtac refl ORELSE' rtac ext ORELSE'
-    eresolve_tac falseEs ORELSE'
-    resolve_tac split_connectI ORELSE'
-    Splitter.split_asm_tac (split_if_asm :: split_asms) ORELSE'
-    Splitter.split_tac (split_if :: splits) ORELSE'
-    eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac ORELSE'
-    etac notE THEN' atac ORELSE'
-    (CHANGED o SELECT_GOAL (unfold_thms_tac ctxt (@{thms fst_conv snd_conv id_def o_def split_def
-       sum.cases} @ mapsx @ map_comps @ map_idents)))));
-
-fun mk_primcorec_ctr_tac ctxt m collapse disc_fun_opt sel_funs =
-  HEADGOAL (rtac ((if null sel_funs then collapse else collapse RS sym) RS trans) THEN'
-    (the_default (K all_tac) (Option.map rtac disc_fun_opt)) THEN' REPEAT_DETERM_N m o atac) THEN
-  unfold_thms_tac ctxt (@{thm split_def} :: unfold_lets @ sel_funs) THEN HEADGOAL (rtac refl);
-
-fun inst_split_eq ctxt split =
-  (case prop_of split of
-    @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ (Var (_, Type (_, [T, _])) $ _) $ _) =>
-    let
-      val s = Name.uu;
-      val eq = Abs (Name.uu, T, HOLogic.mk_eq (Free (s, T), Bound 0));
-      val split' = Drule.instantiate' [] [SOME (certify ctxt eq)] split;
-    in
-      Thm.generalize ([], [s]) (Thm.maxidx_of split' + 1) split'
-    end
-  | _ => split);
-
-fun raw_code_single_tac ctxt distincts discIs splits split_asms m fun_ctr =
-  let
-    val splits' =
-      map (fn th => th RS iffD2) (@{thm split_if_eq2} :: map (inst_split_eq ctxt) splits)
-  in
-    HEADGOAL (REPEAT o (resolve_tac (splits' @ split_connectI))) THEN
-    prelude_tac ctxt [] (fun_ctr RS trans) THEN
-    HEADGOAL ((REPEAT_DETERM_N m o mk_primcorec_assumption_tac ctxt discIs) THEN'
-      SELECT_GOAL (SOLVE (HEADGOAL (REPEAT_DETERM o
-      (rtac refl ORELSE' atac ORELSE'
-       resolve_tac (@{thm Code.abort_def} :: split_connectI) ORELSE'
-       Splitter.split_tac (split_if :: splits) ORELSE'
-       Splitter.split_asm_tac (split_if_asm :: split_asms) ORELSE'
-       mk_primcorec_assumption_tac ctxt discIs ORELSE'
-       distinct_in_prems_tac distincts ORELSE'
-       (TRY o dresolve_tac discIs) THEN' etac notE THEN' atac)))))
-  end;
-
-fun rulify_nchotomy n = funpow (n - 1) (fn thm => thm RS @{thm Meson.make_pos_rule'});
-
-fun mk_primcorec_raw_code_tac ctxt distincts discIs splits split_asms ms fun_ctrs nchotomy_opt =
-  let
-    val n = length ms;
-    val (ms', fun_ctrs') =
-      (case nchotomy_opt of
-        NONE => (ms, fun_ctrs)
-      | SOME nchotomy =>
-        (ms |> split_last ||> K [n - 1] |> op @,
-         fun_ctrs
-         |> split_last
-         ||> unfold_thms ctxt [atomize_conjL]
-         ||> (fn thm => [rulify_nchotomy n nchotomy RS thm] handle THM _ => [thm])
-         |> op @));
-  in
-    EVERY (map2 (raw_code_single_tac ctxt distincts discIs splits split_asms) ms' fun_ctrs') THEN
-    IF_UNSOLVED (unfold_thms_tac ctxt @{thms Code.abort_def} THEN
-      HEADGOAL (REPEAT_DETERM o resolve_tac (refl :: split_connectI)))
-  end;
-
-fun mk_primcorec_code_tac ctxt distincts splits raw =
-  HEADGOAL (rtac raw ORELSE' rtac (raw RS trans) THEN'
-    SELECT_GOAL (unfold_thms_tac ctxt unfold_lets) THEN' REPEAT_DETERM o
-    (rtac refl ORELSE' atac ORELSE'
-     resolve_tac split_connectI ORELSE'
-     Splitter.split_tac (split_if :: splits) ORELSE'
-     distinct_in_prems_tac distincts ORELSE'
-     rtac sym THEN' atac ORELSE'
-     etac notE THEN' atac));
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_gfp_tactics.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1304 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_gfp_tactics.ML
-    Author:     Dmitriy Traytel, TU Muenchen
-    Author:     Andrei Popescu, TU Muenchen
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2012
-
-Tactics for the codatatype construction.
-*)
-
-signature BNF_GFP_TACTICS =
-sig
-  val mk_Lev_sbd_tac: Proof.context -> cterm option list -> thm list -> thm list ->
-    thm list list -> tactic
-  val mk_bis_Gr_tac: thm -> thm list -> thm list -> thm list -> thm list ->
-    {prems: 'a, context: Proof.context} -> tactic
-  val mk_bis_O_tac: Proof.context -> int -> thm -> thm list -> thm list -> tactic
-  val mk_bis_Union_tac: thm -> thm list -> {prems: 'a, context: Proof.context} -> tactic
-  val mk_bis_converse_tac: int -> thm -> thm list -> thm list -> tactic
-  val mk_bis_rel_tac: Proof.context -> int -> thm -> thm list -> thm list -> thm list ->
-    thm list list -> tactic
-  val mk_coalgT_tac: int -> thm list -> thm list -> thm list list ->
-    {prems: 'a, context: Proof.context} -> tactic
-  val mk_coalg_final_tac: int -> thm -> thm list -> thm list -> thm list list -> thm list list ->
-    tactic
-  val mk_coalg_set_tac: thm -> tactic
-  val mk_coind_wit_tac: thm -> thm list -> thm list -> thm list ->
-    {prems: 'a, context: Proof.context} -> tactic
-  val mk_col_bd_tac: int -> int -> cterm option list -> thm list -> thm list -> thm -> thm ->
-    thm list list -> tactic
-  val mk_col_natural_tac: cterm option list -> thm list -> thm list -> thm list -> thm list list ->
-    {prems: 'a, context: Proof.context} -> tactic
-  val mk_congruent_str_final_tac: int -> thm -> thm -> thm -> thm list -> tactic
-  val mk_corec_tac: int -> thm list -> thm -> thm -> thm list ->
-    {prems: 'a, context: Proof.context} -> tactic
-  val mk_corec_unique_mor_tac: thm list -> thm list -> thm -> {prems: 'a, context: Proof.context} ->
-    tactic
-  val mk_dtor_map_coinduct_tac: int -> int list -> thm -> thm -> tactic
-  val mk_dtor_coinduct_tac: int -> thm -> thm -> thm list -> tactic
-  val mk_dtor_rel_tac: Proof.context -> thm list -> int -> thm -> thm -> thm -> thm -> thm list ->
-    thm -> thm -> thm list -> thm list -> thm list list -> tactic
-  val mk_dtor_o_ctor_tac: thm -> thm -> thm -> thm -> thm list ->
-    {prems: 'a, context: Proof.context} -> tactic
-  val mk_equiv_lsbis_tac: thm -> thm -> thm -> thm -> thm -> thm -> tactic
-  val mk_hset_minimal_tac: int -> thm list -> thm -> {prems: 'a, context: Proof.context} -> tactic
-  val mk_hset_rec_minimal_tac: int -> cterm option list -> thm list -> thm list ->
-    {prems: 'a, context: Proof.context} -> tactic
-  val mk_incl_lsbis_tac: int -> int -> thm -> tactic
-  val mk_length_Lev'_tac: thm -> tactic
-  val mk_length_Lev_tac: Proof.context -> cterm option list -> thm list -> thm list -> tactic
-  val mk_map_comp0_tac: thm list -> thm list -> thm -> tactic
-  val mk_mcong_tac: Proof.context -> int -> (int -> tactic) -> thm list -> thm list -> thm list ->
-    thm list list -> thm list list -> thm list list list -> tactic
-  val mk_map_id0_tac: thm list -> thm -> thm -> tactic
-  val mk_map_tac: int -> int -> ctyp option -> thm -> thm -> thm -> tactic
-  val mk_dtor_map_unique_tac: thm -> thm list -> Proof.context -> tactic
-  val mk_mor_Abs_tac: thm list -> thm list -> {prems: 'a, context: Proof.context} -> tactic
-  val mk_mor_Rep_tac: int -> thm list -> thm list -> thm list -> thm list list -> thm list ->
-    thm list -> {prems: 'a, context: Proof.context} -> tactic
-  val mk_mor_T_final_tac: thm -> thm list -> thm list -> tactic
-  val mk_mor_UNIV_tac: thm list -> thm -> tactic
-  val mk_mor_beh_tac: int -> thm -> thm -> thm list -> thm list -> thm list -> thm list ->
-    thm list list -> thm list list -> thm list -> thm list -> thm list -> thm list -> thm list ->
-    thm list -> thm list -> thm list -> thm list list -> thm list list list -> thm list list list ->
-    thm list list list -> thm list list -> thm list list -> thm list -> thm list -> thm list ->
-    {prems: 'a, context: Proof.context} -> tactic
-  val mk_mor_comp_tac: thm -> thm list -> thm list -> thm list -> tactic
-  val mk_mor_elim_tac: thm -> tactic
-  val mk_mor_hset_rec_tac: int -> int -> cterm option list -> int -> thm list -> thm list ->
-    thm list -> thm list list -> thm list list -> tactic
-  val mk_mor_incl_tac: thm -> thm list -> tactic
-  val mk_mor_str_tac: 'a list -> thm -> tactic
-  val mk_mor_sum_case_tac: 'a list -> thm -> tactic
-  val mk_mor_unfold_tac: int -> thm -> thm list -> thm list -> thm list -> thm list -> thm list ->
-    thm list -> tactic
-  val mk_prefCl_Lev_tac: Proof.context -> cterm option list -> thm list -> thm list -> tactic
-  val mk_raw_coind_tac: thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm list ->
-    thm list -> thm list -> thm -> thm list -> tactic
-  val mk_rel_coinduct_tac: thm list -> thm list -> thm list list -> thm list -> thm list ->
-    {prems: thm list, context: Proof.context} -> tactic
-  val mk_rel_coinduct_coind_tac: int -> thm -> int list -> thm list -> thm list -> thm list ->
-    thm list list -> thm list -> thm list -> {prems: 'a, context: Proof.context} -> tactic
-  val mk_rel_coinduct_ind_tac: int -> int list -> thm list -> thm list list -> int -> thm ->
-    {prems: 'a, context: Proof.context} -> tactic
-  val mk_rv_last_tac: ctyp option list -> cterm option list -> thm list -> thm list -> tactic
-  val mk_sbis_lsbis_tac: Proof.context -> thm list -> thm -> thm -> tactic
-  val mk_set_Lev_tac: Proof.context -> cterm option list -> thm list -> thm list -> thm list ->
-    thm list -> thm list list -> tactic
-  val mk_set_bd_tac: thm -> thm -> thm -> tactic
-  val mk_set_hset_incl_hset_tac: int -> thm list -> thm -> int -> tactic
-  val mk_set_image_Lev_tac: Proof.context -> cterm option list -> thm list -> thm list ->
-    thm list -> thm list -> thm list list -> thm list list -> tactic
-  val mk_set_incl_hset_tac: thm -> thm -> tactic
-  val mk_set_ge_tac: int  -> thm -> thm list -> tactic
-  val mk_set_le_tac: int -> thm -> thm list -> thm list list -> tactic
-  val mk_set_map0_tac: thm -> thm -> tactic
-  val mk_set_rv_Lev_tac: Proof.context -> int -> cterm option list -> thm list -> thm list ->
-    thm list -> thm list -> thm list list -> thm list list -> tactic
-  val mk_unfold_unique_mor_tac: thm list -> thm -> thm -> thm list -> tactic
-  val mk_unfold_transfer_tac: int -> thm -> thm list -> thm list ->
-    {prems: thm list, context: Proof.context} -> tactic
-  val mk_wit_tac: int -> thm list -> thm list -> thm list -> thm list -> Proof.context -> tactic
-  val mk_le_rel_OO_tac: thm -> thm list -> thm list -> tactic
-end;
-
-structure BNF_GFP_Tactics : BNF_GFP_TACTICS =
-struct
-
-open BNF_Tactics
-open BNF_Util
-open BNF_FP_Util
-open BNF_GFP_Util
-
-val fst_convol_fun_cong_sym = @{thm fst_convol} RS fun_cong RS sym;
-val list_inject_iffD1 = @{thm list.inject[THEN iffD1]};
-val nat_induct = @{thm nat_induct};
-val o_apply_trans_sym = o_apply RS trans RS sym;
-val ord_eq_le_trans = @{thm ord_eq_le_trans};
-val ord_eq_le_trans_trans_fun_cong_image_id_id_apply =
-  @{thm ord_eq_le_trans[OF trans[OF fun_cong[OF image_id] id_apply]]};
-val ordIso_ordLeq_trans = @{thm ordIso_ordLeq_trans};
-val snd_convol_fun_cong_sym = @{thm snd_convol} RS fun_cong RS sym;
-val sum_case_weak_cong = @{thm sum_case_weak_cong};
-val trans_fun_cong_image_id_id_apply = @{thm trans[OF fun_cong[OF image_id] id_apply]};
-val Collect_splitD_set_mp = @{thm Collect_splitD[OF set_mp]};
-val rev_bspec = Drule.rotate_prems 1 bspec;
-val Un_cong = @{thm arg_cong2[of _ _ _ _ "op \<union>"]};
-val conversep_in_rel_Id_on =
-   @{thm trans[OF conversep_in_rel arg_cong[of _ _ in_rel, OF converse_Id_on]]};
-val relcompp_in_rel_Id_on =
-   @{thm   trans[OF relcompp_in_rel arg_cong[of _ _ in_rel, OF Id_on_Comp[symmetric]]]};
-val converse_shift = @{thm converse_subset_swap} RS iffD1;
-
-fun mk_coalg_set_tac coalg_def =
-  dtac (coalg_def RS iffD1) 1 THEN
-  REPEAT_DETERM (etac conjE 1) THEN
-  EVERY' [dtac rev_bspec, atac] 1 THEN
-  REPEAT_DETERM (eresolve_tac [CollectE, conjE] 1) THEN atac 1;
-
-fun mk_mor_elim_tac mor_def =
-  (dtac (subst OF [mor_def]) THEN'
-  REPEAT o etac conjE THEN'
-  TRY o rtac @{thm image_subsetI} THEN'
-  etac bspec THEN'
-  atac) 1;
-
-fun mk_mor_incl_tac mor_def map_ids =
-  (stac mor_def THEN'
-  rtac conjI THEN'
-  CONJ_WRAP' (K (EVERY' [rtac ballI, etac set_mp, stac id_apply, atac])) map_ids THEN'
-  CONJ_WRAP' (fn thm =>
-   (EVERY' [rtac ballI, rtac (thm RS trans), rtac sym, rtac (id_apply RS arg_cong)])) map_ids) 1;
-
-fun mk_mor_comp_tac mor_def mor_images morEs map_comp_ids =
-  let
-    fun fbetw_tac image = EVERY' [rtac ballI, stac o_apply, etac image, etac image, atac];
-    fun mor_tac ((mor_image, morE), map_comp_id) =
-      EVERY' [rtac ballI, stac o_apply, rtac trans, rtac (map_comp_id RS sym), rtac trans,
-        etac (morE RS arg_cong), atac, etac morE, etac mor_image, atac];
-  in
-    (stac mor_def THEN' rtac conjI THEN'
-    CONJ_WRAP' fbetw_tac mor_images THEN'
-    CONJ_WRAP' mor_tac ((mor_images ~~ morEs) ~~ map_comp_ids)) 1
-  end;
-
-fun mk_mor_UNIV_tac morEs mor_def =
-  let
-    val n = length morEs;
-    fun mor_tac morE = EVERY' [rtac ext, rtac trans, rtac o_apply, rtac trans, etac morE,
-      rtac UNIV_I, rtac sym, rtac o_apply];
-  in
-    EVERY' [rtac iffI, CONJ_WRAP' mor_tac morEs,
-    stac mor_def, rtac conjI, CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) morEs,
-    CONJ_WRAP' (fn i =>
-      EVERY' [dtac (mk_conjunctN n i), rtac ballI, etac @{thm comp_eq_dest}]) (1 upto n)] 1
-  end;
-
-fun mk_mor_str_tac ks mor_UNIV =
-  (stac mor_UNIV THEN' CONJ_WRAP' (K (rtac refl)) ks) 1;
-
-fun mk_mor_sum_case_tac ks mor_UNIV =
-  (stac mor_UNIV THEN' CONJ_WRAP' (K (rtac @{thm sum_case_o_inj(1)[symmetric]})) ks) 1;
-
-fun mk_set_incl_hset_tac def rec_Suc =
-  EVERY' (stac def ::
-    map rtac [@{thm SUP_upper2}, UNIV_I, @{thm ord_le_eq_trans}, @{thm Un_upper1},
-      sym, rec_Suc]) 1;
-
-fun mk_set_hset_incl_hset_tac n defs rec_Suc i =
-  EVERY' (map (TRY oo stac) defs @
-    map rtac [@{thm UN_least}, subsetI, @{thm UN_I}, UNIV_I, set_mp, equalityD2, rec_Suc, UnI2,
-      mk_UnIN n i] @
-    [etac @{thm UN_I}, atac]) 1;
-
-fun mk_hset_rec_minimal_tac m cts rec_0s rec_Sucs {context = ctxt, prems = _} =
-  EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
-    REPEAT_DETERM o rtac allI,
-    CONJ_WRAP' (fn thm => EVERY'
-      [rtac ord_eq_le_trans, rtac thm, rtac @{thm empty_subsetI}]) rec_0s,
-    REPEAT_DETERM o rtac allI,
-    CONJ_WRAP' (fn rec_Suc => EVERY'
-      [rtac ord_eq_le_trans, rtac rec_Suc,
-        if m = 0 then K all_tac
-        else (rtac @{thm Un_least} THEN' Goal.assume_rule_tac ctxt),
-        CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least}))
-          (K (EVERY' [rtac @{thm UN_least}, REPEAT_DETERM o eresolve_tac [allE, conjE],
-            rtac subset_trans, atac, Goal.assume_rule_tac ctxt])) rec_0s])
-      rec_Sucs] 1;
-
-fun mk_hset_minimal_tac n hset_defs hset_rec_minimal {context = ctxt, prems = _} =
-  (CONJ_WRAP' (fn def => (EVERY' [rtac ord_eq_le_trans, rtac def,
-    rtac @{thm UN_least}, rtac rev_mp, rtac hset_rec_minimal,
-    EVERY' (replicate ((n + 1) * n) (Goal.assume_rule_tac ctxt)), rtac impI,
-    REPEAT_DETERM o eresolve_tac [allE, conjE], atac])) hset_defs) 1
-
-fun mk_mor_hset_rec_tac m n cts j rec_0s rec_Sucs morEs set_map0ss coalg_setss =
-  EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
-    REPEAT_DETERM o rtac allI,
-    CONJ_WRAP' (fn thm => EVERY' (map rtac [impI, thm RS trans, thm RS sym])) rec_0s,
-    REPEAT_DETERM o rtac allI,
-    CONJ_WRAP'
-      (fn (rec_Suc, (morE, ((passive_set_map0s, active_set_map0s), coalg_sets))) =>
-        EVERY' [rtac impI, rtac (rec_Suc RS trans), rtac (rec_Suc RS trans RS sym),
-          if m = 0 then K all_tac
-          else EVERY' [rtac Un_cong, rtac box_equals,
-            rtac (nth passive_set_map0s (j - 1) RS sym),
-            rtac trans_fun_cong_image_id_id_apply, etac (morE RS arg_cong), atac],
-          CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 Un_cong))
-            (fn (i, (set_map0, coalg_set)) =>
-              EVERY' [rtac sym, rtac trans, rtac (refl RSN (2, @{thm UN_cong})),
-                etac (morE RS sym RS arg_cong RS trans), atac, rtac set_map0,
-                rtac (@{thm UN_simps(10)} RS trans), rtac (refl RS @{thm UN_cong}),
-                ftac coalg_set, atac, dtac set_mp, atac, rtac mp, rtac (mk_conjunctN n i),
-                REPEAT_DETERM o etac allE, atac, atac])
-            (rev ((1 upto n) ~~ (active_set_map0s ~~ coalg_sets)))])
-      (rec_Sucs ~~ (morEs ~~ (map (chop m) set_map0ss ~~ map (drop m) coalg_setss)))] 1;
-
-fun mk_bis_rel_tac ctxt m bis_def rel_OO_Grps map_comp0s map_cong0s set_map0ss =
-  let
-    val n = length rel_OO_Grps;
-    val thms = ((1 upto n) ~~ map_comp0s ~~ map_cong0s ~~ set_map0ss ~~ rel_OO_Grps);
-
-    fun mk_if_tac ((((i, map_comp0), map_cong0), set_map0s), rel_OO_Grp) =
-      EVERY' [rtac allI, rtac allI, rtac impI, dtac (mk_conjunctN n i),
-        etac allE, etac allE, etac impE, atac, etac bexE, etac conjE,
-        rtac (rel_OO_Grp RS sym RS @{thm eq_refl} RS @{thm predicate2D}),
-        rtac @{thm relcomppI}, rtac @{thm conversepI},
-        EVERY' (map (fn thm =>
-          EVERY' [rtac @{thm GrpI},
-            rtac trans, rtac trans, rtac map_comp0, rtac map_cong0, REPEAT_DETERM_N m o rtac thm,
-            REPEAT_DETERM_N n o rtac (@{thm o_id} RS fun_cong), atac,
-            REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac CollectI,
-            CONJ_WRAP' (fn (i, thm) =>
-              if i <= m
-              then EVERY' [rtac ord_eq_le_trans, rtac thm, rtac subset_trans,
-                etac @{thm image_mono}, rtac @{thm image_subsetI},
-                etac @{thm Collect_split_in_relI[OF Id_onI]}]
-              else EVERY' [rtac ord_eq_le_trans, rtac trans, rtac thm,
-                rtac trans_fun_cong_image_id_id_apply, etac @{thm Collect_split_in_rel_leI}])
-            (1 upto (m + n) ~~ set_map0s)])
-          @{thms fst_diag_id snd_diag_id})];
-
-    fun mk_only_if_tac ((((i, map_comp0), map_cong0), set_map0s), rel_OO_Grp) =
-      EVERY' [dtac (mk_conjunctN n i), rtac allI, rtac allI, rtac impI,
-        etac allE, etac allE, etac impE, atac,
-        dtac (rel_OO_Grp RS @{thm eq_refl} RS @{thm predicate2D}),
-        REPEAT_DETERM o eresolve_tac ([CollectE, conjE, exE] @
-          @{thms GrpE relcomppE conversepE CollectE Collect_split_in_rel_leE}),
-        hyp_subst_tac ctxt,
-        rtac bexI, rtac conjI, rtac trans, rtac map_comp0, rtac trans, rtac map_cong0,
-        REPEAT_DETERM_N m o rtac (@{thm id_o} RS fun_cong),
-        REPEAT_DETERM_N n o rtac (@{thm o_id} RS fun_cong),
-        atac, rtac trans, rtac map_comp0, rtac trans, rtac map_cong0,
-        REPEAT_DETERM_N m o rtac (@{thm id_o} RS fun_cong),
-        REPEAT_DETERM_N n o rtac (@{thm o_id} RS fun_cong),
-        rtac trans, rtac map_cong0,
-        REPEAT_DETERM_N m o EVERY' [rtac @{thm Id_onD'}, etac set_mp, atac],
-        REPEAT_DETERM_N n o rtac refl,
-        atac, rtac CollectI,
-        CONJ_WRAP' (fn (i, thm) =>
-          if i <= m
-          then EVERY' [rtac ord_eq_le_trans, rtac thm, rtac @{thm image_subsetI},
-            rtac @{thm Id_on_fst}, etac set_mp, atac]
-          else EVERY' [rtac ord_eq_le_trans, rtac trans, rtac thm,
-            rtac trans_fun_cong_image_id_id_apply, atac])
-        (1 upto (m + n) ~~ set_map0s)];
-  in
-    EVERY' [rtac (bis_def RS trans),
-      rtac iffI, etac conjE, etac conjI, CONJ_WRAP' mk_if_tac thms,
-      etac conjE, etac conjI, CONJ_WRAP' mk_only_if_tac thms] 1
-  end;
-
-fun mk_bis_converse_tac m bis_rel rel_congs rel_converseps =
-  EVERY' [stac bis_rel, dtac (bis_rel RS iffD1),
-    REPEAT_DETERM o etac conjE, rtac conjI,
-    CONJ_WRAP' (K (EVERY' [rtac converse_shift, etac subset_trans,
-      rtac equalityD2, rtac @{thm converse_Times}])) rel_congs,
-    CONJ_WRAP' (fn (rel_cong, rel_conversep) =>
-      EVERY' [rtac allI, rtac allI, rtac impI,
-        rtac (rel_cong RS @{thm eq_refl} RS @{thm predicate2D}),
-        REPEAT_DETERM_N m o rtac conversep_in_rel_Id_on,
-        REPEAT_DETERM_N (length rel_congs) o rtac @{thm conversep_in_rel},
-        rtac (rel_conversep RS sym RS @{thm eq_refl} RS @{thm predicate2D}),
-        REPEAT_DETERM o etac allE,
-        rtac @{thm conversepI}, etac mp, etac @{thm converseD}]) (rel_congs ~~ rel_converseps)] 1;
-
-fun mk_bis_O_tac ctxt m bis_rel rel_congs rel_OOs =
-  EVERY' [stac bis_rel, REPEAT_DETERM o dtac (bis_rel RS iffD1),
-    REPEAT_DETERM o etac conjE, rtac conjI,
-    CONJ_WRAP' (K (EVERY' [etac @{thm relcomp_subset_Sigma}, atac])) rel_congs,
-    CONJ_WRAP' (fn (rel_cong, rel_OO) =>
-      EVERY' [rtac allI, rtac allI, rtac impI,
-        rtac (rel_cong RS @{thm eq_refl} RS @{thm predicate2D}),
-        REPEAT_DETERM_N m o rtac relcompp_in_rel_Id_on,
-        REPEAT_DETERM_N (length rel_congs) o rtac @{thm relcompp_in_rel},
-        rtac (rel_OO RS sym RS @{thm eq_refl} RS @{thm predicate2D}),
-        etac @{thm relcompE},
-        REPEAT_DETERM o dtac Pair_eqD,
-        etac conjE, hyp_subst_tac ctxt,
-        REPEAT_DETERM o etac allE, rtac @{thm relcomppI},
-        etac mp, atac, etac mp, atac]) (rel_congs ~~ rel_OOs)] 1;
-
-fun mk_bis_Gr_tac bis_rel rel_Grps mor_images morEs coalg_ins
-  {context = ctxt, prems = _} =
-  unfold_thms_tac ctxt (bis_rel :: @{thm Id_on_Gr} :: @{thm in_rel_Gr} :: rel_Grps) THEN
-  EVERY' [rtac conjI,
-    CONJ_WRAP' (fn thm => rtac (@{thm Gr_incl} RS ssubst) THEN' etac thm) mor_images,
-    CONJ_WRAP' (fn (coalg_in, morE) =>
-      EVERY' [rtac allI, rtac allI, rtac impI, rtac @{thm GrpI}, etac (morE RS trans),
-        etac @{thm GrD1}, etac (@{thm GrD2} RS arg_cong), etac coalg_in, etac @{thm GrD1}])
-    (coalg_ins ~~ morEs)] 1;
-
-fun mk_bis_Union_tac bis_def in_monos {context = ctxt, prems = _} =
-  let
-    val n = length in_monos;
-    val ks = 1 upto n;
-  in
-    unfold_thms_tac ctxt [bis_def] THEN
-    EVERY' [rtac conjI,
-      CONJ_WRAP' (fn i =>
-        EVERY' [rtac @{thm UN_least}, dtac bspec, atac,
-          dtac conjunct1, etac (mk_conjunctN n i)]) ks,
-      CONJ_WRAP' (fn (i, in_mono) =>
-        EVERY' [rtac allI, rtac allI, rtac impI, etac @{thm UN_E}, dtac bspec, atac,
-          dtac conjunct2, dtac (mk_conjunctN n i), etac allE, etac allE, dtac mp,
-          atac, etac bexE, rtac bexI, atac, rtac in_mono,
-          REPEAT_DETERM_N n o etac @{thm SUP_upper2[OF _ subset_refl]},
-          atac]) (ks ~~ in_monos)] 1
-  end;
-
-fun mk_sbis_lsbis_tac ctxt lsbis_defs bis_Union bis_cong =
-  let
-    val n = length lsbis_defs;
-  in
-    EVERY' [rtac (Thm.permute_prems 0 1 bis_cong), EVERY' (map rtac lsbis_defs),
-      rtac bis_Union, rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE, exE],
-      hyp_subst_tac ctxt, etac bis_cong, EVERY' (map (rtac o mk_nth_conv n) (1 upto n))] 1
-  end;
-
-fun mk_incl_lsbis_tac n i lsbis_def =
-  EVERY' [rtac @{thm xt1(3)}, rtac lsbis_def, rtac @{thm SUP_upper2}, rtac CollectI,
-    REPEAT_DETERM_N n o rtac exI, rtac conjI, rtac refl, atac, rtac equalityD2,
-    rtac (mk_nth_conv n i)] 1;
-
-fun mk_equiv_lsbis_tac sbis_lsbis lsbis_incl incl_lsbis bis_Id_on bis_converse bis_O =
-  EVERY' [rtac (@{thm equiv_def} RS iffD2),
-
-    rtac conjI, rtac (@{thm refl_on_def} RS iffD2),
-    rtac conjI, rtac lsbis_incl, rtac ballI, rtac set_mp,
-    rtac incl_lsbis, rtac bis_Id_on, atac, etac @{thm Id_onI},
-
-    rtac conjI, rtac (@{thm sym_def} RS iffD2),
-    rtac allI, rtac allI, rtac impI, rtac set_mp,
-    rtac incl_lsbis, rtac bis_converse, rtac sbis_lsbis, etac @{thm converseI},
-
-    rtac (@{thm trans_def} RS iffD2),
-    rtac allI, rtac allI, rtac allI, rtac impI, rtac impI, rtac set_mp,
-    rtac incl_lsbis, rtac bis_O, rtac sbis_lsbis, rtac sbis_lsbis,
-    etac @{thm relcompI}, atac] 1;
-
-fun mk_coalgT_tac m defs strT_defs set_map0ss {context = ctxt, prems = _} =
-  let
-    val n = length strT_defs;
-    val ks = 1 upto n;
-    fun coalg_tac (i, ((passive_sets, active_sets), def)) =
-      EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE],
-        hyp_subst_tac ctxt, rtac (def RS trans RS @{thm ssubst_mem}), etac (arg_cong RS trans),
-        rtac (mk_sum_casesN n i), rtac CollectI,
-        EVERY' (map (fn thm => EVERY' [rtac conjI, rtac (thm RS ord_eq_le_trans),
-          etac ((trans OF [@{thm image_id} RS fun_cong, id_apply]) RS ord_eq_le_trans)])
-          passive_sets),
-        CONJ_WRAP' (fn (i, thm) => EVERY' [rtac (thm RS ord_eq_le_trans),
-          rtac @{thm image_subsetI}, rtac CollectI, rtac exI, rtac exI, rtac conjI, rtac refl,
-          rtac conjI,
-          rtac conjI, etac @{thm empty_Shift}, dtac set_rev_mp,
-            etac equalityD1, etac CollectD,
-          rtac conjI, etac @{thm Shift_clists},
-          rtac conjI, etac @{thm Shift_prefCl},
-          rtac conjI, rtac ballI,
-            rtac conjI, dtac @{thm iffD1[OF ball_conj_distrib]}, dtac conjunct1,
-            SELECT_GOAL (unfold_thms_tac ctxt @{thms Succ_Shift shift_def}),
-            etac bspec, etac @{thm ShiftD},
-            CONJ_WRAP' (fn i => EVERY' [rtac ballI, etac CollectE, dtac @{thm ShiftD},
-              dtac bspec, etac thin_rl, atac, dtac conjunct2, dtac (mk_conjunctN n i),
-              dtac bspec, rtac CollectI, etac @{thm set_mp[OF equalityD1[OF Succ_Shift]]},
-              REPEAT_DETERM o eresolve_tac [exE, conjE], rtac exI,
-              rtac conjI, rtac (@{thm shift_def} RS fun_cong RS trans),
-              rtac (@{thm append_Cons} RS sym RS arg_cong RS trans), atac,
-              REPEAT_DETERM_N m o (rtac conjI THEN' atac),
-              CONJ_WRAP' (K (EVERY' [etac trans, rtac @{thm Collect_cong},
-                rtac @{thm eqset_imp_iff}, rtac sym, rtac trans, rtac @{thm Succ_Shift},
-                rtac (@{thm append_Cons} RS sym RS arg_cong)])) ks]) ks,
-          rtac allI, rtac impI, REPEAT_DETERM o eresolve_tac [allE, impE],
-          etac @{thm not_in_Shift}, rtac trans, rtac (@{thm shift_def} RS fun_cong), atac,
-          dtac bspec, atac, dtac conjunct2, dtac (mk_conjunctN n i), dtac bspec,
-          etac @{thm set_mp[OF equalityD1]}, atac,
-          REPEAT_DETERM o eresolve_tac [exE, conjE], rtac exI,
-          rtac conjI, rtac (@{thm shift_def} RS fun_cong RS trans),
-          etac (@{thm append_Nil} RS sym RS arg_cong RS trans),
-          REPEAT_DETERM_N m o (rtac conjI THEN' atac),
-          CONJ_WRAP' (K (EVERY' [etac trans, rtac @{thm Collect_cong},
-            rtac @{thm eqset_imp_iff}, rtac sym, rtac trans, rtac @{thm Succ_Shift},
-            rtac (@{thm append_Nil} RS sym RS arg_cong)])) ks]) (ks ~~ active_sets)];
-  in
-    unfold_thms_tac ctxt defs THEN
-    CONJ_WRAP' coalg_tac (ks ~~ (map (chop m) set_map0ss ~~ strT_defs)) 1
-  end;
-
-fun mk_Lev_sbd_tac ctxt cts Lev_0s Lev_Sucs to_sbdss =
-  let
-    val n = length Lev_0s;
-    val ks = 1 upto n;
-  in
-    EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
-      REPEAT_DETERM o rtac allI,
-      CONJ_WRAP' (fn Lev_0 =>
-        EVERY' (map rtac [ord_eq_le_trans, Lev_0, @{thm Nil_clists}])) Lev_0s,
-      REPEAT_DETERM o rtac allI,
-      CONJ_WRAP' (fn (Lev_Suc, to_sbds) =>
-        EVERY' [rtac ord_eq_le_trans, rtac Lev_Suc,
-          CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least}))
-            (fn (i, to_sbd) => EVERY' [rtac subsetI,
-              REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac ctxt,
-              rtac @{thm Cons_clists}, rtac (mk_InN_Field n i), etac to_sbd,
-              etac set_rev_mp, REPEAT_DETERM o etac allE,
-              etac (mk_conjunctN n i)])
-          (rev (ks ~~ to_sbds))])
-      (Lev_Sucs ~~ to_sbdss)] 1
-  end;
-
-fun mk_length_Lev_tac ctxt cts Lev_0s Lev_Sucs =
-  let
-    val n = length Lev_0s;
-    val ks = n downto 1;
-  in
-    EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
-      REPEAT_DETERM o rtac allI,
-      CONJ_WRAP' (fn Lev_0 =>
-        EVERY' [rtac impI, dtac (Lev_0 RS equalityD1 RS set_mp),
-          etac @{thm singletonE}, etac ssubst, rtac @{thm list.size(3)}]) Lev_0s,
-      REPEAT_DETERM o rtac allI,
-      CONJ_WRAP' (fn Lev_Suc =>
-        EVERY' [rtac impI, dtac (Lev_Suc RS equalityD1 RS set_mp),
-          CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
-            (fn i =>
-              EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac ctxt,
-                rtac trans, rtac @{thm length_Cons}, rtac @{thm arg_cong[of _ _ Suc]},
-                REPEAT_DETERM o etac allE, dtac (mk_conjunctN n i), etac mp, atac]) ks])
-      Lev_Sucs] 1
-  end;
-
-fun mk_length_Lev'_tac length_Lev =
-  EVERY' [ftac length_Lev, etac ssubst, atac] 1;
-
-fun mk_prefCl_Lev_tac ctxt cts Lev_0s Lev_Sucs =
-  let
-    val n = length Lev_0s;
-    val ks = n downto 1;
-  in
-    EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
-      REPEAT_DETERM o rtac allI,
-      CONJ_WRAP' (fn Lev_0 =>
-        EVERY' [rtac impI, etac conjE, dtac (Lev_0 RS equalityD1 RS set_mp),
-          etac @{thm singletonE}, hyp_subst_tac ctxt,
-          dtac @{thm prefixeq_Nil[THEN subst, of "%x. x"]},
-          hyp_subst_tac ctxt,
-          rtac @{thm set_mp[OF equalityD2[OF trans[OF arg_cong[OF list.size(3)]]]]},
-          rtac Lev_0, rtac @{thm singletonI}]) Lev_0s,
-      REPEAT_DETERM o rtac allI,
-      CONJ_WRAP' (fn (Lev_0, Lev_Suc) =>
-        EVERY' [rtac impI, etac conjE, dtac (Lev_Suc RS equalityD1 RS set_mp),
-          CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
-            (fn i =>
-              EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac ctxt,
-              dtac @{thm prefixeq_Cons[THEN subst, of "%x. x"]}, etac disjE, hyp_subst_tac ctxt,
-              rtac @{thm set_mp[OF equalityD2[OF trans[OF arg_cong[OF list.size(3)]]]]},
-              rtac Lev_0, rtac @{thm singletonI},
-              REPEAT_DETERM o eresolve_tac [exE, conjE], hyp_subst_tac ctxt,
-              rtac @{thm set_mp[OF equalityD2[OF trans[OF arg_cong[OF length_Cons]]]]},
-              rtac Lev_Suc, rtac (mk_UnIN n i), rtac CollectI, REPEAT_DETERM o rtac exI, rtac conjI,
-              rtac refl, etac conjI, REPEAT_DETERM o etac allE, dtac (mk_conjunctN n i),
-              etac mp, etac conjI, atac]) ks])
-      (Lev_0s ~~ Lev_Sucs)] 1
-  end;
-
-fun mk_rv_last_tac cTs cts rv_Nils rv_Conss =
-  let
-    val n = length rv_Nils;
-    val ks = 1 upto n;
-  in
-    EVERY' [rtac (Drule.instantiate' cTs cts @{thm list.induct}),
-      REPEAT_DETERM o rtac allI,
-      CONJ_WRAP' (fn rv_Cons =>
-        CONJ_WRAP' (fn (i, rv_Nil) => (EVERY' [rtac exI,
-          rtac (@{thm append_Nil} RS arg_cong RS trans),
-          rtac (rv_Cons RS trans), rtac (mk_sum_casesN n i RS arg_cong RS trans), rtac rv_Nil]))
-        (ks ~~ rv_Nils))
-      rv_Conss,
-      REPEAT_DETERM o rtac allI, rtac (mk_sumEN n),
-      EVERY' (map (fn i =>
-        CONJ_WRAP' (fn rv_Cons => EVERY' [REPEAT_DETERM o etac allE, dtac (mk_conjunctN n i),
-          CONJ_WRAP' (fn i' => EVERY' [dtac (mk_conjunctN n i'), etac exE, rtac exI,
-            rtac (@{thm append_Cons} RS arg_cong RS trans),
-            rtac (rv_Cons RS trans), etac (sum_case_weak_cong RS arg_cong RS trans),
-            rtac (mk_sum_casesN n i RS arg_cong RS trans), atac])
-          ks])
-        rv_Conss)
-      ks)] 1
-  end;
-
-fun mk_set_rv_Lev_tac ctxt m cts Lev_0s Lev_Sucs rv_Nils rv_Conss coalg_setss from_to_sbdss =
-  let
-    val n = length Lev_0s;
-    val ks = 1 upto n;
-  in
-    EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
-      REPEAT_DETERM o rtac allI,
-      CONJ_WRAP' (fn (i, ((Lev_0, rv_Nil), coalg_sets)) =>
-        EVERY' [rtac impI, REPEAT_DETERM o etac conjE,
-          dtac (Lev_0 RS equalityD1 RS set_mp), etac @{thm singletonE}, etac ssubst,
-          rtac (rv_Nil RS arg_cong RS iffD2),
-          rtac (mk_sum_casesN n i RS iffD2),
-          CONJ_WRAP' (fn thm => etac thm THEN' atac) (take m coalg_sets)])
-      (ks ~~ ((Lev_0s ~~ rv_Nils) ~~ coalg_setss)),
-      REPEAT_DETERM o rtac allI,
-      CONJ_WRAP' (fn ((Lev_Suc, rv_Cons), (from_to_sbds, coalg_sets)) =>
-        EVERY' [rtac impI, etac conjE, dtac (Lev_Suc RS equalityD1 RS set_mp),
-          CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
-            (fn (i, (from_to_sbd, coalg_set)) =>
-              EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac ctxt,
-              rtac (rv_Cons RS arg_cong RS iffD2),
-              rtac (mk_sum_casesN n i RS arg_cong RS trans RS iffD2),
-              etac (from_to_sbd RS arg_cong), REPEAT_DETERM o etac allE,
-              dtac (mk_conjunctN n i), etac mp, etac conjI, etac set_rev_mp,
-              etac coalg_set, atac])
-          (rev (ks ~~ (from_to_sbds ~~ drop m coalg_sets)))])
-      ((Lev_Sucs ~~ rv_Conss) ~~ (from_to_sbdss ~~ coalg_setss))] 1
-  end;
-
-fun mk_set_Lev_tac ctxt cts Lev_0s Lev_Sucs rv_Nils rv_Conss from_to_sbdss =
-  let
-    val n = length Lev_0s;
-    val ks = 1 upto n;
-  in
-    EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
-      REPEAT_DETERM o rtac allI,
-      CONJ_WRAP' (fn ((i, (Lev_0, Lev_Suc)), rv_Nil) =>
-        EVERY' [rtac impI, dtac (Lev_0 RS equalityD1 RS set_mp),
-          etac @{thm singletonE}, hyp_subst_tac ctxt,
-          CONJ_WRAP' (fn i' => rtac impI THEN' dtac (sym RS trans) THEN' rtac rv_Nil THEN'
-            (if i = i'
-            then EVERY' [dtac (mk_InN_inject n i), hyp_subst_tac ctxt,
-              CONJ_WRAP' (fn (i'', Lev_0'') =>
-                EVERY' [rtac impI, rtac @{thm ssubst_mem[OF append_Nil]},
-                  rtac (Lev_Suc RS equalityD2 RS set_mp), rtac (mk_UnIN n i''),
-                  rtac CollectI, REPEAT_DETERM o rtac exI, rtac conjI, rtac refl,
-                  etac conjI, rtac (Lev_0'' RS equalityD2 RS set_mp),
-                  rtac @{thm singletonI}])
-              (ks ~~ Lev_0s)]
-            else etac (mk_InN_not_InM i' i RS notE)))
-          ks])
-      ((ks ~~ (Lev_0s ~~ Lev_Sucs)) ~~ rv_Nils),
-      REPEAT_DETERM o rtac allI,
-      CONJ_WRAP' (fn ((Lev_Suc, rv_Cons), from_to_sbds) =>
-        EVERY' [rtac impI, dtac (Lev_Suc RS equalityD1 RS set_mp),
-          CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
-            (fn (i, from_to_sbd) =>
-              EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac ctxt,
-                CONJ_WRAP' (fn i' => rtac impI THEN'
-                  CONJ_WRAP' (fn i'' =>
-                    EVERY' [rtac impI, rtac (Lev_Suc RS equalityD2 RS set_mp),
-                      rtac @{thm ssubst_mem[OF append_Cons]}, rtac (mk_UnIN n i),
-                      rtac CollectI, REPEAT_DETERM o rtac exI, rtac conjI, rtac refl,
-                      rtac conjI, atac, dtac (sym RS trans RS sym),
-                      rtac (rv_Cons RS trans), rtac (mk_sum_casesN n i RS trans),
-                      etac (from_to_sbd RS arg_cong), REPEAT_DETERM o etac allE,
-                      dtac (mk_conjunctN n i), dtac mp, atac,
-                      dtac (mk_conjunctN n i'), dtac mp, atac,
-                      dtac (mk_conjunctN n i''), etac mp, atac])
-                  ks)
-                ks])
-          (rev (ks ~~ from_to_sbds))])
-      ((Lev_Sucs ~~ rv_Conss) ~~ from_to_sbdss)] 1
-  end;
-
-fun mk_set_image_Lev_tac ctxt cts Lev_0s Lev_Sucs rv_Nils rv_Conss from_to_sbdss to_sbd_injss =
-  let
-    val n = length Lev_0s;
-    val ks = 1 upto n;
-  in
-    EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
-      REPEAT_DETERM o rtac allI,
-      CONJ_WRAP' (fn ((i, (Lev_0, Lev_Suc)), rv_Nil) =>
-        EVERY' [rtac impI, dtac (Lev_0 RS equalityD1 RS set_mp),
-          etac @{thm singletonE}, hyp_subst_tac ctxt,
-          CONJ_WRAP' (fn i' => rtac impI THEN'
-            CONJ_WRAP' (fn i'' => rtac impI  THEN' dtac (sym RS trans) THEN' rtac rv_Nil THEN'
-              (if i = i''
-              then EVERY' [dtac @{thm ssubst_mem[OF sym[OF append_Nil]]},
-                dtac (Lev_Suc RS equalityD1 RS set_mp), dtac (mk_InN_inject n i),
-                hyp_subst_tac ctxt,
-                CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
-                  (fn k => REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE] THEN'
-                    dtac list_inject_iffD1 THEN' etac conjE THEN'
-                    (if k = i'
-                    then EVERY' [dtac (mk_InN_inject n k), hyp_subst_tac ctxt, etac imageI]
-                    else etac (mk_InN_not_InM i' k RS notE)))
-                (rev ks)]
-              else etac (mk_InN_not_InM i'' i RS notE)))
-            ks)
-          ks])
-      ((ks ~~ (Lev_0s ~~ Lev_Sucs)) ~~ rv_Nils),
-      REPEAT_DETERM o rtac allI,
-      CONJ_WRAP' (fn ((Lev_Suc, rv_Cons), (from_to_sbds, to_sbd_injs)) =>
-        EVERY' [rtac impI, dtac (Lev_Suc RS equalityD1 RS set_mp),
-          CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
-            (fn (i, (from_to_sbd, to_sbd_inj)) =>
-              REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE] THEN' hyp_subst_tac ctxt THEN'
-              CONJ_WRAP' (fn i' => rtac impI THEN'
-                dtac @{thm ssubst_mem[OF sym[OF append_Cons]]} THEN'
-                dtac (Lev_Suc RS equalityD1 RS set_mp) THEN'
-                CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE)) (fn k =>
-                  REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE] THEN'
-                  dtac list_inject_iffD1 THEN' etac conjE THEN'
-                  (if k = i
-                  then EVERY' [dtac (mk_InN_inject n i),
-                    dtac (Thm.permute_prems 0 2 (to_sbd_inj RS iffD1)),
-                    atac, atac, hyp_subst_tac ctxt] THEN'
-                    CONJ_WRAP' (fn i'' =>
-                      EVERY' [rtac impI, dtac (sym RS trans),
-                        rtac (rv_Cons RS trans), rtac (mk_sum_casesN n i RS arg_cong RS trans),
-                        etac (from_to_sbd RS arg_cong),
-                        REPEAT_DETERM o etac allE,
-                        dtac (mk_conjunctN n i), dtac mp, atac,
-                        dtac (mk_conjunctN n i'), dtac mp, atac,
-                        dtac (mk_conjunctN n i''), etac mp, etac sym])
-                    ks
-                  else etac (mk_InN_not_InM i k RS notE)))
-                (rev ks))
-              ks)
-          (rev (ks ~~ (from_to_sbds ~~ to_sbd_injs)))])
-      ((Lev_Sucs ~~ rv_Conss) ~~ (from_to_sbdss ~~ to_sbd_injss))] 1
-  end;
-
-fun mk_mor_beh_tac m mor_def mor_cong beh_defs carT_defs strT_defs isNode_defs
-  to_sbd_injss from_to_sbdss Lev_0s Lev_Sucs rv_Nils rv_Conss Lev_sbds length_Levs length_Lev's
-  prefCl_Levs rv_lastss set_rv_Levsss set_Levsss set_image_Levsss set_map0ss coalg_setss
-  map_comp_ids map_cong0s map_arg_congs {context = ctxt, prems = _} =
-  let
-    val n = length beh_defs;
-    val ks = 1 upto n;
-
-    fun fbetw_tac (i, (carT_def, (isNode_def, (Lev_0, (rv_Nil, (Lev_sbd,
-      ((length_Lev, length_Lev'), (prefCl_Lev, (rv_lasts, (set_map0s,
-        (coalg_sets, (set_rv_Levss, (set_Levss, set_image_Levss))))))))))))) =
-      EVERY' [rtac ballI, rtac (carT_def RS equalityD2 RS set_mp),
-        rtac CollectI, REPEAT_DETERM o rtac exI, rtac conjI, rtac refl, rtac conjI,
-        rtac conjI,
-          rtac @{thm UN_I}, rtac UNIV_I, rtac (Lev_0 RS equalityD2 RS set_mp),
-          rtac @{thm singletonI},
-        rtac conjI,
-          rtac @{thm UN_least}, rtac Lev_sbd,
-        rtac conjI,
-          rtac @{thm prefCl_UN}, rtac ssubst, rtac @{thm PrefCl_def}, REPEAT_DETERM o rtac allI,
-          rtac impI, etac conjE, rtac exI, rtac conjI, rtac @{thm ord_le_eq_trans},
-          etac @{thm prefixeq_length_le}, etac length_Lev, rtac prefCl_Lev, etac conjI, atac,
-        rtac conjI,
-          rtac ballI, etac @{thm UN_E}, rtac conjI,
-          if n = 1 then K all_tac else rtac (mk_sumEN n),
-          EVERY' (map6 (fn i => fn isNode_def => fn set_map0s =>
-            fn set_rv_Levs => fn set_Levs => fn set_image_Levs =>
-            EVERY' [rtac (mk_disjIN n i), rtac (isNode_def RS ssubst),
-              rtac exI, rtac conjI,
-              (if n = 1 then rtac @{thm if_P} THEN' etac length_Lev'
-              else rtac (@{thm if_P} RS arg_cong RS trans) THEN' etac length_Lev' THEN'
-                etac (sum_case_weak_cong RS trans) THEN' rtac (mk_sum_casesN n i)),
-              EVERY' (map2 (fn set_map0 => fn set_rv_Lev =>
-                EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac (set_map0 RS trans),
-                  rtac trans_fun_cong_image_id_id_apply,
-                  etac set_rv_Lev, TRY o atac, etac conjI, atac])
-              (take m set_map0s) set_rv_Levs),
-              CONJ_WRAP' (fn (set_map0, (set_Lev, set_image_Lev)) =>
-                EVERY' [rtac (set_map0 RS trans), rtac equalityI, rtac @{thm image_subsetI},
-                  rtac CollectI, rtac @{thm SuccI}, rtac @{thm UN_I}, rtac UNIV_I, etac set_Lev,
-                  if n = 1 then rtac refl else atac, atac, rtac subsetI,
-                  REPEAT_DETERM o eresolve_tac [CollectE, @{thm SuccE}, @{thm UN_E}],
-                  rtac set_image_Lev, atac, dtac length_Lev, hyp_subst_tac ctxt, dtac length_Lev',
-                  etac @{thm set_mp[OF equalityD1[OF arg_cong[OF length_append_singleton]]]},
-                  if n = 1 then rtac refl else atac])
-              (drop m set_map0s ~~ (set_Levs ~~ set_image_Levs))])
-          ks isNode_defs set_map0ss set_rv_Levss set_Levss set_image_Levss),
-          CONJ_WRAP' (fn (i, (rv_last, (isNode_def, (set_map0s,
-            (set_rv_Levs, (set_Levs, set_image_Levs)))))) =>
-            EVERY' [rtac ballI,
-              REPEAT_DETERM o eresolve_tac [CollectE, @{thm SuccE}, @{thm UN_E}],
-              rtac (rev_mp OF [rv_last, impI]), etac exE, rtac (isNode_def RS ssubst),
-              rtac exI, rtac conjI,
-              (if n = 1 then rtac @{thm if_P} THEN' etac length_Lev'
-              else rtac (@{thm if_P} RS trans) THEN' etac length_Lev' THEN'
-                etac (sum_case_weak_cong RS trans) THEN' rtac (mk_sum_casesN n i)),
-              EVERY' (map2 (fn set_map0 => fn set_rv_Lev =>
-                EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac (set_map0 RS trans),
-                  rtac trans_fun_cong_image_id_id_apply,
-                  etac set_rv_Lev, TRY o atac, etac conjI, atac])
-              (take m set_map0s) set_rv_Levs),
-              CONJ_WRAP' (fn (set_map0, (set_Lev, set_image_Lev)) =>
-                EVERY' [rtac (set_map0 RS trans), rtac equalityI, rtac @{thm image_subsetI},
-                  rtac CollectI, rtac @{thm SuccI}, rtac @{thm UN_I}, rtac UNIV_I, etac set_Lev,
-                  if n = 1 then rtac refl else atac, atac, rtac subsetI,
-                  REPEAT_DETERM o eresolve_tac [CollectE, @{thm SuccE}, @{thm UN_E}],
-                  REPEAT_DETERM_N 4 o etac thin_rl,
-                  rtac set_image_Lev,
-                  atac, dtac length_Lev, hyp_subst_tac ctxt, dtac length_Lev',
-                  etac @{thm set_mp[OF equalityD1[OF arg_cong[OF length_append_singleton]]]},
-                  if n = 1 then rtac refl else atac])
-              (drop m set_map0s ~~ (set_Levs ~~ set_image_Levs))])
-          (ks ~~ (rv_lasts ~~ (isNode_defs ~~ (set_map0ss ~~
-            (set_rv_Levss ~~ (set_Levss ~~ set_image_Levss)))))),
-        (**)
-          rtac allI, rtac impI, rtac @{thm if_not_P}, rtac notI,
-          etac notE, etac @{thm UN_I[OF UNIV_I]},
-        (*root isNode*)
-          rtac (isNode_def RS ssubst), rtac exI, rtac conjI, rtac (@{thm if_P} RS trans),
-          rtac length_Lev', rtac (Lev_0 RS equalityD2 RS set_mp), rtac @{thm singletonI},
-          CONVERSION (Conv.top_conv
-            (K (Conv.try_conv (Conv.rewr_conv (rv_Nil RS eq_reflection)))) ctxt),
-          if n = 1 then rtac refl else rtac (mk_sum_casesN n i),
-          EVERY' (map2 (fn set_map0 => fn coalg_set =>
-            EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac (set_map0 RS trans),
-              rtac trans_fun_cong_image_id_id_apply, etac coalg_set, atac])
-            (take m set_map0s) (take m coalg_sets)),
-          CONJ_WRAP' (fn (set_map0, (set_Lev, set_image_Lev)) =>
-            EVERY' [rtac (set_map0 RS trans), rtac equalityI, rtac @{thm image_subsetI},
-              rtac CollectI, rtac @{thm SuccI}, rtac @{thm UN_I}, rtac UNIV_I, rtac set_Lev,
-              rtac (Lev_0 RS equalityD2 RS set_mp), rtac @{thm singletonI}, rtac rv_Nil,
-              atac, rtac subsetI,
-              REPEAT_DETERM o eresolve_tac [CollectE, @{thm SuccE}, @{thm UN_E}],
-              rtac set_image_Lev, rtac (Lev_0 RS equalityD2 RS set_mp),
-              rtac @{thm singletonI}, dtac length_Lev',
-              etac @{thm set_mp[OF equalityD1[OF arg_cong[OF
-                trans[OF length_append_singleton arg_cong[of _ _ Suc, OF list.size(3)]]]]]},
-              rtac rv_Nil])
-          (drop m set_map0s ~~ (nth set_Levss (i - 1) ~~ nth set_image_Levss (i - 1)))];
-
-    fun mor_tac (i, (strT_def, (((Lev_0, Lev_Suc), (rv_Nil, rv_Cons)),
-      ((map_comp_id, (map_cong0, map_arg_cong)), (length_Lev', (from_to_sbds, to_sbd_injs)))))) =
-      EVERY' [rtac ballI, rtac sym, rtac trans, rtac strT_def,
-        rtac (@{thm if_P} RS
-          (if n = 1 then map_arg_cong else sum_case_weak_cong) RS trans),
-        rtac (@{thm list.size(3)} RS arg_cong RS trans RS equalityD2 RS set_mp),
-        rtac Lev_0, rtac @{thm singletonI},
-        CONVERSION (Conv.top_conv
-          (K (Conv.try_conv (Conv.rewr_conv (rv_Nil RS eq_reflection)))) ctxt),
-        if n = 1 then K all_tac
-        else (rtac (sum_case_weak_cong RS trans) THEN'
-          rtac (mk_sum_casesN n i) THEN' rtac (mk_sum_casesN n i RS trans)),
-        rtac (map_comp_id RS trans), rtac (map_cong0 OF replicate m refl),
-        EVERY' (map3 (fn i' => fn to_sbd_inj => fn from_to_sbd =>
-          DETERM o EVERY' [rtac trans, rtac o_apply, rtac Pair_eqI, rtac conjI,
-            rtac trans, rtac @{thm Shift_def},
-            rtac equalityI, rtac subsetI, etac thin_rl, etac thin_rl,
-            REPEAT_DETERM o eresolve_tac [CollectE, @{thm UN_E}], dtac length_Lev', dtac asm_rl,
-            etac thin_rl, dtac @{thm set_rev_mp[OF _ equalityD1]},
-            rtac (@{thm length_Cons} RS arg_cong RS trans), rtac Lev_Suc,
-            CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE)) (fn i'' =>
-              EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE],
-                dtac list_inject_iffD1, etac conjE,
-                if i' = i'' then EVERY' [dtac (mk_InN_inject n i'),
-                  dtac (Thm.permute_prems 0 2 (to_sbd_inj RS iffD1)),
-                  atac, atac, hyp_subst_tac ctxt, etac @{thm UN_I[OF UNIV_I]}]
-                else etac (mk_InN_not_InM i' i'' RS notE)])
-            (rev ks),
-            rtac @{thm UN_least}, rtac subsetI, rtac CollectI, rtac @{thm UN_I[OF UNIV_I]},
-            rtac (Lev_Suc RS equalityD2 RS set_mp), rtac (mk_UnIN n i'), rtac CollectI,
-            REPEAT_DETERM o rtac exI, rtac conjI, rtac refl, etac conjI, atac,
-            rtac trans, rtac @{thm shift_def}, rtac ssubst, rtac @{thm fun_eq_iff}, rtac allI,
-            rtac @{thm if_cong}, rtac (@{thm length_Cons} RS arg_cong RS trans), rtac iffI,
-            dtac asm_rl, dtac asm_rl, dtac asm_rl,
-            dtac (Lev_Suc RS equalityD1 RS set_mp),
-            CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE)) (fn i'' =>
-              EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE],
-                dtac list_inject_iffD1, etac conjE,
-                if i' = i'' then EVERY' [dtac (mk_InN_inject n i'),
-                  dtac (Thm.permute_prems 0 2 (to_sbd_inj RS iffD1)),
-                  atac, atac, hyp_subst_tac ctxt, atac]
-                else etac (mk_InN_not_InM i' i'' RS notE)])
-            (rev ks),
-            rtac (Lev_Suc RS equalityD2 RS set_mp), rtac (mk_UnIN n i'), rtac CollectI,
-            REPEAT_DETERM o rtac exI, rtac conjI, rtac refl, etac conjI, atac,
-            CONVERSION (Conv.top_conv
-              (K (Conv.try_conv (Conv.rewr_conv (rv_Cons RS eq_reflection)))) ctxt),
-            if n = 1 then K all_tac
-            else rtac sum_case_weak_cong THEN' rtac (mk_sum_casesN n i' RS trans),
-            SELECT_GOAL (unfold_thms_tac ctxt [from_to_sbd]), rtac refl,
-            rtac refl])
-        ks to_sbd_injs from_to_sbds)];
-  in
-    (rtac mor_cong THEN'
-    EVERY' (map (fn thm => rtac (thm RS ext)) beh_defs) THEN'
-    stac mor_def THEN' rtac conjI THEN'
-    CONJ_WRAP' fbetw_tac
-      (ks ~~ (carT_defs ~~ (isNode_defs ~~ (Lev_0s ~~ (rv_Nils ~~ (Lev_sbds ~~
-        ((length_Levs ~~ length_Lev's) ~~ (prefCl_Levs ~~ (rv_lastss ~~
-          (set_map0ss ~~ (coalg_setss ~~
-            (set_rv_Levsss ~~ (set_Levsss ~~ set_image_Levsss))))))))))))) THEN'
-    CONJ_WRAP' mor_tac
-      (ks ~~ (strT_defs ~~ (((Lev_0s ~~ Lev_Sucs) ~~ (rv_Nils ~~ rv_Conss)) ~~
-        ((map_comp_ids ~~ (map_cong0s ~~ map_arg_congs)) ~~
-          (length_Lev's ~~ (from_to_sbdss ~~ to_sbd_injss))))))) 1
-  end;
-
-fun mk_congruent_str_final_tac m lsbisE map_comp_id map_cong0 equiv_LSBISs =
-  EVERY' [rtac @{thm congruentI}, dtac lsbisE,
-    REPEAT_DETERM o eresolve_tac [CollectE, conjE, bexE], rtac (o_apply RS trans),
-    etac (sym RS arg_cong RS trans), rtac (map_comp_id RS trans),
-    rtac (map_cong0 RS trans), REPEAT_DETERM_N m o rtac refl,
-    EVERY' (map (fn equiv_LSBIS =>
-      EVERY' [rtac @{thm equiv_proj}, rtac equiv_LSBIS, etac set_mp, atac])
-    equiv_LSBISs), rtac sym, rtac (o_apply RS trans),
-    etac (sym RS arg_cong RS trans), rtac map_comp_id] 1;
-
-fun mk_coalg_final_tac m coalg_def congruent_str_finals equiv_LSBISs set_map0ss coalgT_setss =
-  EVERY' [stac coalg_def,
-    CONJ_WRAP' (fn ((set_map0s, coalgT_sets), (equiv_LSBIS, congruent_str_final)) =>
-      EVERY' [rtac @{thm univ_preserves}, rtac equiv_LSBIS, rtac congruent_str_final,
-        rtac ballI, rtac @{thm ssubst_mem}, rtac o_apply, rtac CollectI,
-        EVERY' (map2 (fn set_map0 => fn coalgT_set =>
-          EVERY' [rtac conjI, rtac (set_map0 RS ord_eq_le_trans),
-            rtac ord_eq_le_trans_trans_fun_cong_image_id_id_apply,
-            etac coalgT_set])
-        (take m set_map0s) (take m coalgT_sets)),
-        CONJ_WRAP' (fn (equiv_LSBIS, (set_map0, coalgT_set)) =>
-          EVERY' [rtac (set_map0 RS ord_eq_le_trans),
-            rtac @{thm image_subsetI}, rtac ssubst, rtac @{thm proj_in_iff},
-            rtac equiv_LSBIS, etac set_rev_mp, etac coalgT_set])
-        (equiv_LSBISs ~~ drop m (set_map0s ~~ coalgT_sets))])
-    ((set_map0ss ~~ coalgT_setss) ~~ (equiv_LSBISs ~~ congruent_str_finals))] 1;
-
-fun mk_mor_T_final_tac mor_def congruent_str_finals equiv_LSBISs =
-  EVERY' [stac mor_def, rtac conjI,
-    CONJ_WRAP' (fn equiv_LSBIS =>
-      EVERY' [rtac ballI, rtac ssubst, rtac @{thm proj_in_iff}, rtac equiv_LSBIS, atac])
-    equiv_LSBISs,
-    CONJ_WRAP' (fn (equiv_LSBIS, congruent_str_final) =>
-      EVERY' [rtac ballI, rtac sym, rtac trans, rtac @{thm univ_commute}, rtac equiv_LSBIS,
-        rtac congruent_str_final, atac, rtac o_apply])
-    (equiv_LSBISs ~~ congruent_str_finals)] 1;
-
-fun mk_mor_Rep_tac m defs Reps Abs_inverses coalg_final_setss map_comp_ids map_cong0Ls
-  {context = ctxt, prems = _} =
-  unfold_thms_tac ctxt defs THEN
-  EVERY' [rtac conjI,
-    CONJ_WRAP' (fn thm => rtac ballI THEN' rtac thm) Reps,
-    CONJ_WRAP' (fn (Rep, ((map_comp_id, map_cong0L), coalg_final_sets)) =>
-      EVERY' [rtac ballI, rtac (map_comp_id RS trans), rtac map_cong0L,
-        EVERY' (map2 (fn Abs_inverse => fn coalg_final_set =>
-          EVERY' [rtac ballI, rtac (o_apply RS trans), rtac Abs_inverse,
-            etac set_rev_mp, rtac coalg_final_set, rtac Rep])
-        Abs_inverses (drop m coalg_final_sets))])
-    (Reps ~~ ((map_comp_ids ~~ map_cong0Ls) ~~ coalg_final_setss))] 1;
-
-fun mk_mor_Abs_tac defs Abs_inverses {context = ctxt, prems = _} =
-  unfold_thms_tac ctxt defs THEN
-  EVERY' [rtac conjI,
-    CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) Abs_inverses,
-    CONJ_WRAP' (fn thm => rtac ballI THEN' etac (thm RS arg_cong RS sym)) Abs_inverses] 1;
-
-fun mk_mor_unfold_tac m mor_UNIV dtor_defs unfold_defs Abs_inverses morEs map_comp_ids map_cong0s =
-  EVERY' [rtac iffD2, rtac mor_UNIV,
-    CONJ_WRAP' (fn ((Abs_inverse, morE), ((dtor_def, unfold_def), (map_comp_id, map_cong0))) =>
-      EVERY' [rtac ext, rtac (o_apply RS trans RS sym), rtac (dtor_def RS trans),
-        rtac (unfold_def RS arg_cong RS trans), rtac (Abs_inverse RS arg_cong RS trans),
-        rtac (morE RS arg_cong RS trans), rtac (map_comp_id RS trans),
-        rtac (o_apply RS trans RS sym), rtac map_cong0,
-        REPEAT_DETERM_N m o rtac refl,
-        EVERY' (map (fn thm => rtac (thm RS trans) THEN' rtac (o_apply RS sym)) unfold_defs)])
-    ((Abs_inverses ~~ morEs) ~~ ((dtor_defs ~~ unfold_defs) ~~ (map_comp_ids ~~ map_cong0s)))] 1;
-
-fun mk_raw_coind_tac bis_def bis_cong bis_O bis_converse bis_Gr tcoalg coalgT mor_T_final
-  sbis_lsbis lsbis_incls incl_lsbiss equiv_LSBISs mor_Rep Rep_injects =
-  let
-    val n = length Rep_injects;
-  in
-    EVERY' [rtac rev_mp, ftac (bis_def RS iffD1),
-      REPEAT_DETERM o etac conjE, rtac bis_cong, rtac bis_O, rtac bis_converse,
-      rtac bis_Gr, rtac tcoalg, rtac mor_Rep, rtac bis_O, atac, rtac bis_Gr, rtac tcoalg,
-      rtac mor_Rep, REPEAT_DETERM_N n o etac @{thm relImage_Gr},
-      rtac impI, rtac rev_mp, rtac bis_cong, rtac bis_O, rtac bis_Gr, rtac coalgT,
-      rtac mor_T_final, rtac bis_O, rtac sbis_lsbis, rtac bis_converse, rtac bis_Gr, rtac coalgT,
-      rtac mor_T_final, EVERY' (map (fn thm => rtac (thm RS @{thm relInvImage_Gr})) lsbis_incls),
-      rtac impI,
-      CONJ_WRAP' (fn (Rep_inject, (equiv_LSBIS , (incl_lsbis, lsbis_incl))) =>
-        EVERY' [rtac subset_trans, rtac @{thm relInvImage_UNIV_relImage}, rtac subset_trans,
-          rtac @{thm relInvImage_mono}, rtac subset_trans, etac incl_lsbis,
-          rtac ord_eq_le_trans, rtac @{thm sym[OF relImage_relInvImage]},
-          rtac @{thm xt1(3)}, rtac @{thm Sigma_cong},
-          rtac @{thm proj_image}, rtac @{thm proj_image}, rtac lsbis_incl,
-          rtac subset_trans, rtac @{thm relImage_mono}, rtac incl_lsbis, atac,
-          rtac @{thm relImage_proj}, rtac equiv_LSBIS, rtac @{thm relInvImage_Id_on},
-          rtac Rep_inject])
-      (Rep_injects ~~ (equiv_LSBISs ~~ (incl_lsbiss ~~ lsbis_incls)))] 1
-  end;
-
-fun mk_unfold_unique_mor_tac raw_coinds bis mor unfold_defs =
-  CONJ_WRAP' (fn (raw_coind, unfold_def) =>
-    EVERY' [rtac ext, etac (bis RS raw_coind RS set_mp RS @{thm IdD}), rtac mor,
-      rtac @{thm image2_eqI}, rtac refl, rtac (unfold_def RS arg_cong RS trans),
-      rtac (o_apply RS sym), rtac UNIV_I]) (raw_coinds ~~ unfold_defs) 1;
-
-fun mk_dtor_o_ctor_tac ctor_def unfold map_comp_id map_cong0L unfold_o_dtors
-  {context = ctxt, prems = _} =
-  unfold_thms_tac ctxt [ctor_def] THEN EVERY' [rtac ext, rtac trans, rtac o_apply,
-    rtac trans, rtac unfold, rtac trans, rtac map_comp_id, rtac trans, rtac map_cong0L,
-    EVERY' (map (fn thm =>
-      rtac ballI THEN' rtac (trans OF [thm RS fun_cong, id_apply])) unfold_o_dtors),
-    rtac sym, rtac id_apply] 1;
-
-fun mk_corec_tac m corec_defs unfold map_cong0 corec_Inls {context = ctxt, prems = _} =
-  unfold_thms_tac ctxt corec_defs THEN EVERY' [rtac trans, rtac (o_apply RS arg_cong),
-    rtac trans, rtac unfold, fo_rtac (@{thm sum.cases(2)} RS arg_cong RS trans) ctxt, rtac map_cong0,
-    REPEAT_DETERM_N m o rtac refl,
-    EVERY' (map (fn thm => rtac @{thm sum_case_expand_Inr} THEN' rtac thm) corec_Inls)] 1;
-
-fun mk_corec_unique_mor_tac corec_defs corec_Inls unfold_unique_mor {context = ctxt, prems = _} =
-  unfold_thms_tac ctxt
-    (corec_defs @ map (fn thm => thm RS @{thm sum_case_expand_Inr'}) corec_Inls) THEN
-  etac unfold_unique_mor 1;
-
-fun mk_dtor_coinduct_tac m raw_coind bis_rel rel_congs =
-  EVERY' [rtac rev_mp, rtac raw_coind, rtac ssubst, rtac bis_rel, rtac conjI,
-    CONJ_WRAP' (K (rtac @{thm ord_le_eq_trans[OF subset_UNIV UNIV_Times_UNIV[THEN sym]]}))
-    rel_congs,
-    CONJ_WRAP' (fn rel_cong => EVERY' [rtac allI, rtac allI, rtac impI,
-      REPEAT_DETERM o etac allE, rtac (rel_cong RS @{thm eq_refl} RS @{thm predicate2D}),
-      REPEAT_DETERM_N m o rtac @{thm in_rel_Id_on_UNIV[symmetric]},
-      REPEAT_DETERM_N (length rel_congs) o rtac @{thm in_rel_Collect_split_eq[symmetric]},
-      etac mp, etac CollectE, etac @{thm splitD}])
-    rel_congs,
-    rtac impI, REPEAT_DETERM o etac conjE,
-    CONJ_WRAP' (K (EVERY' [rtac impI, rtac @{thm IdD}, etac set_mp,
-      rtac CollectI, etac @{thm prod_caseI}])) rel_congs] 1;
-
-fun mk_dtor_map_coinduct_tac m ks raw_coind bis_def =
-  let
-    val n = length ks;
-  in
-    EVERY' [rtac rev_mp, rtac raw_coind, rtac ssubst, rtac bis_def, rtac conjI,
-      CONJ_WRAP' (K (rtac @{thm ord_le_eq_trans[OF subset_UNIV UNIV_Times_UNIV[THEN sym]]})) ks,
-      CONJ_WRAP' (fn i => EVERY' [select_prem_tac n (dtac asm_rl) i, REPEAT_DETERM o rtac allI,
-        rtac impI, REPEAT_DETERM o dtac @{thm meta_spec}, etac CollectE, etac @{thm meta_impE},
-        atac, etac exE, etac conjE, etac conjE, rtac bexI, rtac conjI,
-        etac @{thm fst_conv[THEN subst]}, etac @{thm snd_conv[THEN subst]},
-        rtac CollectI, REPEAT_DETERM_N m o (rtac conjI THEN' rtac subset_UNIV),
-        CONJ_WRAP' (fn i' => EVERY' [rtac subsetI, rtac CollectI, dtac (mk_conjunctN n i'),
-          REPEAT_DETERM o etac allE, etac mp, rtac @{thm ssubst_mem[OF pair_collapse]}, atac])
-        ks])
-      ks,
-      rtac impI,
-      CONJ_WRAP' (fn i => EVERY' [rtac impI, dtac (mk_conjunctN n i),
-        rtac @{thm subst[OF pair_in_Id_conv]}, etac set_mp,
-        rtac CollectI, etac (refl RSN (2, @{thm subst_Pair}))]) ks] 1
-  end;
-
-fun mk_map_tac m n cT unfold map_comp map_cong0 =
-  EVERY' [rtac ext, rtac (o_apply RS trans RS sym), rtac (o_apply RS trans RS sym),
-    rtac (unfold RS trans), rtac (Thm.permute_prems 0 1 (map_comp RS box_equals)), rtac map_cong0,
-    REPEAT_DETERM_N m o rtac (@{thm id_o} RS fun_cong),
-    REPEAT_DETERM_N n o rtac (@{thm o_id} RS fun_cong),
-    rtac (o_apply RS (Drule.instantiate' [cT] [] arg_cong) RS sym)] 1;
-
-fun mk_set_le_tac n hset_minimal set_hsets set_hset_hsetss =
-  EVERY' [rtac hset_minimal,
-    REPEAT_DETERM_N n o rtac @{thm Un_upper1},
-    REPEAT_DETERM_N n o
-      EVERY' (map3 (fn i => fn set_hset => fn set_hset_hsets =>
-        EVERY' [rtac subsetI, rtac @{thm UnI2}, rtac (mk_UnIN n i), etac @{thm UN_I},
-          etac UnE, etac set_hset, REPEAT_DETERM_N (n - 1) o etac UnE,
-          EVERY' (map (fn thm => EVERY' [etac @{thm UN_E}, etac thm, atac]) set_hset_hsets)])
-      (1 upto n) set_hsets set_hset_hsetss)] 1;
-
-fun mk_set_ge_tac n set_incl_hset set_hset_incl_hsets =
-  EVERY' [rtac @{thm Un_least}, rtac set_incl_hset,
-    REPEAT_DETERM_N (n - 1) o rtac @{thm Un_least},
-    EVERY' (map (fn thm => rtac @{thm UN_least} THEN' etac thm) set_hset_incl_hsets)] 1;
-
-fun mk_map_id0_tac maps unfold_unique unfold_dtor =
-  EVERY' [rtac (unfold_unique RS trans), EVERY' (map (rtac o mk_sym) maps),
-    rtac unfold_dtor] 1;
-
-fun mk_map_comp0_tac maps map_comp0s map_unique =
-  EVERY' [rtac map_unique,
-    EVERY' (map2 (fn map_thm => fn map_comp0 =>
-      EVERY' (map rtac
-        [@{thm o_assoc} RS trans,
-        @{thm arg_cong2[of _ _ _ _ "op o"]} OF [map_thm, refl] RS trans,
-        @{thm o_assoc} RS sym RS trans, map_thm RS arg_cong RS trans, @{thm o_assoc} RS trans,
-        @{thm arg_cong2[of _ _ _ _ "op o"]} OF [map_comp0 RS sym, refl]]))
-    maps map_comp0s)] 1;
-
-fun mk_mcong_tac ctxt m coinduct_tac map_comps dtor_maps map_cong0s set_map0ss set_hsetss
-  set_hset_hsetsss =
-  let
-    val n = length map_comps;
-    val ks = 1 upto n;
-  in
-    EVERY' ([rtac rev_mp,
-      coinduct_tac] @
-      maps (fn (((((map_comp_trans, dtor_maps_trans), map_cong0), set_map0s), set_hsets),
-        set_hset_hsetss) =>
-        [REPEAT_DETERM o eresolve_tac [exE, conjE], hyp_subst_tac ctxt, rtac exI, rtac conjI,
-         rtac conjI, rtac map_comp_trans, rtac sym, rtac dtor_maps_trans, rtac map_cong0,
-         REPEAT_DETERM_N m o (rtac o_apply_trans_sym THEN' rtac id_apply),
-         REPEAT_DETERM_N n o rtac fst_convol_fun_cong_sym,
-         rtac map_comp_trans, rtac sym, rtac dtor_maps_trans, rtac map_cong0,
-         EVERY' (maps (fn set_hset =>
-           [rtac o_apply_trans_sym, rtac (id_apply RS trans), etac CollectE,
-           REPEAT_DETERM o etac conjE, etac bspec, etac set_hset]) set_hsets),
-         REPEAT_DETERM_N n o rtac snd_convol_fun_cong_sym,
-         CONJ_WRAP' (fn (set_map0, set_hset_hsets) =>
-           EVERY' [REPEAT_DETERM o rtac allI, rtac impI, rtac @{thm image_convolD},
-             etac set_rev_mp, rtac ord_eq_le_trans, rtac set_map0,
-             rtac @{thm image_mono}, rtac subsetI, rtac CollectI, etac CollectE,
-             REPEAT_DETERM o etac conjE,
-             CONJ_WRAP' (fn set_hset_hset =>
-               EVERY' [rtac ballI, etac bspec, etac set_hset_hset, atac]) set_hset_hsets])
-           (drop m set_map0s ~~ set_hset_hsetss)])
-        (map (fn th => th RS trans) map_comps ~~ map (fn th => th RS trans) dtor_maps ~~
-          map_cong0s ~~ set_map0ss ~~ set_hsetss ~~ set_hset_hsetsss) @
-      [rtac impI,
-       CONJ_WRAP' (fn k =>
-         EVERY' [rtac impI, dtac (mk_conjunctN n k), etac mp, rtac exI, rtac conjI, etac CollectI,
-           rtac conjI, rtac refl, rtac refl]) ks]) 1
-  end
-
-fun mk_dtor_map_unique_tac unfold_unique sym_map_comps ctxt =
-  rtac unfold_unique 1 THEN
-  unfold_thms_tac ctxt (sym_map_comps @ @{thms o_assoc id_o o_id}) THEN
-  ALLGOALS (etac sym);
-
-fun mk_col_natural_tac cts rec_0s rec_Sucs dtor_maps set_map0ss
-  {context = ctxt, prems = _} =
-  let
-    val n = length dtor_maps;
-  in
-    EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
-      REPEAT_DETERM o rtac allI, SELECT_GOAL (unfold_thms_tac ctxt rec_0s),
-      CONJ_WRAP' (K (rtac @{thm image_empty})) rec_0s,
-      REPEAT_DETERM o rtac allI,
-      CONJ_WRAP' (fn (rec_Suc, (dtor_map, set_nats)) => EVERY'
-        [SELECT_GOAL (unfold_thms_tac ctxt
-          (rec_Suc :: dtor_map :: set_nats @ @{thms image_Un image_UN UN_simps(10)})),
-        rtac Un_cong, rtac refl,
-        CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 Un_cong))
-          (fn i => EVERY' [rtac @{thm UN_cong[OF refl]},
-            REPEAT_DETERM o etac allE, etac (mk_conjunctN n i)]) (n downto 1)])
-      (rec_Sucs ~~ (dtor_maps ~~ set_map0ss))] 1
-  end;
-
-fun mk_set_map0_tac hset_def col_natural =
-  EVERY' (map rtac [ext, (o_apply RS trans), (hset_def RS trans), sym,
-    (o_apply RS trans), (@{thm image_cong} OF [hset_def, refl] RS trans),
-    (@{thm image_UN} RS trans), (refl RS @{thm UN_cong}), col_natural]) 1;
-
-fun mk_col_bd_tac m j cts rec_0s rec_Sucs sbd_Card_order sbd_Cinfinite set_sbdss =
-  let
-    val n = length rec_0s;
-  in
-    EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
-      REPEAT_DETERM o rtac allI,
-      CONJ_WRAP' (fn rec_0 => EVERY' (map rtac [ordIso_ordLeq_trans,
-        @{thm card_of_ordIso_subst}, rec_0, @{thm Card_order_empty}, sbd_Card_order])) rec_0s,
-      REPEAT_DETERM o rtac allI,
-      CONJ_WRAP' (fn (rec_Suc, set_sbds) => EVERY'
-        [rtac ordIso_ordLeq_trans, rtac @{thm card_of_ordIso_subst}, rtac rec_Suc,
-        rtac (sbd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})), rtac (nth set_sbds (j - 1)),
-        REPEAT_DETERM_N (n - 1) o rtac (sbd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})),
-        EVERY' (map2 (fn i => fn set_sbd => EVERY' [rtac @{thm UNION_Cinfinite_bound},
-          rtac set_sbd, rtac ballI, REPEAT_DETERM o etac allE,
-          etac (mk_conjunctN n i), rtac sbd_Cinfinite]) (1 upto n) (drop m set_sbds))])
-      (rec_Sucs ~~ set_sbdss)] 1
-  end;
-
-fun mk_set_bd_tac sbd_Cinfinite hset_def col_bd =
-  EVERY' (map rtac [ordIso_ordLeq_trans, @{thm card_of_ordIso_subst}, hset_def,
-    @{thm UNION_Cinfinite_bound}, ordIso_ordLeq_trans, @{thm card_of_nat},
-    @{thm natLeq_ordLeq_cinfinite}, sbd_Cinfinite, ballI, col_bd, sbd_Cinfinite]) 1;
-
-fun mk_le_rel_OO_tac coinduct rel_Jrels rel_OOs =
-  EVERY' (rtac coinduct :: map2 (fn rel_Jrel => fn rel_OO =>
-    let val Jrel_imp_rel = rel_Jrel RS iffD1;
-    in
-      EVERY' [rtac (rel_OO RS sym RS @{thm eq_refl} RS @{thm predicate2D}), etac @{thm relcomppE},
-      rtac @{thm relcomppI}, etac Jrel_imp_rel, etac Jrel_imp_rel]
-    end)
-  rel_Jrels rel_OOs) 1;
-
-fun mk_wit_tac n dtor_ctors dtor_set wit coind_wits ctxt =
-  ALLGOALS (TRY o (eresolve_tac coind_wits THEN' rtac refl)) THEN
-  REPEAT_DETERM (atac 1 ORELSE
-    EVERY' [dtac set_rev_mp, rtac equalityD1, resolve_tac dtor_set,
-    K (unfold_thms_tac ctxt dtor_ctors),
-    REPEAT_DETERM_N n o etac UnE,
-    REPEAT_DETERM o
-      (TRY o REPEAT_DETERM o etac UnE THEN' TRY o etac @{thm UN_E} THEN'
-        (eresolve_tac wit ORELSE'
-        (dresolve_tac wit THEN'
-          (etac FalseE ORELSE'
-          EVERY' [hyp_subst_tac ctxt, dtac set_rev_mp, rtac equalityD1, resolve_tac dtor_set,
-            K (unfold_thms_tac ctxt dtor_ctors), REPEAT_DETERM_N n o etac UnE]))))] 1);
-
-fun mk_coind_wit_tac induct unfolds set_nats wits {context = ctxt, prems = _} =
-  rtac induct 1 THEN ALLGOALS (TRY o rtac impI THEN' TRY o hyp_subst_tac ctxt) THEN
-  unfold_thms_tac ctxt (unfolds @ set_nats @ @{thms image_id id_apply}) THEN
-  ALLGOALS (REPEAT_DETERM o etac imageE THEN' TRY o hyp_subst_tac ctxt) THEN
-  ALLGOALS (TRY o
-    FIRST' [rtac TrueI, rtac refl, etac (refl RSN (2, mp)), dresolve_tac wits THEN' etac FalseE])
-
-fun mk_dtor_rel_tac ctxt in_Jrels i in_rel map_comp0 map_cong0 dtor_map dtor_sets dtor_inject
-    dtor_ctor set_map0s dtor_set_incls dtor_set_set_inclss =
-  let
-    val m = length dtor_set_incls;
-    val n = length dtor_set_set_inclss;
-    val (passive_set_map0s, active_set_map0s) = chop m set_map0s;
-    val in_Jrel = nth in_Jrels (i - 1);
-    val if_tac =
-      EVERY' [dtac (in_Jrel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
-        rtac (in_rel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
-        EVERY' (map2 (fn set_map0 => fn set_incl =>
-          EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac set_map0,
-            rtac ord_eq_le_trans, rtac trans_fun_cong_image_id_id_apply,
-            etac (set_incl RS @{thm subset_trans})])
-        passive_set_map0s dtor_set_incls),
-        CONJ_WRAP' (fn (in_Jrel, (set_map0, dtor_set_set_incls)) =>
-          EVERY' [rtac ord_eq_le_trans, rtac set_map0, rtac @{thm image_subsetI}, rtac CollectI,
-            rtac @{thm prod_caseI}, rtac (in_Jrel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
-            CONJ_WRAP' (fn thm => etac (thm RS @{thm subset_trans}) THEN' atac) dtor_set_set_incls,
-            rtac conjI, rtac refl, rtac refl])
-        (in_Jrels ~~ (active_set_map0s ~~ dtor_set_set_inclss)),
-        CONJ_WRAP' (fn conv =>
-          EVERY' [rtac trans, rtac map_comp0, rtac trans, rtac map_cong0,
-          REPEAT_DETERM_N m o rtac @{thm fun_cong[OF o_id]},
-          REPEAT_DETERM_N n o EVERY' (map rtac [trans, o_apply, conv]),
-          rtac trans, rtac sym, rtac dtor_map, rtac (dtor_inject RS iffD2), atac])
-        @{thms fst_conv snd_conv}];
-    val only_if_tac =
-      EVERY' [dtac (in_rel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
-        rtac (in_Jrel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
-        CONJ_WRAP' (fn (dtor_set, passive_set_map0) =>
-          EVERY' [rtac ord_eq_le_trans, rtac dtor_set, rtac @{thm Un_least},
-            rtac ord_eq_le_trans, rtac box_equals, rtac passive_set_map0,
-            rtac (dtor_ctor RS sym RS arg_cong), rtac trans_fun_cong_image_id_id_apply, atac,
-            CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least}))
-              (fn (active_set_map0, in_Jrel) => EVERY' [rtac ord_eq_le_trans,
-                rtac @{thm UN_cong[OF _ refl]}, rtac @{thm box_equals[OF _ _ refl]},
-                rtac active_set_map0, rtac (dtor_ctor RS sym RS arg_cong), rtac @{thm UN_least},
-                dtac set_rev_mp, etac @{thm image_mono}, etac imageE,
-                dtac @{thm ssubst_mem[OF pair_collapse]},
-                REPEAT_DETERM o eresolve_tac (CollectE :: conjE ::
-                  @{thms prod_caseE iffD1[OF Pair_eq, elim_format]}),
-                hyp_subst_tac ctxt,
-                dtac (in_Jrel RS iffD1),
-                dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE,
-                TRY o
-                  EVERY' [dtac (Thm.permute_prems 0 1 @{thm ssubst_mem}), atac, hyp_subst_tac ctxt],
-                REPEAT_DETERM o eresolve_tac [CollectE, conjE], atac])
-            (rev (active_set_map0s ~~ in_Jrels))])
-        (dtor_sets ~~ passive_set_map0s),
-        rtac conjI,
-        REPEAT_DETERM_N 2 o EVERY'[rtac (dtor_inject RS iffD1), rtac trans, rtac dtor_map,
-          rtac box_equals, rtac map_comp0, rtac (dtor_ctor RS sym RS arg_cong), rtac trans,
-          rtac map_cong0, REPEAT_DETERM_N m o rtac @{thm fun_cong[OF o_id]},
-          EVERY' (map (fn in_Jrel => EVERY' [rtac trans, rtac o_apply, dtac set_rev_mp, atac,
-            dtac @{thm ssubst_mem[OF pair_collapse]},
-            REPEAT_DETERM o
-              eresolve_tac (CollectE :: conjE :: @{thms prod_caseE iffD1[OF Pair_eq, elim_format]}),
-            hyp_subst_tac ctxt, dtac (in_Jrel RS iffD1),
-            dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE, atac]) in_Jrels),
-          atac]]
-  in
-    EVERY' [rtac iffI, if_tac, only_if_tac] 1
-  end;
-
-fun mk_rel_coinduct_coind_tac m coinduct ks map_comp0s map_congs map_arg_congs set_map0ss 
-  dtor_unfolds dtor_maps {context = ctxt, prems = _} =
-  let val n = length ks;
-  in
-    EVERY' [DETERM o rtac coinduct,
-      EVERY' (map7 (fn i => fn map_comp0 => fn map_cong => fn map_arg_cong => fn set_map0s =>
-          fn dtor_unfold => fn dtor_map =>
-        EVERY' [REPEAT_DETERM o eresolve_tac [exE, conjE],
-          select_prem_tac (length ks) (dtac @{thm spec2}) i, dtac mp, atac,
-          REPEAT_DETERM o eresolve_tac [CollectE, conjE], hyp_subst_tac ctxt,
-          rtac exI, rtac conjI, rtac conjI,
-          rtac (map_comp0 RS trans), rtac (dtor_map RS trans RS sym),
-          rtac (dtor_unfold RS map_arg_cong RS trans), rtac (trans OF [map_comp0, map_cong]),
-          REPEAT_DETERM_N m o rtac @{thm trans[OF fun_cong[OF o_id] sym[OF fun_cong[OF id_o]]]},
-          REPEAT_DETERM_N n o (rtac @{thm sym[OF trans[OF o_apply]]} THEN' rtac @{thm fst_conv}),
-          rtac (map_comp0 RS trans), rtac (map_cong RS trans),
-          REPEAT_DETERM_N m o rtac @{thm fun_cong[OF id_o]},
-          REPEAT_DETERM_N n o (rtac @{thm trans[OF o_apply]} THEN' rtac @{thm snd_conv}),
-          etac (@{thm prod.cases} RS map_arg_cong RS trans),
-          SELECT_GOAL (unfold_thms_tac ctxt @{thms prod.cases}), 
-          CONJ_WRAP' (fn set_map0 =>
-            EVERY' [REPEAT_DETERM o resolve_tac [allI, impI],
-              dtac (set_map0 RS equalityD1 RS set_mp),
-              REPEAT_DETERM o
-                eresolve_tac (imageE :: conjE :: @{thms iffD1[OF Pair_eq, elim_format]}),
-              hyp_subst_tac ctxt, rtac exI, rtac conjI, etac Collect_splitD_set_mp, atac,
-              rtac (o_apply RS trans), rtac (@{thm surjective_pairing} RS arg_cong)])
-          (drop m set_map0s)])
-      ks map_comp0s map_congs map_arg_congs set_map0ss dtor_unfolds dtor_maps)] 1
-  end;
-
-val split_id_unfolds = @{thms prod.cases image_id id_apply};
-
-fun mk_rel_coinduct_ind_tac m ks unfolds set_map0ss j set_induct {context = ctxt, prems = _} =
-  let val n = length ks;
-  in
-    rtac set_induct 1 THEN
-    EVERY' (map3 (fn unfold => fn set_map0s => fn i =>
-      EVERY' [REPEAT_DETERM o resolve_tac [allI, impI], etac conjE,
-        select_prem_tac n (dtac @{thm spec2}) i, dtac mp, atac,
-        REPEAT_DETERM o eresolve_tac [CollectE, conjE, Collect_splitD_set_mp, set_rev_mp],
-        hyp_subst_tac ctxt,
-        SELECT_GOAL (unfold_thms_tac ctxt ([unfold, nth set_map0s (j - 1)] @ split_id_unfolds)),
-        rtac subset_refl])
-    unfolds set_map0ss ks) 1 THEN
-    EVERY' (map3 (fn unfold => fn set_map0s => fn i =>
-      EVERY' (map (fn set_map0 =>
-        EVERY' [REPEAT_DETERM o resolve_tac [allI, impI], etac conjE,
-        select_prem_tac n (dtac @{thm spec2}) i, dtac mp, atac,
-        REPEAT_DETERM o eresolve_tac [CollectE, conjE], hyp_subst_tac ctxt,
-        SELECT_GOAL (unfold_thms_tac ctxt ([unfold, set_map0] @ split_id_unfolds)),
-        etac imageE, hyp_subst_tac ctxt, REPEAT_DETERM o eresolve_tac [allE, mp],
-        rtac conjI, etac Collect_splitD_set_mp, atac, rtac (@{thm surjective_pairing} RS arg_cong)])
-      (drop m set_map0s)))
-    unfolds set_map0ss ks) 1
-  end;
-
-fun mk_rel_coinduct_tac in_rels in_Jrels helper_indss helper_coind1s helper_coind2s
-  {context = ctxt, prems = CIHs} =
-  let val n = length in_rels;
-  in
-    Method.insert_tac CIHs 1 THEN
-    unfold_thms_tac ctxt (@{thm choice_iff} :: @{thm ex_simps(6)[symmetric]} :: in_rels) THEN
-    REPEAT_DETERM (etac exE 1) THEN
-    CONJ_WRAP' (fn (in_Jrel, (helper_inds, (helper_coind1, helper_coind2))) =>
-      EVERY' [rtac @{thm predicate2I}, rtac (in_Jrel RS iffD2), rtac exI, rtac conjI,
-        if null helper_inds then rtac UNIV_I
-        else rtac CollectI THEN'
-          CONJ_WRAP' (fn helper_ind =>
-            EVERY' [rtac (helper_ind RS rev_mp), REPEAT_DETERM_N n o atac,
-              REPEAT_DETERM_N n o etac thin_rl, rtac impI,
-              REPEAT_DETERM o resolve_tac [subsetI, CollectI, @{thm iffD2[OF split_beta]}],
-              dtac bspec, atac, REPEAT_DETERM o eresolve_tac [allE, mp, conjE],
-              etac (refl RSN (2, conjI))])
-          helper_inds,
-        rtac conjI,
-        rtac (helper_coind1 RS rev_mp), REPEAT_DETERM_N n o atac, REPEAT_DETERM_N n o etac thin_rl,
-        rtac impI, etac mp, rtac exI, etac (refl RSN (2, conjI)),
-        rtac (helper_coind2 RS rev_mp), REPEAT_DETERM_N n o atac, REPEAT_DETERM_N n o etac thin_rl,
-        rtac impI, etac mp, rtac exI, etac (refl RSN (2, conjI))])
-    (in_Jrels ~~ (helper_indss ~~ (helper_coind1s ~~ helper_coind2s))) 1
-  end;
-
-fun mk_unfold_transfer_tac m rel_coinduct map_transfers unfolds {context = ctxt, prems = _} =
-  let
-    val n = length map_transfers;
-  in
-    unfold_thms_tac ctxt
-      @{thms fun_rel_def_butlast all_conj_distrib[symmetric] imp_conjR[symmetric]} THEN
-    unfold_thms_tac ctxt @{thms fun_rel_iff_geq_image2p} THEN
-    HEADGOAL (EVERY'
-      [REPEAT_DETERM o resolve_tac [allI, impI], rtac rel_coinduct,
-      EVERY' (map (fn map_transfer => EVERY'
-        [REPEAT_DETERM o resolve_tac [allI, impI], etac @{thm image2pE}, hyp_subst_tac ctxt,
-        SELECT_GOAL (unfold_thms_tac ctxt unfolds),
-        rtac (funpow (m + n + 1) (fn thm => thm RS @{thm fun_relD}) map_transfer),
-        REPEAT_DETERM_N m o rtac @{thm id_transfer},
-        REPEAT_DETERM_N n o rtac @{thm fun_rel_image2p},
-        etac @{thm predicate2D}, etac @{thm image2pI}])
-      map_transfers)])
-  end;
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_gfp_util.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,182 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_gfp_util.ML
-    Author:     Dmitriy Traytel, TU Muenchen
-    Copyright   2012
-
-Library for the codatatype construction.
-*)
-
-signature BNF_GFP_UTIL =
-sig
-  val mk_rec_simps: int -> thm -> thm list -> thm list list
-
-  val dest_listT: typ -> typ
-
-  val mk_Cons: term -> term -> term
-  val mk_Shift: term -> term -> term
-  val mk_Succ: term -> term -> term
-  val mk_Times: term * term -> term
-  val mk_append: term * term -> term
-  val mk_congruent: term -> term -> term
-  val mk_clists: term -> term
-  val mk_Id_on: term -> term
-  val mk_in_rel: term -> term
-  val mk_equiv: term -> term -> term
-  val mk_fromCard: term -> term -> term
-  val mk_list_rec: term -> term -> term
-  val mk_nat_rec: term -> term -> term
-  val mk_prefCl: term -> term
-  val mk_prefixeq: term -> term -> term
-  val mk_proj: term -> term
-  val mk_quotient: term -> term -> term
-  val mk_shift: term -> term -> term
-  val mk_size: term -> term
-  val mk_toCard: term -> term -> term
-  val mk_undefined: typ -> term
-  val mk_univ: term -> term
-
-  val mk_specN: int -> thm -> thm
-
-  val mk_InN_Field: int -> int -> thm
-  val mk_InN_inject: int -> int -> thm
-  val mk_InN_not_InM: int -> int -> thm
-end;
-
-structure BNF_GFP_Util : BNF_GFP_UTIL =
-struct
-
-open BNF_Util
-
-val mk_append = HOLogic.mk_binop @{const_name append};
-
-fun mk_equiv B R =
-  Const (@{const_name equiv}, fastype_of B --> fastype_of R --> HOLogic.boolT) $ B $ R;
-
-fun mk_Sigma (A, B) =
-  let
-    val AT = fastype_of A;
-    val BT = fastype_of B;
-    val ABT = mk_relT (HOLogic.dest_setT AT, HOLogic.dest_setT (range_type BT));
-  in Const (@{const_name Sigma}, AT --> BT --> ABT) $ A $ B end;
-
-fun mk_Id_on A =
-  let
-    val AT = fastype_of A;
-    val AAT = mk_relT (HOLogic.dest_setT AT, HOLogic.dest_setT AT);
-  in Const (@{const_name Id_on}, AT --> AAT) $ A end;
-
-fun mk_in_rel R =
-  let
-    val ((A, B), RT) = `dest_relT (fastype_of R);
-  in Const (@{const_name in_rel}, RT --> mk_pred2T A B) $ R end;
-
-fun mk_Times (A, B) =
-  let val AT = HOLogic.dest_setT (fastype_of A);
-  in mk_Sigma (A, Term.absdummy AT B) end;
-
-fun dest_listT (Type (@{type_name list}, [T])) = T
-  | dest_listT T = raise TYPE ("dest_setT: set type expected", [T], []);
-
-fun mk_Succ Kl kl =
-  let val T = fastype_of kl;
-  in
-    Const (@{const_name Succ},
-      HOLogic.mk_setT T --> T --> HOLogic.mk_setT (dest_listT T)) $ Kl $ kl
-  end;
-
-fun mk_Shift Kl k =
-  let val T = fastype_of Kl;
-  in
-    Const (@{const_name Shift}, T --> dest_listT (HOLogic.dest_setT T) --> T) $ Kl $ k
-  end;
-
-fun mk_shift lab k =
-  let val T = fastype_of lab;
-  in
-    Const (@{const_name shift}, T --> dest_listT (Term.domain_type T) --> T) $ lab $ k
-  end;
-
-fun mk_prefCl A =
-  Const (@{const_name prefCl}, fastype_of A --> HOLogic.boolT) $ A;
-
-fun mk_prefixeq xs ys =
-  Const (@{const_name prefixeq}, fastype_of xs --> fastype_of ys --> HOLogic.boolT) $ xs $ ys;
-
-fun mk_clists r =
-  let val T = fastype_of r;
-  in Const (@{const_name clists}, T --> mk_relT (`I (HOLogic.listT (fst (dest_relT T))))) $ r end;
-
-fun mk_toCard A r =
-  let
-    val AT = fastype_of A;
-    val rT = fastype_of r;
-  in
-    Const (@{const_name toCard},
-      AT --> rT --> HOLogic.dest_setT AT --> fst (dest_relT rT)) $ A $ r
-  end;
-
-fun mk_fromCard A r =
-  let
-    val AT = fastype_of A;
-    val rT = fastype_of r;
-  in
-    Const (@{const_name fromCard},
-      AT --> rT --> fst (dest_relT rT) --> HOLogic.dest_setT AT) $ A $ r
-  end;
-
-fun mk_Cons x xs =
-  let val T = fastype_of xs;
-  in Const (@{const_name Cons}, dest_listT T --> T --> T) $ x $ xs end;
-
-fun mk_size t = HOLogic.size_const (fastype_of t) $ t;
-
-fun mk_quotient A R =
-  let val T = fastype_of A;
-  in Const (@{const_name quotient}, T --> fastype_of R --> HOLogic.mk_setT T) $ A $ R end;
-
-fun mk_proj R =
-  let val ((AT, BT), T) = `dest_relT (fastype_of R);
-  in Const (@{const_name proj}, T --> AT --> HOLogic.mk_setT BT) $ R end;
-
-fun mk_univ f =
-  let val ((AT, BT), T) = `dest_funT (fastype_of f);
-  in Const (@{const_name univ}, T --> HOLogic.mk_setT AT --> BT) $ f end;
-
-fun mk_congruent R f =
-  Const (@{const_name congruent}, fastype_of R --> fastype_of f --> HOLogic.boolT) $ R $ f;
-
-fun mk_undefined T = Const (@{const_name undefined}, T);
-
-fun mk_nat_rec Zero Suc =
-  let val T = fastype_of Zero;
-  in Const (@{const_name nat_rec}, T --> fastype_of Suc --> HOLogic.natT --> T) $ Zero $ Suc end;
-
-fun mk_list_rec Nil Cons =
-  let
-    val T = fastype_of Nil;
-    val (U, consT) = `(Term.domain_type) (fastype_of Cons);
-  in
-    Const (@{const_name list_rec}, T --> consT --> HOLogic.listT U --> T) $ Nil $ Cons
-  end;
-
-fun mk_InN_not_InM 1 _ = @{thm Inl_not_Inr}
-  | mk_InN_not_InM n m =
-    if n > m then mk_InN_not_InM m n RS @{thm not_sym}
-    else mk_InN_not_InM (n - 1) (m - 1) RS @{thm not_arg_cong_Inr};
-
-fun mk_InN_Field 1 1 = @{thm TrueE[OF TrueI]}
-  | mk_InN_Field _ 1 = @{thm Inl_Field_csum}
-  | mk_InN_Field 2 2 = @{thm Inr_Field_csum}
-  | mk_InN_Field n m = mk_InN_Field (n - 1) (m - 1) RS @{thm Inr_Field_csum};
-
-fun mk_InN_inject 1 _ = @{thm TrueE[OF TrueI]}
-  | mk_InN_inject _ 1 = @{thm Inl_inject}
-  | mk_InN_inject 2 2 = @{thm Inr_inject}
-  | mk_InN_inject n m = @{thm Inr_inject} RS mk_InN_inject (n - 1) (m - 1);
-
-fun mk_specN 0 thm = thm
-  | mk_specN n thm = mk_specN (n - 1) (thm RS spec);
-
-fun mk_rec_simps n rec_thm defs = map (fn i =>
-  map (fn def => def RS rec_thm RS mk_nthI n i RS fun_cong) defs) (1 upto n);
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_lfp.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1881 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_lfp.ML
-    Author:     Dmitriy Traytel, TU Muenchen
-    Author:     Andrei Popescu, TU Muenchen
-    Copyright   2012
-
-Datatype construction.
-*)
-
-signature BNF_LFP =
-sig
-  val construct_lfp: mixfix list -> binding list -> binding list -> binding list list ->
-    binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list ->
-    local_theory -> BNF_FP_Util.fp_result * local_theory
-end;
-
-structure BNF_LFP : BNF_LFP =
-struct
-
-open BNF_Def
-open BNF_Util
-open BNF_Tactics
-open BNF_Comp
-open BNF_FP_Util
-open BNF_FP_Def_Sugar
-open BNF_LFP_Rec_Sugar
-open BNF_LFP_Util
-open BNF_LFP_Tactics
-
-(*all BNFs have the same lives*)
-fun construct_lfp mixfixes map_bs rel_bs set_bss0 bs resBs (resDs, Dss) bnfs lthy =
-  let
-    val time = time lthy;
-    val timer = time (Timer.startRealTimer ());
-
-    val live = live_of_bnf (hd bnfs);
-    val n = length bnfs; (*active*)
-    val ks = 1 upto n;
-    val m = live - n; (*passive, if 0 don't generate a new BNF*)
-
-    val note_all = Config.get lthy bnf_note_all;
-    val b_names = map Binding.name_of bs;
-    val b_name = mk_common_name b_names;
-    val b = Binding.name b_name;
-    val mk_internal_b = Binding.name #> Binding.prefix true b_name #> Binding.conceal;
-    fun mk_internal_bs name =
-      map (fn b =>
-        Binding.prefix true b_name (Binding.prefix_name (name ^ "_") b) |> Binding.conceal) bs;
-    val external_bs = map2 (Binding.prefix false) b_names bs
-      |> note_all = false ? map Binding.conceal;
-
-    (* TODO: check if m, n, etc., are sane *)
-
-    val deads = fold (union (op =)) Dss resDs;
-    val names_lthy = fold Variable.declare_typ deads lthy;
-    val passives = map fst (subtract (op = o apsnd TFree) deads resBs);
-
-    (* tvars *)
-    val (((((passiveAs, activeAs), passiveBs), activeBs), passiveCs), activeCs) =
-      names_lthy
-      |> variant_tfrees passives
-      ||>> mk_TFrees n
-      ||>> variant_tfrees passives
-      ||>> mk_TFrees n
-      ||>> variant_tfrees passives
-      ||>> mk_TFrees n
-      |> fst;
-
-    val allAs = passiveAs @ activeAs;
-    val allBs' = passiveBs @ activeBs;
-    val Ass = replicate n allAs;
-    val allBs = passiveAs @ activeBs;
-    val Bss = replicate n allBs;
-    val allCs = passiveAs @ activeCs;
-    val allCs' = passiveBs @ activeCs;
-    val Css' = replicate n allCs';
-
-    (* types *)
-    val dead_poss =
-      map (fn x => if member (op =) deads (TFree x) then SOME (TFree x) else NONE) resBs;
-    fun mk_param NONE passive = (hd passive, tl passive)
-      | mk_param (SOME a) passive = (a, passive);
-    val mk_params = fold_map mk_param dead_poss #> fst;
-
-    fun mk_FTs Ts = map2 (fn Ds => mk_T_of_bnf Ds Ts) Dss bnfs;
-    val (params, params') = `(map Term.dest_TFree) (mk_params passiveAs);
-    val FTsAs = mk_FTs allAs;
-    val FTsBs = mk_FTs allBs;
-    val FTsCs = mk_FTs allCs;
-    val ATs = map HOLogic.mk_setT passiveAs;
-    val BTs = map HOLogic.mk_setT activeAs;
-    val B'Ts = map HOLogic.mk_setT activeBs;
-    val B''Ts = map HOLogic.mk_setT activeCs;
-    val sTs = map2 (curry op -->) FTsAs activeAs;
-    val s'Ts = map2 (curry op -->) FTsBs activeBs;
-    val s''Ts = map2 (curry op -->) FTsCs activeCs;
-    val fTs = map2 (curry op -->) activeAs activeBs;
-    val inv_fTs = map2 (curry op -->) activeBs activeAs;
-    val self_fTs = map2 (curry op -->) activeAs activeAs;
-    val gTs = map2 (curry op -->) activeBs activeCs;
-    val all_gTs = map2 (curry op -->) allBs allCs';
-    val prodBsAs = map2 (curry HOLogic.mk_prodT) activeBs activeAs;
-    val prodFTs = mk_FTs (passiveAs @ prodBsAs);
-    val prod_sTs = map2 (curry op -->) prodFTs activeAs;
-
-    (* terms *)
-    val mapsAsAs = map4 mk_map_of_bnf Dss Ass Ass bnfs;
-    val mapsAsBs = map4 mk_map_of_bnf Dss Ass Bss bnfs;
-    val mapsBsAs = map4 mk_map_of_bnf Dss Bss Ass bnfs;
-    val mapsBsCs' = map4 mk_map_of_bnf Dss Bss Css' bnfs;
-    val mapsAsCs' = map4 mk_map_of_bnf Dss Ass Css' bnfs;
-    val map_fsts = map4 mk_map_of_bnf Dss (replicate n (passiveAs @ prodBsAs)) Bss bnfs;
-    val map_fsts_rev = map4 mk_map_of_bnf Dss Bss (replicate n (passiveAs @ prodBsAs)) bnfs;
-    fun mk_setss Ts = map3 mk_sets_of_bnf (map (replicate live) Dss)
-      (map (replicate live) (replicate n Ts)) bnfs;
-    val setssAs = mk_setss allAs;
-    val bd0s = map3 mk_bd_of_bnf Dss Ass bnfs;
-    val bds =
-      map3 (fn bd0 => fn Ds => fn bnf => mk_csum bd0
-        (mk_card_of (HOLogic.mk_UNIV
-          (mk_T_of_bnf Ds (replicate live (fst (dest_relT (fastype_of bd0)))) bnf))))
-      bd0s Dss bnfs;
-    val witss = map wits_of_bnf bnfs;
-
-    val (((((((((((((((((((zs, zs'), As), Bs), Bs_copy), B's), B''s), ss), prod_ss), s's), s''s),
-      fs), fs_copy), inv_fs), self_fs), gs), all_gs), (xFs, xFs')), (yFs, yFs')),
-      names_lthy) = lthy
-      |> mk_Frees' "z" activeAs
-      ||>> mk_Frees "A" ATs
-      ||>> mk_Frees "B" BTs
-      ||>> mk_Frees "B" BTs
-      ||>> mk_Frees "B'" B'Ts
-      ||>> mk_Frees "B''" B''Ts
-      ||>> mk_Frees "s" sTs
-      ||>> mk_Frees "prods" prod_sTs
-      ||>> mk_Frees "s'" s'Ts
-      ||>> mk_Frees "s''" s''Ts
-      ||>> mk_Frees "f" fTs
-      ||>> mk_Frees "f" fTs
-      ||>> mk_Frees "f" inv_fTs
-      ||>> mk_Frees "f" self_fTs
-      ||>> mk_Frees "g" gTs
-      ||>> mk_Frees "g" all_gTs
-      ||>> mk_Frees' "x" FTsAs
-      ||>> mk_Frees' "y" FTsBs;
-
-    val passive_UNIVs = map HOLogic.mk_UNIV passiveAs;
-    val active_UNIVs = map HOLogic.mk_UNIV activeAs;
-    val prod_UNIVs = map HOLogic.mk_UNIV prodBsAs;
-    val passive_ids = map HOLogic.id_const passiveAs;
-    val active_ids = map HOLogic.id_const activeAs;
-    val fsts = map fst_const prodBsAs;
-
-    (* thms *)
-    val bd0_card_orders = map bd_card_order_of_bnf bnfs;
-    val bd0_Card_orders = map bd_Card_order_of_bnf bnfs;
-    val bd0_Cinfinites = map bd_Cinfinite_of_bnf bnfs;
-    val set_bd0ss = map set_bd_of_bnf bnfs;
-
-    val bd_card_orders =
-      map (fn thm => @{thm card_order_csum} OF [thm, @{thm card_of_card_order_on}]) bd0_card_orders;
-    val bd_Card_order = @{thm Card_order_csum};
-    val bd_Card_orders = replicate n bd_Card_order;
-    val bd_Cinfinites = map (fn thm => thm RS @{thm Cinfinite_csum1}) bd0_Cinfinites;
-    val bd_Cnotzeros = map (fn thm => thm RS @{thm Cinfinite_Cnotzero}) bd_Cinfinites;
-    val bd_Cinfinite = hd bd_Cinfinites;
-    val bd_Cnotzero = hd bd_Cnotzeros;
-    val set_bdss =
-      map2 (fn set_bd0s => fn bd0_Card_order =>
-        map (fn thm => ctrans OF [thm, bd0_Card_order RS @{thm ordLeq_csum1}]) set_bd0s)
-      set_bd0ss bd0_Card_orders;
-    val in_bds = map in_bd_of_bnf bnfs;
-    val sym_map_comps = map (fn bnf => map_comp0_of_bnf bnf RS sym) bnfs;
-    val map_comps = map map_comp_of_bnf bnfs;
-    val map_cong0s = map map_cong0_of_bnf bnfs;
-    val map_id0s = map map_id0_of_bnf bnfs;
-    val map_ids = map map_id_of_bnf bnfs;
-    val set_mapss = map set_map_of_bnf bnfs;
-    val rel_mono_strongs = map rel_mono_strong_of_bnf bnfs;
-    val rel_OOs = map rel_OO_of_bnf bnfs;
-
-    val timer = time (timer "Extracted terms & thms");
-
-    (* nonemptiness check *)
-    fun new_wit X (wit: nonemptiness_witness) = subset (op =) (#I wit, (0 upto m - 1) @ map snd X);
-
-    val all = m upto m + n - 1;
-
-    fun enrich X = map_filter (fn i =>
-      (case find_first (fn (_, i') => i = i') X of
-        NONE =>
-          (case find_index (new_wit X) (nth witss (i - m)) of
-            ~1 => NONE
-          | j => SOME (j, i))
-      | SOME ji => SOME ji)) all;
-    val reachable = fixpoint (op =) enrich [];
-    val _ = (case subtract (op =) (map snd reachable) all of
-        [] => ()
-      | i :: _ => error ("Cannot define empty datatype " ^ quote (Binding.name_of (nth bs (i - m)))));
-
-    val wit_thms = flat (map2 (fn bnf => fn (j, _) => nth (wit_thmss_of_bnf bnf) j) bnfs reachable);
-
-    val timer = time (timer "Checked nonemptiness");
-
-    (* derived thms *)
-
-    (*map g1 ... gm g(m+1) ... g(m+n) (map id ... id f(m+1) ... f(m+n) x) =
-      map g1 ... gm (g(m+1) o f(m+1)) ... (g(m+n) o f(m+n)) x*)
-    fun mk_map_comp_id x mapAsBs mapBsCs mapAsCs map_comp0 =
-      let
-        val lhs = Term.list_comb (mapBsCs, all_gs) $
-          (Term.list_comb (mapAsBs, passive_ids @ fs) $ x);
-        val rhs = Term.list_comb (mapAsCs,
-          take m all_gs @ map HOLogic.mk_comp (drop m all_gs ~~ fs)) $ x;
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (x :: fs @ all_gs) (mk_Trueprop_eq (lhs, rhs)))
-          (K (mk_map_comp_id_tac map_comp0))
-        |> Thm.close_derivation
-      end;
-
-    val map_comp_id_thms = map5 mk_map_comp_id xFs mapsAsBs mapsBsCs' mapsAsCs' map_comps;
-
-    (*forall a : set(m+1) x. f(m+1) a = a; ...; forall a : set(m+n) x. f(m+n) a = a ==>
-      map id ... id f(m+1) ... f(m+n) x = x*)
-    fun mk_map_cong0L x mapAsAs sets map_cong0 map_id =
-      let
-        fun mk_prem set f z z' = HOLogic.mk_Trueprop
-          (mk_Ball (set $ x) (Term.absfree z' (HOLogic.mk_eq (f $ z, z))));
-        val prems = map4 mk_prem (drop m sets) self_fs zs zs';
-        val goal = mk_Trueprop_eq (Term.list_comb (mapAsAs, passive_ids @ self_fs) $ x, x);
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (x :: self_fs) (Logic.list_implies (prems, goal)))
-          (K (mk_map_cong0L_tac m map_cong0 map_id))
-        |> Thm.close_derivation
-      end;
-
-    val map_cong0L_thms = map5 mk_map_cong0L xFs mapsAsAs setssAs map_cong0s map_ids;
-    val in_mono'_thms = map (fn bnf => in_mono_of_bnf bnf OF (replicate m subset_refl)) bnfs;
-    val in_cong'_thms = map (fn bnf => in_cong_of_bnf bnf OF (replicate m refl)) bnfs;
-
-    val timer = time (timer "Derived simple theorems");
-
-    (* algebra *)
-
-    val alg_bind = mk_internal_b algN;
-    val alg_name = Binding.name_of alg_bind;
-    val alg_def_bind = (Thm.def_binding alg_bind, []);
-
-    (*forall i = 1 ... n: (\<forall>x \<in> Fi_in A1 .. Am B1 ... Bn. si x \<in> Bi)*)
-    val alg_spec =
-      let
-        val algT = Library.foldr (op -->) (ATs @ BTs @ sTs, HOLogic.boolT);
-
-        val ins = map3 mk_in (replicate n (As @ Bs)) setssAs FTsAs;
-        fun mk_alg_conjunct B s X x x' =
-          mk_Ball X (Term.absfree x' (HOLogic.mk_mem (s $ x, B)));
-
-        val lhs = Term.list_comb (Free (alg_name, algT), As @ Bs @ ss);
-        val rhs = Library.foldr1 HOLogic.mk_conj (map5 mk_alg_conjunct Bs ss ins xFs xFs')
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((alg_free, (_, alg_def_free)), (lthy, lthy_old)) =
-        lthy
-        |> Specification.definition (SOME (alg_bind, NONE, NoSyn), (alg_def_bind, alg_spec))
-        ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-    val alg = fst (Term.dest_Const (Morphism.term phi alg_free));
-    val alg_def = Morphism.thm phi alg_def_free;
-
-    fun mk_alg As Bs ss =
-      let
-        val args = As @ Bs @ ss;
-        val Ts = map fastype_of args;
-        val algT = Library.foldr (op -->) (Ts, HOLogic.boolT);
-      in
-        Term.list_comb (Const (alg, algT), args)
-      end;
-
-    val alg_set_thms =
-      let
-        val alg_prem = HOLogic.mk_Trueprop (mk_alg As Bs ss);
-        fun mk_prem x set B = HOLogic.mk_Trueprop (mk_leq (set $ x) B);
-        fun mk_concl s x B = HOLogic.mk_Trueprop (HOLogic.mk_mem (s $ x, B));
-        val premss = map2 ((fn x => fn sets =>  map2 (mk_prem x) sets (As @ Bs))) xFs setssAs;
-        val concls = map3 mk_concl ss xFs Bs;
-        val goals = map3 (fn x => fn prems => fn concl =>
-          fold_rev Logic.all (x :: As @ Bs @ ss)
-            (Logic.list_implies (alg_prem :: prems, concl))) xFs premss concls;
-      in
-        map (fn goal =>
-          Goal.prove_sorry lthy [] [] goal (K (mk_alg_set_tac alg_def)) |> Thm.close_derivation)
-        goals
-      end;
-
-    fun mk_talg ATs BTs = mk_alg (map HOLogic.mk_UNIV ATs) (map HOLogic.mk_UNIV BTs);
-
-    val talg_thm =
-      let
-        val goal = fold_rev Logic.all ss
-          (HOLogic.mk_Trueprop (mk_talg passiveAs activeAs ss))
-      in
-        Goal.prove_sorry lthy [] [] goal
-          (K (stac alg_def 1 THEN CONJ_WRAP (K (EVERY' [rtac ballI, rtac UNIV_I] 1)) ss))
-        |> Thm.close_derivation
-      end;
-
-    val timer = time (timer "Algebra definition & thms");
-
-    val alg_not_empty_thms =
-      let
-        val alg_prem =
-          HOLogic.mk_Trueprop (mk_alg passive_UNIVs Bs ss);
-        val concls = map (HOLogic.mk_Trueprop o mk_not_empty) Bs;
-        val goals =
-          map (fn concl =>
-            fold_rev Logic.all (Bs @ ss) (Logic.mk_implies (alg_prem, concl))) concls;
-      in
-        map2 (fn goal => fn alg_set =>
-          Goal.prove_sorry lthy [] []
-            goal (K (mk_alg_not_empty_tac lthy alg_set alg_set_thms wit_thms))
-          |> Thm.close_derivation)
-        goals alg_set_thms
-      end;
-
-    val timer = time (timer "Proved nonemptiness");
-
-    (* morphism *)
-
-    val mor_bind = mk_internal_b morN;
-    val mor_name = Binding.name_of mor_bind;
-    val mor_def_bind = (Thm.def_binding mor_bind, []);
-
-    (*fbetw) forall i = 1 ... n: (\<forall>x \<in> Bi. f x \<in> B'i)*)
-    (*mor) forall i = 1 ... n: (\<forall>x \<in> Fi_in UNIV ... UNIV B1 ... Bn.
-       f (s1 x) = s1' (Fi_map id ... id f1 ... fn x))*)
-    val mor_spec =
-      let
-        val morT = Library.foldr (op -->) (BTs @ sTs @ B'Ts @ s'Ts @ fTs, HOLogic.boolT);
-
-        fun mk_fbetw f B1 B2 z z' =
-          mk_Ball B1 (Term.absfree z' (HOLogic.mk_mem (f $ z, B2)));
-        fun mk_mor sets mapAsBs f s s' T x x' =
-          mk_Ball (mk_in (passive_UNIVs @ Bs) sets T)
-            (Term.absfree x' (HOLogic.mk_eq (f $ (s $ x), s' $
-              (Term.list_comb (mapAsBs, passive_ids @ fs) $ x))));
-        val lhs = Term.list_comb (Free (mor_name, morT), Bs @ ss @ B's @ s's @ fs);
-        val rhs = HOLogic.mk_conj
-          (Library.foldr1 HOLogic.mk_conj (map5 mk_fbetw fs Bs B's zs zs'),
-          Library.foldr1 HOLogic.mk_conj
-            (map8 mk_mor setssAs mapsAsBs fs ss s's FTsAs xFs xFs'))
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((mor_free, (_, mor_def_free)), (lthy, lthy_old)) =
-        lthy
-        |> Specification.definition (SOME (mor_bind, NONE, NoSyn), (mor_def_bind, mor_spec))
-        ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-    val mor = fst (Term.dest_Const (Morphism.term phi mor_free));
-    val mor_def = Morphism.thm phi mor_def_free;
-
-    fun mk_mor Bs1 ss1 Bs2 ss2 fs =
-      let
-        val args = Bs1 @ ss1 @ Bs2 @ ss2 @ fs;
-        val Ts = map fastype_of (Bs1 @ ss1 @ Bs2 @ ss2 @ fs);
-        val morT = Library.foldr (op -->) (Ts, HOLogic.boolT);
-      in
-        Term.list_comb (Const (mor, morT), args)
-      end;
-
-    val (mor_image_thms, morE_thms) =
-      let
-        val prem = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs);
-        fun mk_image_goal f B1 B2 = fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs)
-          (Logic.mk_implies (prem, HOLogic.mk_Trueprop (mk_leq (mk_image f $ B1) B2)));
-        val image_goals = map3 mk_image_goal fs Bs B's;
-        fun mk_elim_prem sets x T = HOLogic.mk_Trueprop
-          (HOLogic.mk_mem (x, mk_in (passive_UNIVs @ Bs) sets T));
-        fun mk_elim_goal sets mapAsBs f s s' x T =
-          fold_rev Logic.all (x :: Bs @ ss @ B's @ s's @ fs)
-            (Logic.list_implies ([prem, mk_elim_prem sets x T],
-              mk_Trueprop_eq (f $ (s $ x), s' $ Term.list_comb (mapAsBs, passive_ids @ fs @ [x]))));
-        val elim_goals = map7 mk_elim_goal setssAs mapsAsBs fs ss s's xFs FTsAs;
-        fun prove goal =
-          Goal.prove_sorry lthy [] [] goal (K (mk_mor_elim_tac mor_def)) |> Thm.close_derivation;
-      in
-        (map prove image_goals, map prove elim_goals)
-      end;
-
-    val mor_incl_thm =
-      let
-        val prems = map2 (HOLogic.mk_Trueprop oo mk_leq) Bs Bs_copy;
-        val concl = HOLogic.mk_Trueprop (mk_mor Bs ss Bs_copy ss active_ids);
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (Bs @ ss @ Bs_copy) (Logic.list_implies (prems, concl)))
-          (K (mk_mor_incl_tac mor_def map_ids))
-        |> Thm.close_derivation
-      end;
-
-    val mor_comp_thm =
-      let
-        val prems =
-          [HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs),
-           HOLogic.mk_Trueprop (mk_mor B's s's B''s s''s gs)];
-        val concl =
-          HOLogic.mk_Trueprop (mk_mor Bs ss B''s s''s (map2 (curry HOLogic.mk_comp) gs fs));
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (Bs @ ss @ B's @ s's @ B''s @ s''s @ fs @ gs)
-             (Logic.list_implies (prems, concl)))
-          (K (mk_mor_comp_tac mor_def set_mapss map_comp_id_thms))
-        |> Thm.close_derivation
-      end;
-
-    val mor_inv_thm =
-      let
-        fun mk_inv_prem f inv_f B B' = HOLogic.mk_conj (mk_leq (mk_image inv_f $ B') B,
-          HOLogic.mk_conj (mk_inver inv_f f B, mk_inver f inv_f B'));
-        val prems = map HOLogic.mk_Trueprop
-          ([mk_mor Bs ss B's s's fs,
-          mk_alg passive_UNIVs Bs ss,
-          mk_alg passive_UNIVs B's s's] @
-          map4 mk_inv_prem fs inv_fs Bs B's);
-        val concl = HOLogic.mk_Trueprop (mk_mor B's s's Bs ss inv_fs);
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs @ inv_fs)
-            (Logic.list_implies (prems, concl)))
-          (K (mk_mor_inv_tac alg_def mor_def set_mapss morE_thms map_comp_id_thms map_cong0L_thms))
-        |> Thm.close_derivation
-      end;
-
-    val mor_cong_thm =
-      let
-        val prems = map HOLogic.mk_Trueprop
-         (map2 (curry HOLogic.mk_eq) fs_copy fs @ [mk_mor Bs ss B's s's fs])
-        val concl = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs_copy);
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs @ fs_copy)
-             (Logic.list_implies (prems, concl)))
-          (K ((hyp_subst_tac lthy THEN' atac) 1))
-        |> Thm.close_derivation
-      end;
-
-    val mor_str_thm =
-      let
-        val maps = map2 (fn Ds => fn bnf => Term.list_comb
-          (mk_map_of_bnf Ds (passiveAs @ FTsAs) allAs bnf, passive_ids @ ss)) Dss bnfs;
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all ss (HOLogic.mk_Trueprop
-            (mk_mor (map HOLogic.mk_UNIV FTsAs) maps active_UNIVs ss ss)))
-          (K (mk_mor_str_tac ks mor_def))
-        |> Thm.close_derivation
-      end;
-
-    val mor_convol_thm =
-      let
-        val maps = map3 (fn s => fn prod_s => fn mapx =>
-          mk_convol (HOLogic.mk_comp (s, Term.list_comb (mapx, passive_ids @ fsts)), prod_s))
-          s's prod_ss map_fsts;
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (s's @ prod_ss) (HOLogic.mk_Trueprop
-            (mk_mor prod_UNIVs maps (map HOLogic.mk_UNIV activeBs) s's fsts)))
-          (K (mk_mor_convol_tac ks mor_def))
-        |> Thm.close_derivation
-      end;
-
-    val mor_UNIV_thm =
-      let
-        fun mk_conjunct mapAsBs f s s' = HOLogic.mk_eq
-            (HOLogic.mk_comp (f, s),
-            HOLogic.mk_comp (s', Term.list_comb (mapAsBs, passive_ids @ fs)));
-        val lhs = mk_mor active_UNIVs ss (map HOLogic.mk_UNIV activeBs) s's fs;
-        val rhs = Library.foldr1 HOLogic.mk_conj (map4 mk_conjunct mapsAsBs fs ss s's);
-      in
-        Goal.prove_sorry lthy [] [] (fold_rev Logic.all (ss @ s's @ fs) (mk_Trueprop_eq (lhs, rhs)))
-          (K (mk_mor_UNIV_tac m morE_thms mor_def))
-        |> Thm.close_derivation
-      end;
-
-    val timer = time (timer "Morphism definition & thms");
-
-    (* isomorphism *)
-
-    (*mor Bs1 ss1 Bs2 ss2 fs \<and> (\<exists>gs. mor Bs2 ss2 Bs1 ss1 fs \<and>
-       forall i = 1 ... n. (inver gs[i] fs[i] Bs1[i] \<and> inver fs[i] gs[i] Bs2[i]))*)
-    fun mk_iso Bs1 ss1 Bs2 ss2 fs gs =
-      let
-        val ex_inv_mor = list_exists_free gs
-          (HOLogic.mk_conj (mk_mor Bs2 ss2 Bs1 ss1 gs,
-            Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_conj)
-              (map3 mk_inver gs fs Bs1) (map3 mk_inver fs gs Bs2))));
-      in
-        HOLogic.mk_conj (mk_mor Bs1 ss1 Bs2 ss2 fs, ex_inv_mor)
-      end;
-
-    val iso_alt_thm =
-      let
-        val prems = map HOLogic.mk_Trueprop
-         [mk_alg passive_UNIVs Bs ss,
-         mk_alg passive_UNIVs B's s's]
-        val concl = mk_Trueprop_eq (mk_iso Bs ss B's s's fs inv_fs,
-          HOLogic.mk_conj (mk_mor Bs ss B's s's fs,
-            Library.foldr1 HOLogic.mk_conj (map3 mk_bij_betw fs Bs B's)));
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs) (Logic.list_implies (prems, concl)))
-          (K (mk_iso_alt_tac mor_image_thms mor_inv_thm))
-        |> Thm.close_derivation
-      end;
-
-    val timer = time (timer "Isomorphism definition & thms");
-
-    (* algebra copies *)
-
-    val (copy_alg_thm, ex_copy_alg_thm) =
-      let
-        val prems = map HOLogic.mk_Trueprop
-         (mk_alg passive_UNIVs Bs ss :: map3 mk_bij_betw inv_fs B's Bs);
-        val inver_prems = map HOLogic.mk_Trueprop
-          (map3 mk_inver inv_fs fs Bs @ map3 mk_inver fs inv_fs B's);
-        val all_prems = prems @ inver_prems;
-        fun mk_s f s mapT y y' = Term.absfree y' (f $ (s $
-          (Term.list_comb (mapT, passive_ids @ inv_fs) $ y)));
-
-        val alg = HOLogic.mk_Trueprop
-          (mk_alg passive_UNIVs B's (map5 mk_s fs ss mapsBsAs yFs yFs'));
-        val copy_str_thm = Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (Bs @ ss @ B's @ inv_fs @ fs)
-            (Logic.list_implies (all_prems, alg)))
-          (K (mk_copy_str_tac set_mapss alg_def alg_set_thms))
-          |> Thm.close_derivation;
-
-        val iso = HOLogic.mk_Trueprop
-          (mk_iso B's (map5 mk_s fs ss mapsBsAs yFs yFs') Bs ss inv_fs fs_copy);
-        val copy_alg_thm = Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (Bs @ ss @ B's @ inv_fs @ fs)
-            (Logic.list_implies (all_prems, iso)))
-          (K (mk_copy_alg_tac set_mapss alg_set_thms mor_def iso_alt_thm copy_str_thm))
-          |> Thm.close_derivation;
-
-        val ex = HOLogic.mk_Trueprop
-          (list_exists_free s's
-            (HOLogic.mk_conj (mk_alg passive_UNIVs B's s's,
-              mk_iso B's s's Bs ss inv_fs fs_copy)));
-        val ex_copy_alg_thm = Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (Bs @ ss @ B's @ inv_fs @ fs)
-             (Logic.list_implies (prems, ex)))
-          (K (mk_ex_copy_alg_tac n copy_str_thm copy_alg_thm))
-          |> Thm.close_derivation;
-      in
-        (copy_alg_thm, ex_copy_alg_thm)
-      end;
-
-    val timer = time (timer "Copy thms");
-
-
-    (* bounds *)
-
-    val sum_Card_order = if n = 1 then bd_Card_order else @{thm Card_order_csum};
-    val sum_Cnotzero = if n = 1 then bd_Cnotzero else bd_Cnotzero RS @{thm csum_Cnotzero1};
-    val sum_Cinfinite = if n = 1 then bd_Cinfinite else bd_Cinfinite RS @{thm Cinfinite_csum1};
-    fun mk_set_bd_sums i bd_Card_order bds =
-      if n = 1 then bds
-      else map (fn thm => bd_Card_order RS mk_ordLeq_csum n i thm) bds;
-    val set_bd_sumss = map3 mk_set_bd_sums ks bd_Card_orders set_bdss;
-
-    fun mk_in_bd_sum i Co Cnz bd =
-      if n = 1 then bd
-      else Cnz RS ((Co RS mk_ordLeq_csum n i (Co RS @{thm ordLeq_refl})) RS
-        (bd RS @{thm ordLeq_transitive[OF _ cexp_mono2_Cnotzero[OF _ Card_order_csum]]}));
-    val in_bd_sums = map4 mk_in_bd_sum ks bd_Card_orders bd_Cnotzeros in_bds;
-
-    val sum_bd = Library.foldr1 (uncurry mk_csum) bds;
-    val suc_bd = mk_cardSuc sum_bd;
-    val field_suc_bd = mk_Field suc_bd;
-    val suc_bdT = fst (dest_relT (fastype_of suc_bd));
-    fun mk_Asuc_bd [] = mk_cexp ctwo suc_bd
-      | mk_Asuc_bd As =
-        mk_cexp (mk_csum (Library.foldr1 (uncurry mk_csum) (map mk_card_of As)) ctwo) suc_bd;
-
-    val suc_bd_Card_order = if n = 1 then bd_Card_order RS @{thm cardSuc_Card_order}
-      else @{thm cardSuc_Card_order[OF Card_order_csum]};
-    val suc_bd_Cinfinite = if n = 1 then bd_Cinfinite RS @{thm Cinfinite_cardSuc}
-      else bd_Cinfinite RS @{thm Cinfinite_cardSuc[OF Cinfinite_csum1]};
-    val suc_bd_Cnotzero = suc_bd_Cinfinite RS @{thm Cinfinite_Cnotzero};
-    val suc_bd_worel = suc_bd_Card_order RS @{thm Card_order_wo_rel}
-    val basis_Asuc = if m = 0 then @{thm ordLeq_refl[OF Card_order_ctwo]}
-        else @{thm ordLeq_csum2[OF Card_order_ctwo]};
-    val Asuc_bd_Cinfinite = suc_bd_Cinfinite RS (basis_Asuc RS @{thm Cinfinite_cexp});
-
-    val suc_bd_Asuc_bd = @{thm ordLess_ordLeq_trans[OF ordLess_ctwo_cexp cexp_mono1]} OF
-      [suc_bd_Card_order, basis_Asuc, suc_bd_Card_order];
-
-    val Asuc_bdT = fst (dest_relT (fastype_of (mk_Asuc_bd As)));
-    val II_BTs = replicate n (HOLogic.mk_setT Asuc_bdT);
-    val II_sTs = map2 (fn Ds => fn bnf =>
-      mk_T_of_bnf Ds (passiveAs @ replicate n Asuc_bdT) bnf --> Asuc_bdT) Dss bnfs;
-
-    val (((((((idxs, Asi_name), (idx, idx')), (jdx, jdx')), II_Bs), II_ss), Asuc_fs),
-      names_lthy) = names_lthy
-      |> mk_Frees "i" (replicate n suc_bdT)
-      ||>> (fn ctxt => apfst the_single (mk_fresh_names ctxt 1 "Asi"))
-      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "i") suc_bdT
-      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "j") suc_bdT
-      ||>> mk_Frees "IIB" II_BTs
-      ||>> mk_Frees "IIs" II_sTs
-      ||>> mk_Frees "f" (map (fn T => Asuc_bdT --> T) activeAs);
-
-    val suc_bd_limit_thm =
-      let
-        val prem = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
-          (map (fn idx => HOLogic.mk_mem (idx, field_suc_bd)) idxs));
-        fun mk_conjunct idx = HOLogic.mk_conj (mk_not_eq idx jdx,
-          HOLogic.mk_mem (HOLogic.mk_prod (idx, jdx), suc_bd));
-        val concl = HOLogic.mk_Trueprop (mk_Bex field_suc_bd
-          (Term.absfree jdx' (Library.foldr1 HOLogic.mk_conj (map mk_conjunct idxs))));
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all idxs (Logic.list_implies ([prem], concl)))
-          (K (mk_bd_limit_tac n suc_bd_Cinfinite))
-        |> Thm.close_derivation
-      end;
-
-    val timer = time (timer "Bounds");
-
-
-    (* minimal algebra *)
-
-    fun mk_minG Asi i k = mk_UNION (mk_underS suc_bd $ i)
-      (Term.absfree jdx' (mk_nthN n (Asi $ jdx) k));
-
-    fun mk_minH_component As Asi i sets Ts s k =
-      HOLogic.mk_binop @{const_name "sup"}
-      (mk_minG Asi i k, mk_image s $ mk_in (As @ map (mk_minG Asi i) ks) sets Ts);
-
-    fun mk_min_algs As ss =
-      let
-        val BTs = map (range_type o fastype_of) ss;
-        val Ts = map (HOLogic.dest_setT o fastype_of) As @ BTs;
-        val (Asi, Asi') = `Free (Asi_name, suc_bdT -->
-          Library.foldr1 HOLogic.mk_prodT (map HOLogic.mk_setT BTs));
-      in
-         mk_worec suc_bd (Term.absfree Asi' (Term.absfree idx' (HOLogic.mk_tuple
-           (map4 (mk_minH_component As Asi idx) (mk_setss Ts) (mk_FTs Ts) ss ks))))
-      end;
-
-    val (min_algs_thms, min_algs_mono_thms, card_of_min_algs_thm, least_min_algs_thm) =
-      let
-        val i_field = HOLogic.mk_mem (idx, field_suc_bd);
-        val min_algs = mk_min_algs As ss;
-        val min_algss = map (fn k => mk_nthN n (min_algs $ idx) k) ks;
-
-        val concl = HOLogic.mk_Trueprop
-          (HOLogic.mk_eq (min_algs $ idx, HOLogic.mk_tuple
-            (map4 (mk_minH_component As min_algs idx) setssAs FTsAs ss ks)));
-        val goal = fold_rev Logic.all (idx :: As @ ss)
-          (Logic.mk_implies (HOLogic.mk_Trueprop i_field, concl));
-
-        val min_algs_thm = Goal.prove_sorry lthy [] [] goal
-          (K (mk_min_algs_tac suc_bd_worel in_cong'_thms))
-          |> Thm.close_derivation;
-
-        val min_algs_thms = map (fn k => min_algs_thm RS mk_nthI n k) ks;
-
-        fun mk_mono_goal min_alg =
-          fold_rev Logic.all (As @ ss) (HOLogic.mk_Trueprop (mk_relChain suc_bd
-            (Term.absfree idx' min_alg)));
-
-        val monos =
-          map2 (fn goal => fn min_algs =>
-            Goal.prove_sorry lthy [] [] goal (K (mk_min_algs_mono_tac lthy min_algs))
-            |> Thm.close_derivation)
-          (map mk_mono_goal min_algss) min_algs_thms;
-
-        val Asuc_bd = mk_Asuc_bd As;
-
-        fun mk_card_conjunct min_alg = mk_ordLeq (mk_card_of min_alg) Asuc_bd;
-        val card_conjunction = Library.foldr1 HOLogic.mk_conj (map mk_card_conjunct min_algss);
-        val card_cT = certifyT lthy suc_bdT;
-        val card_ct = certify lthy (Term.absfree idx' card_conjunction);
-
-        val card_of = singleton (Proof_Context.export names_lthy lthy)
-          (Goal.prove_sorry lthy [] []
-            (HOLogic.mk_Trueprop (HOLogic.mk_imp (i_field, card_conjunction)))
-            (K (mk_min_algs_card_of_tac card_cT card_ct
-              m suc_bd_worel min_algs_thms in_bd_sums
-              sum_Card_order sum_Cnotzero suc_bd_Card_order suc_bd_Cinfinite suc_bd_Cnotzero
-              suc_bd_Asuc_bd Asuc_bd_Cinfinite)))
-          |> Thm.close_derivation;
-
-        val least_prem = HOLogic.mk_Trueprop (mk_alg As Bs ss);
-        val least_conjunction = Library.foldr1 HOLogic.mk_conj (map2 mk_leq min_algss Bs);
-        val least_cT = certifyT lthy suc_bdT;
-        val least_ct = certify lthy (Term.absfree idx' least_conjunction);
-
-        val least = singleton (Proof_Context.export names_lthy lthy)
-          (Goal.prove_sorry lthy [] []
-            (Logic.mk_implies (least_prem,
-              HOLogic.mk_Trueprop (HOLogic.mk_imp (i_field, least_conjunction))))
-            (K (mk_min_algs_least_tac least_cT least_ct
-              suc_bd_worel min_algs_thms alg_set_thms)))
-          |> Thm.close_derivation;
-      in
-        (min_algs_thms, monos, card_of, least)
-      end;
-
-    val timer = time (timer "min_algs definition & thms");
-
-    val min_alg_binds = mk_internal_bs min_algN;
-    fun min_alg_bind i = nth min_alg_binds (i - 1);
-    fun min_alg_name i = Binding.name_of (min_alg_bind i);
-    val min_alg_def_bind = rpair [] o Thm.def_binding o min_alg_bind;
-
-    fun min_alg_spec i =
-      let
-        val min_algT =
-          Library.foldr (op -->) (ATs @ sTs, HOLogic.mk_setT (nth activeAs (i - 1)));
-
-        val lhs = Term.list_comb (Free (min_alg_name i, min_algT), As @ ss);
-        val rhs = mk_UNION (field_suc_bd)
-          (Term.absfree idx' (mk_nthN n (mk_min_algs As ss $ idx) i));
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((min_alg_frees, (_, min_alg_def_frees)), (lthy, lthy_old)) =
-        lthy
-        |> fold_map (fn i => Specification.definition
-          (SOME (min_alg_bind i, NONE, NoSyn), (min_alg_def_bind i, min_alg_spec i))) ks
-        |>> apsnd split_list o split_list
-        ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-    val min_algs = map (fst o Term.dest_Const o Morphism.term phi) min_alg_frees;
-    val min_alg_defs = map (Morphism.thm phi) min_alg_def_frees;
-
-    fun mk_min_alg As ss i =
-      let
-        val T = HOLogic.mk_setT (range_type (fastype_of (nth ss (i - 1))))
-        val args = As @ ss;
-        val Ts = map fastype_of args;
-        val min_algT = Library.foldr (op -->) (Ts, T);
-      in
-        Term.list_comb (Const (nth min_algs (i - 1), min_algT), args)
-      end;
-
-    val (alg_min_alg_thm, card_of_min_alg_thms, least_min_alg_thms, mor_incl_min_alg_thm) =
-      let
-        val min_algs = map (mk_min_alg As ss) ks;
-
-        val goal = fold_rev Logic.all (As @ ss) (HOLogic.mk_Trueprop (mk_alg As min_algs ss));
-        val alg_min_alg = Goal.prove_sorry lthy [] [] goal
-          (K (mk_alg_min_alg_tac m alg_def min_alg_defs suc_bd_limit_thm sum_Cinfinite
-            set_bd_sumss min_algs_thms min_algs_mono_thms))
-          |> Thm.close_derivation;
-
-        val Asuc_bd = mk_Asuc_bd As;
-        fun mk_card_of_thm min_alg def = Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (As @ ss)
-            (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of min_alg) Asuc_bd)))
-          (K (mk_card_of_min_alg_tac def card_of_min_algs_thm
-            suc_bd_Card_order suc_bd_Asuc_bd Asuc_bd_Cinfinite))
-          |> Thm.close_derivation;
-
-        val least_prem = HOLogic.mk_Trueprop (mk_alg As Bs ss);
-        fun mk_least_thm min_alg B def = Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (As @ Bs @ ss)
-            (Logic.mk_implies (least_prem, HOLogic.mk_Trueprop (mk_leq min_alg B))))
-          (K (mk_least_min_alg_tac def least_min_algs_thm))
-          |> Thm.close_derivation;
-
-        val leasts = map3 mk_least_thm min_algs Bs min_alg_defs;
-
-        val incl_prem = HOLogic.mk_Trueprop (mk_alg passive_UNIVs Bs ss);
-        val incl_min_algs = map (mk_min_alg passive_UNIVs ss) ks;
-        val incl = Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (Bs @ ss)
-            (Logic.mk_implies (incl_prem,
-              HOLogic.mk_Trueprop (mk_mor incl_min_algs ss Bs ss active_ids))))
-          (K (EVERY' (rtac mor_incl_thm :: map etac leasts) 1))
-          |> Thm.close_derivation;
-      in
-        (alg_min_alg, map2 mk_card_of_thm min_algs min_alg_defs, leasts, incl)
-      end;
-
-    val timer = time (timer "Minimal algebra definition & thms");
-
-    val II_repT = HOLogic.mk_prodT (HOLogic.mk_tupleT II_BTs, HOLogic.mk_tupleT II_sTs);
-    val IIT_bind = mk_internal_b IITN;
-
-    val ((IIT_name, (IIT_glob_info, IIT_loc_info)), lthy) =
-      typedef (IIT_bind, params, NoSyn)
-        (HOLogic.mk_UNIV II_repT) NONE (EVERY' [rtac exI, rtac UNIV_I] 1) lthy;
-
-    val IIT = Type (IIT_name, params');
-    val Abs_IIT = Const (#Abs_name IIT_glob_info, II_repT --> IIT);
-    val Rep_IIT = Const (#Rep_name IIT_glob_info, IIT --> II_repT);
-    val Abs_IIT_inverse_thm = UNIV_I RS #Abs_inverse IIT_loc_info;
-
-    val initT = IIT --> Asuc_bdT;
-    val active_initTs = replicate n initT;
-    val init_FTs = map2 (fn Ds => mk_T_of_bnf Ds (passiveAs @ active_initTs)) Dss bnfs;
-    val init_fTs = map (fn T => initT --> T) activeAs;
-
-    val (((((((iidx, iidx'), init_xs), (init_xFs, init_xFs')),
-      init_fs), init_fs_copy), init_phis), names_lthy) = names_lthy
-      |> yield_singleton (apfst (op ~~) oo mk_Frees' "i") IIT
-      ||>> mk_Frees "ix" active_initTs
-      ||>> mk_Frees' "x" init_FTs
-      ||>> mk_Frees "f" init_fTs
-      ||>> mk_Frees "f" init_fTs
-      ||>> mk_Frees "P" (replicate n (mk_pred1T initT));
-
-    val II = HOLogic.mk_Collect (fst iidx', IIT, list_exists_free (II_Bs @ II_ss)
-      (HOLogic.mk_conj (HOLogic.mk_eq (iidx,
-        Abs_IIT $ (HOLogic.mk_prod (HOLogic.mk_tuple II_Bs, HOLogic.mk_tuple II_ss))),
-        mk_alg passive_UNIVs II_Bs II_ss)));
-
-    val select_Bs = map (mk_nthN n (HOLogic.mk_fst (Rep_IIT $ iidx))) ks;
-    val select_ss = map (mk_nthN n (HOLogic.mk_snd (Rep_IIT $ iidx))) ks;
-
-    val str_init_binds = mk_internal_bs str_initN;
-    fun str_init_bind i = nth str_init_binds (i - 1);
-    val str_init_name = Binding.name_of o str_init_bind;
-    val str_init_def_bind = rpair [] o Thm.def_binding o str_init_bind;
-
-    fun str_init_spec i =
-      let
-        val T = nth init_FTs (i - 1);
-        val init_xF = nth init_xFs (i - 1)
-        val select_s = nth select_ss (i - 1);
-        val map = mk_map_of_bnf (nth Dss (i - 1))
-          (passiveAs @ active_initTs) (passiveAs @ replicate n Asuc_bdT)
-          (nth bnfs (i - 1));
-        val map_args = passive_ids @ replicate n (mk_rapp iidx Asuc_bdT);
-        val str_initT = T --> IIT --> Asuc_bdT;
-
-        val lhs = Term.list_comb (Free (str_init_name i, str_initT), [init_xF, iidx]);
-        val rhs = select_s $ (Term.list_comb (map, map_args) $ init_xF);
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((str_init_frees, (_, str_init_def_frees)), (lthy, lthy_old)) =
-      lthy
-      |> fold_map (fn i => Specification.definition
-        (SOME (str_init_bind i, NONE, NoSyn), (str_init_def_bind i, str_init_spec i))) ks
-      |>> apsnd split_list o split_list
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-    val str_inits =
-      map (Term.subst_atomic_types (map (`(Morphism.typ phi)) params') o Morphism.term phi)
-        str_init_frees;
-
-    val str_init_defs = map (Morphism.thm phi) str_init_def_frees;
-
-    val car_inits = map (mk_min_alg passive_UNIVs str_inits) ks;
-
-    (*TODO: replace with instantiate? (problem: figure out right type instantiation)*)
-    val alg_init_thm = Goal.prove_sorry lthy [] []
-      (HOLogic.mk_Trueprop (mk_alg passive_UNIVs car_inits str_inits))
-      (K (rtac alg_min_alg_thm 1))
-      |> Thm.close_derivation;
-
-    val alg_select_thm = Goal.prove_sorry lthy [] []
-      (HOLogic.mk_Trueprop (mk_Ball II
-        (Term.absfree iidx' (mk_alg passive_UNIVs select_Bs select_ss))))
-      (mk_alg_select_tac Abs_IIT_inverse_thm)
-      |> Thm.close_derivation;
-
-    val mor_select_thm =
-      let
-        val alg_prem = HOLogic.mk_Trueprop (mk_alg passive_UNIVs Bs ss);
-        val i_prem = HOLogic.mk_Trueprop (HOLogic.mk_mem (iidx, II));
-        val mor_prem = HOLogic.mk_Trueprop (mk_mor select_Bs select_ss Bs ss Asuc_fs);
-        val prems = [alg_prem, i_prem, mor_prem];
-        val concl = HOLogic.mk_Trueprop
-          (mk_mor car_inits str_inits Bs ss
-            (map (fn f => HOLogic.mk_comp (f, mk_rapp iidx Asuc_bdT)) Asuc_fs));
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (iidx :: Bs @ ss @ Asuc_fs) (Logic.list_implies (prems, concl)))
-          (K (mk_mor_select_tac mor_def mor_cong_thm mor_comp_thm mor_incl_min_alg_thm alg_def
-            alg_select_thm alg_set_thms set_mapss str_init_defs))
-        |> Thm.close_derivation
-      end;
-
-    val (init_ex_mor_thm, init_unique_mor_thms) =
-      let
-        val prem = HOLogic.mk_Trueprop (mk_alg passive_UNIVs Bs ss);
-        val concl = HOLogic.mk_Trueprop
-          (list_exists_free init_fs (mk_mor car_inits str_inits Bs ss init_fs));
-        val ex_mor = Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (Bs @ ss) (Logic.mk_implies (prem, concl)))
-          (mk_init_ex_mor_tac Abs_IIT_inverse_thm ex_copy_alg_thm alg_min_alg_thm
-            card_of_min_alg_thms mor_comp_thm mor_select_thm mor_incl_min_alg_thm)
-          |> Thm.close_derivation;
-
-        val prems = map2 (HOLogic.mk_Trueprop oo curry HOLogic.mk_mem) init_xs car_inits
-        val mor_prems = map HOLogic.mk_Trueprop
-          [mk_mor car_inits str_inits Bs ss init_fs,
-          mk_mor car_inits str_inits Bs ss init_fs_copy];
-        fun mk_fun_eq f g x = HOLogic.mk_eq (f $ x, g $ x);
-        val unique = HOLogic.mk_Trueprop
-          (Library.foldr1 HOLogic.mk_conj (map3 mk_fun_eq init_fs init_fs_copy init_xs));
-        val unique_mor = Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (init_xs @ Bs @ ss @ init_fs @ init_fs_copy)
-            (Logic.list_implies (prems @ mor_prems, unique)))
-          (K (mk_init_unique_mor_tac m alg_def alg_init_thm least_min_alg_thms
-            in_mono'_thms alg_set_thms morE_thms map_cong0s))
-          |> Thm.close_derivation;
-      in
-        (ex_mor, split_conj_thm unique_mor)
-      end;
-
-    val init_setss = mk_setss (passiveAs @ active_initTs);
-    val active_init_setss = map (drop m) init_setss;
-    val init_ins = map2 (fn sets => mk_in (passive_UNIVs @ car_inits) sets) init_setss init_FTs;
-
-    fun mk_closed phis =
-      let
-        fun mk_conjunct phi str_init init_sets init_in x x' =
-          let
-            val prem = Library.foldr1 HOLogic.mk_conj
-              (map2 (fn set => mk_Ball (set $ x)) init_sets phis);
-            val concl = phi $ (str_init $ x);
-          in
-            mk_Ball init_in (Term.absfree x' (HOLogic.mk_imp (prem, concl)))
-          end;
-      in
-        Library.foldr1 HOLogic.mk_conj
-          (map6 mk_conjunct phis str_inits active_init_setss init_ins init_xFs init_xFs')
-      end;
-
-    val init_induct_thm =
-      let
-        val prem = HOLogic.mk_Trueprop (mk_closed init_phis);
-        val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
-          (map2 mk_Ball car_inits init_phis));
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all init_phis (Logic.mk_implies (prem, concl)))
-          (K (mk_init_induct_tac m alg_def alg_init_thm least_min_alg_thms alg_set_thms))
-        |> Thm.close_derivation
-      end;
-
-    val timer = time (timer "Initiality definition & thms");
-
-    val ((T_names, (T_glob_infos, T_loc_infos)), lthy) =
-      lthy
-      |> fold_map3 (fn b => fn mx => fn car_init =>
-        typedef (Binding.conceal b, params, mx) car_init NONE
-          (EVERY' [rtac ssubst, rtac @{thm ex_in_conv}, resolve_tac alg_not_empty_thms,
-            rtac alg_init_thm] 1)) bs mixfixes car_inits
-      |>> apsnd split_list o split_list;
-
-    val Ts = map (fn name => Type (name, params')) T_names;
-    fun mk_Ts passive = map (Term.typ_subst_atomic (passiveAs ~~ passive)) Ts;
-    val Ts' = mk_Ts passiveBs;
-    val Rep_Ts = map2 (fn info => fn T => Const (#Rep_name info, T --> initT)) T_glob_infos Ts;
-    val Abs_Ts = map2 (fn info => fn T => Const (#Abs_name info, initT --> T)) T_glob_infos Ts;
-
-    val type_defs = map #type_definition T_loc_infos;
-    val Reps = map #Rep T_loc_infos;
-    val Rep_casess = map #Rep_cases T_loc_infos;
-    val Rep_injects = map #Rep_inject T_loc_infos;
-    val Rep_inverses = map #Rep_inverse T_loc_infos;
-    val Abs_inverses = map #Abs_inverse T_loc_infos;
-
-    fun mk_inver_thm mk_tac rep abs X thm =
-      Goal.prove_sorry lthy [] []
-        (HOLogic.mk_Trueprop (mk_inver rep abs X))
-        (K (EVERY' [rtac ssubst, rtac @{thm inver_def}, rtac ballI, mk_tac thm] 1))
-      |> Thm.close_derivation;
-
-    val inver_Reps = map4 (mk_inver_thm rtac) Abs_Ts Rep_Ts (map HOLogic.mk_UNIV Ts) Rep_inverses;
-    val inver_Abss = map4 (mk_inver_thm etac) Rep_Ts Abs_Ts car_inits Abs_inverses;
-
-    val timer = time (timer "THE TYPEDEFs & Rep/Abs thms");
-
-    val UNIVs = map HOLogic.mk_UNIV Ts;
-    val FTs = mk_FTs (passiveAs @ Ts);
-    val FTs' = mk_FTs (passiveBs @ Ts');
-    fun mk_set_Ts T = passiveAs @ replicate n (HOLogic.mk_setT T);
-    val setFTss = map (mk_FTs o mk_set_Ts) passiveAs;
-    val FTs_setss = mk_setss (passiveAs @ Ts);
-    val FTs'_setss = mk_setss (passiveBs @ Ts');
-    val map_FT_inits = map2 (fn Ds =>
-      mk_map_of_bnf Ds (passiveAs @ Ts) (passiveAs @ active_initTs)) Dss bnfs;
-    val fTs = map2 (curry op -->) Ts activeAs;
-    val foldT = Library.foldr1 HOLogic.mk_prodT (map2 (curry op -->) Ts activeAs);
-    val rec_sTs = map (Term.typ_subst_atomic (activeBs ~~ Ts)) prod_sTs;
-    val rec_maps = map (Term.subst_atomic_types (activeBs ~~ Ts)) map_fsts;
-    val rec_maps_rev = map (Term.subst_atomic_types (activeBs ~~ Ts)) map_fsts_rev;
-    val rec_fsts = map (Term.subst_atomic_types (activeBs ~~ Ts)) fsts;
-    val rec_UNIVs = map2 (HOLogic.mk_UNIV oo curry HOLogic.mk_prodT) Ts activeAs;
-
-    val (((((((((Izs1, Izs1'), (Izs2, Izs2')), (xFs, xFs')), yFs), (AFss, AFss')),
-      (fold_f, fold_f')), fs), rec_ss), names_lthy) = names_lthy
-      |> mk_Frees' "z1" Ts
-      ||>> mk_Frees' "z2" Ts'
-      ||>> mk_Frees' "x" FTs
-      ||>> mk_Frees "y" FTs'
-      ||>> mk_Freess' "z" setFTss
-      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "f") foldT
-      ||>> mk_Frees "f" fTs
-      ||>> mk_Frees "s" rec_sTs;
-
-    val Izs = map2 retype_free Ts zs;
-    val phis = map2 retype_free (map mk_pred1T Ts) init_phis;
-    val phi2s = map2 retype_free (map2 mk_pred2T Ts Ts') init_phis;
-
-    fun ctor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctorN ^ "_");
-    val ctor_name = Binding.name_of o ctor_bind;
-    val ctor_def_bind = rpair [] o Binding.conceal o Thm.def_binding o ctor_bind;
-
-    fun ctor_spec i abs str map_FT_init x x' =
-      let
-        val ctorT = nth FTs (i - 1) --> nth Ts (i - 1);
-
-        val lhs = Free (ctor_name i, ctorT);
-        val rhs = Term.absfree x' (abs $ (str $
-          (Term.list_comb (map_FT_init, map HOLogic.id_const passiveAs @ Rep_Ts) $ x)));
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((ctor_frees, (_, ctor_def_frees)), (lthy, lthy_old)) =
-      lthy
-      |> fold_map6 (fn i => fn abs => fn str => fn mapx => fn x => fn x' =>
-        Specification.definition
-          (SOME (ctor_bind i, NONE, NoSyn), (ctor_def_bind i, ctor_spec i abs str mapx x x')))
-          ks Abs_Ts str_inits map_FT_inits xFs xFs'
-      |>> apsnd split_list o split_list
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-    fun mk_ctors passive =
-      map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ (mk_params passive)) o
-        Morphism.term phi) ctor_frees;
-    val ctors = mk_ctors passiveAs;
-    val ctor's = mk_ctors passiveBs;
-    val ctor_defs = map (Morphism.thm phi) ctor_def_frees;
-
-    val (mor_Rep_thm, mor_Abs_thm) =
-      let
-        val copy = alg_init_thm RS copy_alg_thm;
-        fun mk_bij inj Rep cases = @{thm bij_betwI'} OF [inj, Rep, cases];
-        val bijs = map3 mk_bij Rep_injects Reps Rep_casess;
-        val mor_Rep =
-          Goal.prove_sorry lthy [] []
-            (HOLogic.mk_Trueprop (mk_mor UNIVs ctors car_inits str_inits Rep_Ts))
-            (mk_mor_Rep_tac ctor_defs copy bijs inver_Abss inver_Reps)
-          |> Thm.close_derivation;
-
-        val inv = mor_inv_thm OF [mor_Rep, talg_thm, alg_init_thm];
-        val mor_Abs =
-          Goal.prove_sorry lthy [] []
-            (HOLogic.mk_Trueprop (mk_mor car_inits str_inits UNIVs ctors Abs_Ts))
-            (K (mk_mor_Abs_tac inv inver_Abss inver_Reps))
-          |> Thm.close_derivation;
-      in
-        (mor_Rep, mor_Abs)
-      end;
-
-    val timer = time (timer "ctor definitions & thms");
-
-    val fold_fun = Term.absfree fold_f'
-      (mk_mor UNIVs ctors active_UNIVs ss (map (mk_nthN n fold_f) ks));
-    val foldx = HOLogic.choice_const foldT $ fold_fun;
-
-    fun fold_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctor_foldN ^ "_");
-    val fold_name = Binding.name_of o fold_bind;
-    val fold_def_bind = rpair [] o Binding.conceal o Thm.def_binding o fold_bind;
-
-    fun fold_spec i T AT =
-      let
-        val foldT = Library.foldr (op -->) (sTs, T --> AT);
-
-        val lhs = Term.list_comb (Free (fold_name i, foldT), ss);
-        val rhs = mk_nthN n foldx i;
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((fold_frees, (_, fold_def_frees)), (lthy, lthy_old)) =
-      lthy
-      |> fold_map3 (fn i => fn T => fn AT =>
-        Specification.definition
-          (SOME (fold_bind i, NONE, NoSyn), (fold_def_bind i, fold_spec i T AT)))
-          ks Ts activeAs
-      |>> apsnd split_list o split_list
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-    val folds = map (Morphism.term phi) fold_frees;
-    val fold_names = map (fst o dest_Const) folds;
-    fun mk_folds passives actives =
-      map3 (fn name => fn T => fn active =>
-        Const (name, Library.foldr (op -->)
-          (map2 (curry op -->) (mk_FTs (passives @ actives)) actives, T --> active)))
-      fold_names (mk_Ts passives) actives;
-    fun mk_fold Ts ss i = Term.list_comb (Const (nth fold_names (i - 1), Library.foldr (op -->)
-      (map fastype_of ss, nth Ts (i - 1) --> range_type (fastype_of (nth ss (i - 1))))), ss);
-    val fold_defs = map (Morphism.thm phi) fold_def_frees;
-
-    val mor_fold_thm =
-      let
-        val ex_mor = talg_thm RS init_ex_mor_thm;
-        val mor_cong = mor_cong_thm OF (map (mk_nth_conv n) ks);
-        val mor_comp = mor_Rep_thm RS mor_comp_thm;
-        val cT = certifyT lthy foldT;
-        val ct = certify lthy fold_fun
-      in
-        singleton (Proof_Context.export names_lthy lthy)
-          (Goal.prove_sorry lthy [] []
-            (HOLogic.mk_Trueprop (mk_mor UNIVs ctors active_UNIVs ss (map (mk_fold Ts ss) ks)))
-            (K (mk_mor_fold_tac cT ct fold_defs ex_mor (mor_comp RS mor_cong))))
-        |> Thm.close_derivation
-      end;
-
-    val ctor_fold_thms = map (fn morE => rule_by_tactic lthy
-      ((rtac CollectI THEN' CONJ_WRAP' (K (rtac @{thm subset_UNIV})) (1 upto m + n)) 1)
-      (mor_fold_thm RS morE)) morE_thms;
-
-    val (fold_unique_mor_thms, fold_unique_mor_thm) =
-      let
-        val prem = HOLogic.mk_Trueprop (mk_mor UNIVs ctors active_UNIVs ss fs);
-        fun mk_fun_eq f i = HOLogic.mk_eq (f, mk_fold Ts ss i);
-        val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_fun_eq fs ks));
-        val unique_mor = Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (ss @ fs) (Logic.mk_implies (prem, unique)))
-          (K (mk_fold_unique_mor_tac type_defs init_unique_mor_thms Reps
-            mor_comp_thm mor_Abs_thm mor_fold_thm))
-          |> Thm.close_derivation;
-      in
-        `split_conj_thm unique_mor
-      end;
-
-    val (ctor_fold_unique_thms, ctor_fold_unique_thm) =
-      `split_conj_thm (mk_conjIN n RS
-        (mor_UNIV_thm RS iffD2 RS fold_unique_mor_thm))
-
-    val fold_ctor_thms =
-      map (fn thm => (mor_incl_thm OF replicate n @{thm subset_UNIV}) RS thm RS sym)
-        fold_unique_mor_thms;
-
-    val ctor_o_fold_thms =
-      let
-        val mor = mor_comp_thm OF [mor_fold_thm, mor_str_thm];
-      in
-        map2 (fn unique => fn fold_ctor =>
-          trans OF [mor RS unique, fold_ctor]) fold_unique_mor_thms fold_ctor_thms
-      end;
-
-    val timer = time (timer "fold definitions & thms");
-
-    val map_ctors = map2 (fn Ds => fn bnf =>
-      Term.list_comb (mk_map_of_bnf Ds (passiveAs @ FTs) (passiveAs @ Ts) bnf,
-        map HOLogic.id_const passiveAs @ ctors)) Dss bnfs;
-
-    fun dtor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtorN ^ "_");
-    val dtor_name = Binding.name_of o dtor_bind;
-    val dtor_def_bind = rpair [] o Binding.conceal o Thm.def_binding o dtor_bind;
-
-    fun dtor_spec i FT T =
-      let
-        val dtorT = T --> FT;
-
-        val lhs = Free (dtor_name i, dtorT);
-        val rhs = mk_fold Ts map_ctors i;
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((dtor_frees, (_, dtor_def_frees)), (lthy, lthy_old)) =
-      lthy
-      |> fold_map3 (fn i => fn FT => fn T =>
-        Specification.definition
-          (SOME (dtor_bind i, NONE, NoSyn), (dtor_def_bind i, dtor_spec i FT T))) ks FTs Ts
-      |>> apsnd split_list o split_list
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-    fun mk_dtors params =
-      map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ params) o Morphism.term phi)
-        dtor_frees;
-    val dtors = mk_dtors params';
-    val dtor_defs = map (Morphism.thm phi) dtor_def_frees;
-
-    val ctor_o_dtor_thms = map2 (fold_thms lthy o single) dtor_defs ctor_o_fold_thms;
-
-    val dtor_o_ctor_thms =
-      let
-        fun mk_goal dtor ctor FT =
-          mk_Trueprop_eq (HOLogic.mk_comp (dtor, ctor), HOLogic.id_const FT);
-        val goals = map3 mk_goal dtors ctors FTs;
-      in
-        map5 (fn goal => fn dtor_def => fn foldx => fn map_comp_id => fn map_cong0L =>
-          Goal.prove_sorry lthy [] [] goal
-            (K (mk_dtor_o_ctor_tac dtor_def foldx map_comp_id map_cong0L ctor_o_fold_thms))
-          |> Thm.close_derivation)
-        goals dtor_defs ctor_fold_thms map_comp_id_thms map_cong0L_thms
-      end;
-
-    val dtor_ctor_thms = map (fn thm => thm RS @{thm pointfree_idE}) dtor_o_ctor_thms;
-    val ctor_dtor_thms = map (fn thm => thm RS @{thm pointfree_idE}) ctor_o_dtor_thms;
-
-    val bij_dtor_thms =
-      map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) ctor_o_dtor_thms dtor_o_ctor_thms;
-    val inj_dtor_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_dtor_thms;
-    val surj_dtor_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_dtor_thms;
-    val dtor_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_dtor_thms;
-    val dtor_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_dtor_thms;
-    val dtor_exhaust_thms = map (fn thm => thm RS exE) dtor_nchotomy_thms;
-
-    val bij_ctor_thms =
-      map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) dtor_o_ctor_thms ctor_o_dtor_thms;
-    val inj_ctor_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_ctor_thms;
-    val surj_ctor_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_ctor_thms;
-    val ctor_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_ctor_thms;
-    val ctor_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_ctor_thms;
-    val ctor_exhaust_thms = map (fn thm => thm RS exE) ctor_nchotomy_thms;
-
-    val timer = time (timer "dtor definitions & thms");
-
-    val fst_rec_pair_thms =
-      let
-        val mor = mor_comp_thm OF [mor_fold_thm, mor_convol_thm];
-      in
-        map2 (fn unique => fn fold_ctor =>
-          trans OF [mor RS unique, fold_ctor]) fold_unique_mor_thms fold_ctor_thms
-      end;
-
-    fun rec_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctor_recN ^ "_");
-    val rec_name = Binding.name_of o rec_bind;
-    val rec_def_bind = rpair [] o Binding.conceal o Thm.def_binding o rec_bind;
-
-    val rec_strs =
-      map3 (fn ctor => fn prod_s => fn mapx =>
-        mk_convol (HOLogic.mk_comp (ctor, Term.list_comb (mapx, passive_ids @ rec_fsts)), prod_s))
-      ctors rec_ss rec_maps;
-
-    fun rec_spec i T AT =
-      let
-        val recT = Library.foldr (op -->) (rec_sTs, T --> AT);
-
-        val lhs = Term.list_comb (Free (rec_name i, recT), rec_ss);
-        val rhs = HOLogic.mk_comp (snd_const (HOLogic.mk_prodT (T, AT)), mk_fold Ts rec_strs i);
-      in
-        mk_Trueprop_eq (lhs, rhs)
-      end;
-
-    val ((rec_frees, (_, rec_def_frees)), (lthy, lthy_old)) =
-      lthy
-      |> fold_map3 (fn i => fn T => fn AT =>
-        Specification.definition
-          (SOME (rec_bind i, NONE, NoSyn), (rec_def_bind i, rec_spec i T AT)))
-          ks Ts activeAs
-      |>> apsnd split_list o split_list
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-    val recs = map (Morphism.term phi) rec_frees;
-    val rec_names = map (fst o dest_Const) recs;
-    fun mk_rec ss i = Term.list_comb (Const (nth rec_names (i - 1), Library.foldr (op -->)
-      (map fastype_of ss, nth Ts (i - 1) --> range_type (fastype_of (nth ss (i - 1))))), ss);
-    val rec_defs = map (Morphism.thm phi) rec_def_frees;
-
-    val convols = map2 (fn T => fn i => mk_convol (HOLogic.id_const T, mk_rec rec_ss i)) Ts ks;
-    val ctor_rec_thms =
-      let
-        fun mk_goal i rec_s rec_map ctor x =
-          let
-            val lhs = mk_rec rec_ss i $ (ctor $ x);
-            val rhs = rec_s $ (Term.list_comb (rec_map, passive_ids @ convols) $ x);
-          in
-            fold_rev Logic.all (x :: rec_ss) (mk_Trueprop_eq (lhs, rhs))
-          end;
-        val goals = map5 mk_goal ks rec_ss rec_maps_rev ctors xFs;
-      in
-        map2 (fn goal => fn foldx =>
-          Goal.prove_sorry lthy [] [] goal (mk_rec_tac rec_defs foldx fst_rec_pair_thms)
-          |> Thm.close_derivation)
-        goals ctor_fold_thms
-      end;
-
-    val rec_unique_mor_thm =
-      let
-        val id_fs = map2 (fn T => fn f => mk_convol (HOLogic.id_const T, f)) Ts fs;
-        val prem = HOLogic.mk_Trueprop (mk_mor UNIVs ctors rec_UNIVs rec_strs id_fs);
-        fun mk_fun_eq f i = HOLogic.mk_eq (f, mk_rec rec_ss i);
-        val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_fun_eq fs ks));
-      in
-        Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (rec_ss @ fs) (Logic.mk_implies (prem, unique)))
-          (mk_rec_unique_mor_tac rec_defs fst_rec_pair_thms fold_unique_mor_thm)
-          |> Thm.close_derivation
-      end;
-
-    val (ctor_rec_unique_thms, ctor_rec_unique_thm) =
-      `split_conj_thm (split_conj_prems n
-        (mor_UNIV_thm RS iffD2 RS rec_unique_mor_thm)
-        |> Local_Defs.unfold lthy (@{thms convol_o o_id id_o o_assoc[symmetric] fst_convol} @
-           map_id0s @ sym_map_comps) OF replicate n @{thm arg_cong2[of _ _ _ _ convol, OF refl]});
-
-    val timer = time (timer "rec definitions & thms");
-
-    val (ctor_induct_thm, induct_params) =
-      let
-        fun mk_prem phi ctor sets x =
-          let
-            fun mk_IH phi set z =
-              let
-                val prem = HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set $ x));
-                val concl = HOLogic.mk_Trueprop (phi $ z);
-              in
-                Logic.all z (Logic.mk_implies (prem, concl))
-              end;
-
-            val IHs = map3 mk_IH phis (drop m sets) Izs;
-            val concl = HOLogic.mk_Trueprop (phi $ (ctor $ x));
-          in
-            Logic.all x (Logic.list_implies (IHs, concl))
-          end;
-
-        val prems = map4 mk_prem phis ctors FTs_setss xFs;
-
-        fun mk_concl phi z = phi $ z;
-        val concl =
-          HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_concl phis Izs));
-
-        val goal = Logic.list_implies (prems, concl);
-      in
-        (Goal.prove_sorry lthy [] []
-          (fold_rev Logic.all (phis @ Izs) goal)
-          (K (mk_ctor_induct_tac lthy m set_mapss init_induct_thm morE_thms mor_Abs_thm
-            Rep_inverses Abs_inverses Reps))
-        |> Thm.close_derivation,
-        rev (Term.add_tfrees goal []))
-      end;
-
-    val cTs = map (SOME o certifyT lthy o TFree) induct_params;
-
-    val weak_ctor_induct_thms =
-      let fun insts i = (replicate (i - 1) TrueI) @ (asm_rl :: replicate (n - i) TrueI);
-      in map (fn i => (ctor_induct_thm OF insts i) RS mk_conjunctN n i) ks end;
-
-    val (ctor_induct2_thm, induct2_params) =
-      let
-        fun mk_prem phi ctor ctor' sets sets' x y =
-          let
-            fun mk_IH phi set set' z1 z2 =
-              let
-                val prem1 = HOLogic.mk_Trueprop (HOLogic.mk_mem (z1, (set $ x)));
-                val prem2 = HOLogic.mk_Trueprop (HOLogic.mk_mem (z2, (set' $ y)));
-                val concl = HOLogic.mk_Trueprop (phi $ z1 $ z2);
-              in
-                fold_rev Logic.all [z1, z2] (Logic.list_implies ([prem1, prem2], concl))
-              end;
-
-            val IHs = map5 mk_IH phi2s (drop m sets) (drop m sets') Izs1 Izs2;
-            val concl = HOLogic.mk_Trueprop (phi $ (ctor $ x) $ (ctor' $ y));
-          in
-            fold_rev Logic.all [x, y] (Logic.list_implies (IHs, concl))
-          end;
-
-        val prems = map7 mk_prem phi2s ctors ctor's FTs_setss FTs'_setss xFs yFs;
-
-        fun mk_concl phi z1 z2 = phi $ z1 $ z2;
-        val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
-          (map3 mk_concl phi2s Izs1 Izs2));
-        fun mk_t phi (z1, z1') (z2, z2') =
-          Term.absfree z1' (HOLogic.mk_all (fst z2', snd z2', phi $ z1 $ z2));
-        val cts = map3 (SOME o certify lthy ooo mk_t) phi2s (Izs1 ~~ Izs1') (Izs2 ~~ Izs2');
-        val goal = Logic.list_implies (prems, concl);
-      in
-        (singleton (Proof_Context.export names_lthy lthy)
-          (Goal.prove_sorry lthy [] [] goal
-            (mk_ctor_induct2_tac cTs cts ctor_induct_thm weak_ctor_induct_thms))
-          |> Thm.close_derivation,
-        rev (Term.add_tfrees goal []))
-      end;
-
-    val timer = time (timer "induction");
-
-    fun mk_ctor_map_DEADID_thm ctor_inject map_id0 =
-      trans OF [id_apply, iffD2 OF [ctor_inject, map_id0 RS sym]];
-
-    fun mk_ctor_Irel_DEADID_thm ctor_inject bnf =
-      trans OF [ctor_inject, rel_eq_of_bnf bnf RS @{thm predicate2_eqD} RS sym];
-
-    val IphiTs = map2 mk_pred2T passiveAs passiveBs;
-    val Ipsi1Ts = map2 mk_pred2T passiveAs passiveCs;
-    val Ipsi2Ts = map2 mk_pred2T passiveCs passiveBs;
-    val activephiTs = map2 mk_pred2T activeAs activeBs;
-    val activeIphiTs = map2 mk_pred2T Ts Ts';
-    val (((((Iphis, Ipsi1s), Ipsi2s), activephis), activeIphis), names_lthy) = names_lthy
-      |> mk_Frees "R" IphiTs
-      ||>> mk_Frees "R" Ipsi1Ts
-      ||>> mk_Frees "Q" Ipsi2Ts
-      ||>> mk_Frees "S" activephiTs
-      ||>> mk_Frees "IR" activeIphiTs;
-    val rels = map2 (fn Ds => mk_rel_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
-
-    (*register new datatypes as BNFs*)
-    val (timer, Ibnfs, (ctor_Imap_o_thms, ctor_Imap_thms), ctor_Iset_thmss',
-        ctor_Irel_thms, Ibnf_notes, lthy) =
-      if m = 0 then
-        (timer, replicate n DEADID_bnf,
-        map_split (`(mk_pointfree lthy)) (map2 mk_ctor_map_DEADID_thm ctor_inject_thms map_ids),
-        replicate n [], map2 mk_ctor_Irel_DEADID_thm ctor_inject_thms bnfs, [], lthy)
-      else let
-        val fTs = map2 (curry op -->) passiveAs passiveBs;
-        val uTs = map2 (curry op -->) Ts Ts';
-
-        val (((((fs, fs'), fs_copy), us), (ys, ys')),
-          names_lthy) = names_lthy
-          |> mk_Frees' "f" fTs
-          ||>> mk_Frees "f" fTs
-          ||>> mk_Frees "u" uTs
-          ||>> mk_Frees' "y" passiveAs;
-
-        val map_FTFT's = map2 (fn Ds =>
-          mk_map_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
-        fun mk_passive_maps ATs BTs Ts =
-          map2 (fn Ds => mk_map_of_bnf Ds (ATs @ Ts) (BTs @ Ts)) Dss bnfs;
-        fun mk_map_fold_arg fs Ts ctor fmap =
-          HOLogic.mk_comp (ctor, Term.list_comb (fmap, fs @ map HOLogic.id_const Ts));
-        fun mk_map Ts fs Ts' ctors mk_maps =
-          mk_fold Ts (map2 (mk_map_fold_arg fs Ts') ctors (mk_maps Ts'));
-        val pmapsABT' = mk_passive_maps passiveAs passiveBs;
-        val fs_maps = map (mk_map Ts fs Ts' ctor's pmapsABT') ks;
-
-        val ls = 1 upto m;
-        val setsss = map (mk_setss o mk_set_Ts) passiveAs;
-
-        fun mk_col l T z z' sets =
-          let
-            fun mk_UN set = mk_Union T $ (set $ z);
-          in
-            Term.absfree z'
-              (mk_union (nth sets (l - 1) $ z,
-                Library.foldl1 mk_union (map mk_UN (drop m sets))))
-          end;
-
-        val colss = map5 (fn l => fn T => map3 (mk_col l T)) ls passiveAs AFss AFss' setsss;
-        val setss_by_range = map (fn cols => map (mk_fold Ts cols) ks) colss;
-        val setss_by_bnf = transpose setss_by_range;
-
-        val set_bss =
-          map (flat o map2 (fn B => fn b =>
-            if member (op =) resDs (TFree B) then [] else [b]) resBs) set_bss0;
-
-        val ctor_witss =
-          let
-            val witss = map2 (fn Ds => fn bnf => mk_wits_of_bnf
-              (replicate (nwits_of_bnf bnf) Ds)
-              (replicate (nwits_of_bnf bnf) (passiveAs @ Ts)) bnf) Dss bnfs;
-            fun close_wit (I, wit) = fold_rev Term.absfree (map (nth ys') I) wit;
-            fun wit_apply (arg_I, arg_wit) (fun_I, fun_wit) =
-              (union (op =) arg_I fun_I, fun_wit $ arg_wit);
-
-            fun gen_arg support i =
-              if i < m then [([i], nth ys i)]
-              else maps (mk_wit support (nth ctors (i - m)) (i - m)) (nth support (i - m))
-            and mk_wit support ctor i (I, wit) =
-              let val args = map (gen_arg (nth_map i (remove (op =) (I, wit)) support)) I;
-              in
-                (args, [([], wit)])
-                |-> fold (map_product wit_apply)
-                |> map (apsnd (fn t => ctor $ t))
-                |> minimize_wits
-              end;
-          in
-            map3 (fn ctor => fn i => map close_wit o minimize_wits o maps (mk_wit witss ctor i))
-              ctors (0 upto n - 1) witss
-          end;
-
-        val (Ibnf_consts, lthy) =
-          fold_map8 (fn b => fn map_b => fn rel_b => fn set_bs => fn mapx => fn sets => fn wits =>
-              fn T => fn lthy =>
-            define_bnf_consts Dont_Inline (user_policy Note_Some lthy) (SOME deads)
-              map_b rel_b set_bs
-              ((((((b, T), fold_rev Term.absfree fs' mapx), sets), sum_bd), wits), NONE) lthy)
-          bs map_bs rel_bs set_bss fs_maps setss_by_bnf ctor_witss Ts lthy;
-
-        val (_, Iconsts, Iconst_defs, mk_Iconsts) = split_list4 Ibnf_consts;
-        val (_, Isetss, Ibds_Ds, Iwitss_Ds, _) = split_list5 Iconsts;
-        val (Imap_defs, Iset_defss, Ibd_defs, Iwit_defss, Irel_defs) = split_list5 Iconst_defs;
-        val (mk_Imaps_Ds, mk_It_Ds, _, mk_Irels_Ds, _) = split_list5 mk_Iconsts;
-
-        val Irel_unabs_defs = map (fn def => mk_unabs_def m (def RS meta_eq_to_obj_eq)) Irel_defs;
-        val Iset_defs = flat Iset_defss;
-
-        fun mk_Imaps As Bs = map (fn mk => mk deads As Bs) mk_Imaps_Ds;
-        fun mk_Isetss As = map2 (fn mk => fn Isets => map (mk deads As) Isets) mk_It_Ds Isetss;
-        val Ibds = map2 (fn mk => mk deads passiveAs) mk_It_Ds Ibds_Ds;
-        val Iwitss =
-          map2 (fn mk => fn Iwits => map (mk deads passiveAs o snd) Iwits) mk_It_Ds Iwitss_Ds;
-        fun mk_Irels As Bs = map (fn mk => mk deads As Bs) mk_Irels_Ds;
-
-        val Imaps = mk_Imaps passiveAs passiveBs;
-        val fs_Imaps = map (fn m => Term.list_comb (m, fs)) Imaps;
-        val fs_copy_Imaps = map (fn m => Term.list_comb (m, fs_copy)) Imaps;
-        val (Isetss_by_range, Isetss_by_bnf) = `transpose (mk_Isetss passiveAs);
-
-        val map_setss = map (fn T => map2 (fn Ds =>
-          mk_map_of_bnf Ds (passiveAs @ Ts) (mk_set_Ts T)) Dss bnfs) passiveAs;
-
-        val timer = time (timer "bnf constants for the new datatypes");
-
-        val (ctor_Imap_thms, ctor_Imap_o_thms) =
-          let
-            fun mk_goal fs_map map ctor ctor' = fold_rev Logic.all fs
-              (mk_Trueprop_eq (HOLogic.mk_comp (fs_map, ctor),
-                HOLogic.mk_comp (ctor', Term.list_comb (map, fs @ fs_Imaps))));
-            val goals = map4 mk_goal fs_Imaps map_FTFT's ctors ctor's;
-            val maps =
-              map4 (fn goal => fn foldx => fn map_comp_id => fn map_cong0 =>
-                Goal.prove_sorry lthy [] [] goal
-                  (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Imap_defs THEN
-                    mk_map_tac m n foldx map_comp_id map_cong0)
-                |> Thm.close_derivation)
-              goals ctor_fold_thms map_comp_id_thms map_cong0s;
-          in
-            `(map (fn thm => thm RS @{thm comp_eq_dest})) maps
-          end;
-
-        val (ctor_Imap_unique_thms, ctor_Imap_unique_thm) =
-          let
-            fun mk_prem u map ctor ctor' =
-              mk_Trueprop_eq (HOLogic.mk_comp (u, ctor),
-                HOLogic.mk_comp (ctor', Term.list_comb (map, fs @ us)));
-            val prems = map4 mk_prem us map_FTFT's ctors ctor's;
-            val goal =
-              HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
-                (map2 (curry HOLogic.mk_eq) us fs_Imaps));
-            val unique = Goal.prove_sorry lthy [] []
-              (fold_rev Logic.all (us @ fs) (Logic.list_implies (prems, goal)))
-              (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Imap_defs THEN
-                mk_ctor_map_unique_tac ctor_fold_unique_thm sym_map_comps ctxt)
-              |> Thm.close_derivation;
-          in
-            `split_conj_thm unique
-          end;
-
-        val timer = time (timer "map functions for the new datatypes");
-
-        val ctor_Iset_thmss =
-          let
-            fun mk_goal sets ctor set col map =
-              mk_Trueprop_eq (HOLogic.mk_comp (set, ctor),
-                HOLogic.mk_comp (col, Term.list_comb (map, passive_ids @ sets)));
-            val goalss =
-              map3 (fn sets => map4 (mk_goal sets) ctors sets) Isetss_by_range colss map_setss;
-            val setss = map (map2 (fn foldx => fn goal =>
-                Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} =>
-                  unfold_thms_tac ctxt Iset_defs THEN mk_set_tac foldx)
-                |> Thm.close_derivation)
-              ctor_fold_thms) goalss;
-
-            fun mk_simp_goal pas_set act_sets sets ctor z set =
-              Logic.all z (mk_Trueprop_eq (set $ (ctor $ z),
-                mk_union (pas_set $ z,
-                  Library.foldl1 mk_union (map2 (fn X => mk_UNION (X $ z)) act_sets sets))));
-            val simp_goalss =
-              map2 (fn i => fn sets =>
-                map4 (fn Fsets => mk_simp_goal (nth Fsets (i - 1)) (drop m Fsets) sets)
-                  FTs_setss ctors xFs sets)
-                ls Isetss_by_range;
-
-            val ctor_setss = map3 (fn i => map3 (fn set_nats => fn goal => fn set =>
-                Goal.prove_sorry lthy [] [] goal
-                  (K (mk_ctor_set_tac set (nth set_nats (i - 1)) (drop m set_nats)))
-                |> Thm.close_derivation)
-              set_mapss) ls simp_goalss setss;
-          in
-            ctor_setss
-          end;
-
-        fun mk_set_thms ctor_set = (@{thm xt1(3)} OF [ctor_set, @{thm Un_upper1}]) ::
-          map (fn i => (@{thm xt1(3)} OF [ctor_set, @{thm Un_upper2}]) RS
-            (mk_Un_upper n i RS subset_trans) RSN
-            (2, @{thm UN_upper} RS subset_trans))
-            (1 upto n);
-        val set_Iset_thmsss = transpose (map (map mk_set_thms) ctor_Iset_thmss);
-
-        val timer = time (timer "set functions for the new datatypes");
-
-        val cxs = map (SOME o certify lthy) Izs;
-        val Isetss_by_range' =
-          map (map (Term.subst_atomic_types (passiveAs ~~ passiveBs))) Isetss_by_range;
-
-        val Iset_Imap0_thmss =
-          let
-            fun mk_set_map0 f map z set set' =
-              HOLogic.mk_eq (mk_image f $ (set $ z), set' $ (map $ z));
-
-            fun mk_cphi f map z set set' = certify lthy
-              (Term.absfree (dest_Free z) (mk_set_map0 f map z set set'));
-
-            val csetss = map (map (certify lthy)) Isetss_by_range';
-
-            val cphiss = map3 (fn f => fn sets => fn sets' =>
-              (map4 (mk_cphi f) fs_Imaps Izs sets sets')) fs Isetss_by_range Isetss_by_range';
-
-            val inducts = map (fn cphis =>
-              Drule.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm) cphiss;
-
-            val goals =
-              map3 (fn f => fn sets => fn sets' =>
-                HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
-                  (map4 (mk_set_map0 f) fs_Imaps Izs sets sets')))
-                  fs Isetss_by_range Isetss_by_range';
-
-            fun mk_tac induct = mk_set_nat_tac m (rtac induct) set_mapss ctor_Imap_thms;
-            val thms =
-              map5 (fn goal => fn csets => fn ctor_sets => fn induct => fn i =>
-                singleton (Proof_Context.export names_lthy lthy)
-                  (Goal.prove_sorry lthy [] [] goal (mk_tac induct csets ctor_sets i))
-                |> Thm.close_derivation)
-              goals csetss ctor_Iset_thmss inducts ls;
-          in
-            map split_conj_thm thms
-          end;
-
-        val Iset_bd_thmss =
-          let
-            fun mk_set_bd z bd set = mk_ordLeq (mk_card_of (set $ z)) bd;
-
-            fun mk_cphi z set = certify lthy (Term.absfree (dest_Free z) (mk_set_bd z sum_bd set));
-
-            val cphiss = map (map2 mk_cphi Izs) Isetss_by_range;
-
-            val inducts = map (fn cphis =>
-              Drule.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm) cphiss;
-
-            val goals =
-              map (fn sets =>
-                HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
-                  (map3 mk_set_bd Izs Ibds sets))) Isetss_by_range;
-
-            fun mk_tac induct = mk_set_bd_tac m (rtac induct) sum_Cinfinite set_bd_sumss;
-            val thms =
-              map4 (fn goal => fn ctor_sets => fn induct => fn i =>
-                singleton (Proof_Context.export names_lthy lthy)
-                  (Goal.prove_sorry lthy [] [] goal
-                    (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Ibd_defs THEN
-                      mk_tac induct ctor_sets i ctxt))
-                |> Thm.close_derivation)
-              goals ctor_Iset_thmss inducts ls;
-          in
-            map split_conj_thm thms
-          end;
-
-        val Imap_cong0_thms =
-          let
-            fun mk_prem z set f g y y' =
-              mk_Ball (set $ z) (Term.absfree y' (HOLogic.mk_eq (f $ y, g $ y)));
-
-            fun mk_map_cong0 sets z fmap gmap =
-              HOLogic.mk_imp
-                (Library.foldr1 HOLogic.mk_conj (map5 (mk_prem z) sets fs fs_copy ys ys'),
-                HOLogic.mk_eq (fmap $ z, gmap $ z));
-
-            fun mk_cphi sets z fmap gmap =
-              certify lthy (Term.absfree (dest_Free z) (mk_map_cong0 sets z fmap gmap));
-
-            val cphis = map4 mk_cphi Isetss_by_bnf Izs fs_Imaps fs_copy_Imaps;
-
-            val induct = Drule.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm;
-
-            val goal =
-              HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
-                (map4 mk_map_cong0 Isetss_by_bnf Izs fs_Imaps fs_copy_Imaps));
-
-            val thm = singleton (Proof_Context.export names_lthy lthy)
-              (Goal.prove_sorry lthy [] [] goal
-              (mk_mcong_tac (rtac induct) set_Iset_thmsss map_cong0s ctor_Imap_thms))
-              |> Thm.close_derivation;
-          in
-            split_conj_thm thm
-          end;
-
-        val in_rels = map in_rel_of_bnf bnfs;
-        val in_Irels = map (fn def => trans OF [def, @{thm OO_Grp_alt}] RS @{thm predicate2_eqD})
-            Irel_unabs_defs;
-
-        val ctor_Iset_incl_thmss = map (map hd) set_Iset_thmsss;
-        val ctor_set_Iset_incl_thmsss = map (transpose o map tl) set_Iset_thmsss;
-        val ctor_Iset_thmss' = transpose ctor_Iset_thmss;
-
-        val Irels = mk_Irels passiveAs passiveBs;
-        val Irelphis = map (fn rel => Term.list_comb (rel, Iphis)) Irels;
-        val relphis = map (fn rel => Term.list_comb (rel, Iphis @ Irelphis)) rels;
-        val Irelpsi1s = map (fn rel => Term.list_comb (rel, Ipsi1s)) (mk_Irels passiveAs passiveCs);
-        val Irelpsi2s = map (fn rel => Term.list_comb (rel, Ipsi2s)) (mk_Irels passiveCs passiveBs);
-        val Irelpsi12s = map (fn rel =>
-            Term.list_comb (rel, map2 (curry mk_rel_compp) Ipsi1s Ipsi2s)) Irels;
-
-        val ctor_Irel_thms =
-          let
-            fun mk_goal xF yF ctor ctor' Irelphi relphi = fold_rev Logic.all (xF :: yF :: Iphis)
-              (mk_Trueprop_eq (Irelphi $ (ctor $ xF) $ (ctor' $ yF), relphi $ xF $ yF));
-            val goals = map6 mk_goal xFs yFs ctors ctor's Irelphis relphis;
-          in
-            map12 (fn i => fn goal => fn in_rel => fn map_comp0 => fn map_cong0 =>
-              fn ctor_map => fn ctor_sets => fn ctor_inject => fn ctor_dtor =>
-              fn set_map0s => fn ctor_set_incls => fn ctor_set_set_inclss =>
-              Goal.prove_sorry lthy [] [] goal
-               (K (mk_ctor_rel_tac lthy in_Irels i in_rel map_comp0 map_cong0 ctor_map ctor_sets
-                 ctor_inject ctor_dtor set_map0s ctor_set_incls ctor_set_set_inclss))
-              |> Thm.close_derivation)
-            ks goals in_rels map_comps map_cong0s ctor_Imap_thms ctor_Iset_thmss'
-              ctor_inject_thms ctor_dtor_thms set_mapss ctor_Iset_incl_thmss
-              ctor_set_Iset_incl_thmsss
-          end;
-
-        val le_Irel_OO_thm =
-          let
-            fun mk_le_Irel_OO Irelpsi1 Irelpsi2 Irelpsi12 Iz1 Iz2 =
-              HOLogic.mk_imp (mk_rel_compp (Irelpsi1, Irelpsi2) $ Iz1 $ Iz2,
-                Irelpsi12 $ Iz1 $ Iz2);
-            val goals = map5 mk_le_Irel_OO Irelpsi1s Irelpsi2s Irelpsi12s Izs1 Izs2;
-
-            val cTs = map (SOME o certifyT lthy o TFree) induct2_params;
-            val cxs = map (SOME o certify lthy) (splice Izs1 Izs2);
-            fun mk_cphi z1 z2 goal = SOME (certify lthy (Term.absfree z1 (Term.absfree z2 goal)));
-            val cphis = map3 mk_cphi Izs1' Izs2' goals;
-            val induct = Drule.instantiate' cTs (cphis @ cxs) ctor_induct2_thm;
-
-            val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj goals);
-          in
-            singleton (Proof_Context.export names_lthy lthy)
-              (Goal.prove_sorry lthy [] [] goal
-                (mk_le_rel_OO_tac m induct ctor_nchotomy_thms ctor_Irel_thms rel_mono_strongs
-                  rel_OOs))
-              |> Thm.close_derivation
-          end;
-
-        val timer = time (timer "helpers for BNF properties");
-
-        val map_id0_tacs = map (K o mk_map_id0_tac map_id0s) ctor_Imap_unique_thms;
-        val map_comp0_tacs =
-          map2 (K oo mk_map_comp0_tac map_comps ctor_Imap_thms) ctor_Imap_unique_thms ks;
-        val map_cong0_tacs = map (mk_map_cong0_tac m) Imap_cong0_thms;
-        val set_map0_tacss = map (map (K o mk_set_map0_tac)) (transpose Iset_Imap0_thmss);
-        val bd_co_tacs = replicate n (fn {context = ctxt, prems = _} =>
-          unfold_thms_tac ctxt Ibd_defs THEN mk_bd_card_order_tac bd_card_orders);
-        val bd_cinf_tacs = replicate n (fn {context = ctxt, prems = _} =>
-          unfold_thms_tac ctxt Ibd_defs THEN rtac (sum_Cinfinite RS conjunct1) 1);
-        val set_bd_tacss = map (map (fn thm => K (rtac thm 1))) (transpose Iset_bd_thmss);
-        val le_rel_OO_tacs = map (fn i =>
-          K ((rtac @{thm predicate2I} THEN' etac (le_Irel_OO_thm RS mk_conjunctN n i RS mp)) 1)) ks;
-
-        val rel_OO_Grp_tacs = map (fn def => K (rtac def 1)) Irel_unabs_defs;
-
-        val tacss = map9 zip_axioms map_id0_tacs map_comp0_tacs map_cong0_tacs set_map0_tacss
-          bd_co_tacs bd_cinf_tacs set_bd_tacss le_rel_OO_tacs rel_OO_Grp_tacs;
-
-        fun wit_tac {context = ctxt, prems = _} = unfold_thms_tac ctxt (flat Iwit_defss) THEN
-          mk_wit_tac ctxt n (flat ctor_Iset_thmss) (maps wit_thms_of_bnf bnfs);
-
-        val (Ibnfs, lthy) =
-          fold_map5 (fn tacs => fn map_b => fn rel_b => fn set_bs => fn consts => fn lthy =>
-            bnf_def Do_Inline (user_policy Note_Some) I tacs wit_tac (SOME deads)
-              map_b rel_b set_bs consts lthy
-            |> register_bnf (Local_Theory.full_name lthy b))
-          tacss map_bs rel_bs set_bss
-            ((((((bs ~~ Ts) ~~ Imaps) ~~ Isetss_by_bnf) ~~ Ibds) ~~ Iwitss) ~~ map SOME Irels)
-            lthy;
-
-        val timer = time (timer "registered new datatypes as BNFs");
-
-        val ls' = if m = 1 then [0] else ls
-
-        val Ibnf_common_notes =
-          [(ctor_map_uniqueN, [ctor_Imap_unique_thm])]
-          |> map (fn (thmN, thms) =>
-            ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]));
-
-        val Ibnf_notes =
-          [(ctor_mapN, map single ctor_Imap_thms),
-          (ctor_relN, map single ctor_Irel_thms),
-          (ctor_set_inclN, ctor_Iset_incl_thmss),
-          (ctor_set_set_inclN, map flat ctor_set_Iset_incl_thmsss)] @
-          map2 (fn i => fn thms => (mk_ctor_setN i, map single thms)) ls' ctor_Iset_thmss
-          |> maps (fn (thmN, thmss) =>
-            map2 (fn b => fn thms =>
-              ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]))
-            bs thmss)
-      in
-        (timer, Ibnfs, (ctor_Imap_o_thms, ctor_Imap_thms), ctor_Iset_thmss',
-          ctor_Irel_thms, Ibnf_common_notes @ Ibnf_notes, lthy)
-      end;
-
-      val ctor_fold_o_Imap_thms = mk_xtor_un_fold_o_map_thms Least_FP false m ctor_fold_unique_thm
-        ctor_Imap_o_thms (map (mk_pointfree lthy) ctor_fold_thms) sym_map_comps map_cong0s;
-      val ctor_rec_o_Imap_thms = mk_xtor_un_fold_o_map_thms Least_FP true m ctor_rec_unique_thm
-        ctor_Imap_o_thms (map (mk_pointfree lthy) ctor_rec_thms) sym_map_comps map_cong0s;
-
-      val Irels = if m = 0 then map HOLogic.eq_const Ts
-        else map (mk_rel_of_bnf deads passiveAs passiveBs) Ibnfs;
-      val Irel_induct_thm =
-        mk_rel_xtor_co_induct_thm Least_FP rels activeIphis Irels Iphis xFs yFs ctors ctor's
-          (mk_rel_induct_tac m ctor_induct2_thm ks ctor_Irel_thms rel_mono_strongs) lthy;
-
-      val rels = map2 (fn Ds => mk_rel_of_bnf Ds allAs allBs') Dss bnfs;
-      val ctor_fold_transfer_thms =
-        mk_un_fold_transfer_thms Least_FP rels activephis Irels Iphis
-          (mk_folds passiveAs activeAs) (mk_folds passiveBs activeBs)
-          (mk_fold_transfer_tac m Irel_induct_thm (map map_transfer_of_bnf bnfs) ctor_fold_thms)
-          lthy;
-
-      val timer = time (timer "relator induction");
-
-      val common_notes =
-        [(ctor_inductN, [ctor_induct_thm]),
-        (ctor_induct2N, [ctor_induct2_thm]),
-        (rel_inductN, [Irel_induct_thm]),
-        (ctor_fold_transferN, ctor_fold_transfer_thms)]
-        |> map (fn (thmN, thms) =>
-          ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]));
-
-      val notes =
-        [(ctor_dtorN, ctor_dtor_thms),
-        (ctor_exhaustN, ctor_exhaust_thms),
-        (ctor_foldN, ctor_fold_thms),
-        (ctor_fold_uniqueN, ctor_fold_unique_thms),
-        (ctor_rec_uniqueN, ctor_rec_unique_thms),
-        (ctor_fold_o_mapN, ctor_fold_o_Imap_thms),
-        (ctor_rec_o_mapN, ctor_rec_o_Imap_thms),
-        (ctor_injectN, ctor_inject_thms),
-        (ctor_recN, ctor_rec_thms),
-        (dtor_ctorN, dtor_ctor_thms),
-        (dtor_exhaustN, dtor_exhaust_thms),
-        (dtor_injectN, dtor_inject_thms)]
-        |> map (apsnd (map single))
-        |> maps (fn (thmN, thmss) =>
-          map2 (fn b => fn thms =>
-            ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]))
-          bs thmss);
-
-    (*FIXME: once the package exports all the necessary high-level characteristic theorems,
-       those should not only be concealed but rather not noted at all*)
-    val maybe_conceal_notes = note_all = false ? map (apfst (apfst Binding.conceal));
-  in
-    timer;
-    ({Ts = Ts, bnfs = Ibnfs, ctors = ctors, dtors = dtors, xtor_co_iterss = transpose [folds, recs],
-      xtor_co_induct = ctor_induct_thm, dtor_ctors = dtor_ctor_thms, ctor_dtors = ctor_dtor_thms,
-      ctor_injects = ctor_inject_thms, dtor_injects = dtor_inject_thms,
-      xtor_map_thms = ctor_Imap_thms, xtor_set_thmss = ctor_Iset_thmss',
-      xtor_rel_thms = ctor_Irel_thms,
-      xtor_co_iter_thmss = transpose [ctor_fold_thms, ctor_rec_thms],
-      xtor_co_iter_o_map_thmss = transpose [ctor_fold_o_Imap_thms, ctor_rec_o_Imap_thms],
-      rel_xtor_co_induct_thm = Irel_induct_thm},
-     lthy |> Local_Theory.notes (maybe_conceal_notes (common_notes @ notes @ Ibnf_notes)) |> snd)
-  end;
-
-val _ =
-  Outer_Syntax.local_theory @{command_spec "datatype_new"} "define new-style inductive datatypes"
-    (parse_co_datatype_cmd Least_FP construct_lfp);
-
-val _ = Outer_Syntax.local_theory @{command_spec "primrec_new"}
-  "define primitive recursive functions"
-  (Parse.fixes -- Parse_Spec.where_alt_specs >> (snd oo uncurry add_primrec_cmd));
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_lfp_compat.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,172 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_lfp_compat.ML
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2013
-
-Compatibility layer with the old datatype package.
-*)
-
-signature BNF_LFP_COMPAT =
-sig
-  val datatype_new_compat_cmd : string list -> local_theory -> local_theory
-end;
-
-structure BNF_LFP_Compat : BNF_LFP_COMPAT =
-struct
-
-open Ctr_Sugar
-open BNF_Util
-open BNF_FP_Util
-open BNF_FP_Def_Sugar
-open BNF_FP_N2M_Sugar
-
-fun dtyp_of_typ _ (TFree a) = Datatype_Aux.DtTFree a
-  | dtyp_of_typ recTs (T as Type (s, Ts)) =
-    (case find_index (curry (op =) T) recTs of
-      ~1 => Datatype_Aux.DtType (s, map (dtyp_of_typ recTs) Ts)
-    | kk => Datatype_Aux.DtRec kk);
-
-val compatN = "compat_";
-
-(* TODO: graceful failure for local datatypes -- perhaps by making the command global *)
-fun datatype_new_compat_cmd raw_fpT_names lthy =
-  let
-    val thy = Proof_Context.theory_of lthy;
-
-    fun not_datatype s = error (quote s ^ " is not a new-style datatype");
-    fun not_mutually_recursive ss =
-      error ("{" ^ commas ss ^ "} is not a complete set of mutually recursive new-style datatypes");
-
-    val (fpT_names as fpT_name1 :: _) =
-      map (fst o dest_Type o Proof_Context.read_type_name_proper lthy false) raw_fpT_names;
-
-    fun lfp_sugar_of s =
-      (case fp_sugar_of lthy s of
-        SOME (fp_sugar as {fp = Least_FP, ...}) => fp_sugar
-      | _ => not_datatype s);
-
-    val {ctr_sugars, ...} = lfp_sugar_of fpT_name1;
-    val fpTs0 as Type (_, var_As) :: _ = map (body_type o fastype_of o hd o #ctrs) ctr_sugars;
-    val fpT_names' = map (fst o dest_Type) fpTs0;
-
-    val _ = eq_set (op =) (fpT_names, fpT_names') orelse not_mutually_recursive fpT_names;
-
-    val (unsorted_As, _) = lthy |> mk_TFrees (length var_As);
-    val As = map2 (resort_tfree o Type.sort_of_atyp) var_As unsorted_As;
-    val fpTs as fpT1 :: _ = map (fn s => Type (s, As)) fpT_names';
-
-    fun add_nested_types_of (T as Type (s, _)) seen =
-      if member (op =) seen T then
-        seen
-      else if s = @{type_name fun} then
-        (warning "Partial support for recursion through functions -- 'primrec' will fail"; seen)
-      else
-        (case try lfp_sugar_of s of
-          SOME ({T = T0, fp_res = {Ts = mutual_Ts0, ...}, ctr_sugars, ...}) =>
-          let
-            val rho = Vartab.fold (cons o apsnd snd) (Sign.typ_match thy (T0, T) Vartab.empty) [];
-            val substT = Term.typ_subst_TVars rho;
-
-            val mutual_Ts = map substT mutual_Ts0;
-
-            fun add_interesting_subtypes (U as Type (_, Us)) =
-                (case filter (exists_subtype_in mutual_Ts) Us of [] => I
-                | Us' => insert (op =) U #> fold add_interesting_subtypes Us')
-              | add_interesting_subtypes _ = I;
-
-            val ctrs = maps #ctrs ctr_sugars;
-            val ctr_Ts = maps (binder_types o substT o fastype_of) ctrs |> distinct (op =);
-            val subTs = fold add_interesting_subtypes ctr_Ts [];
-          in
-            fold add_nested_types_of subTs (seen @ mutual_Ts)
-          end
-        | NONE => error ("Unsupported recursion via type constructor " ^ quote s ^
-            " not corresponding to new-style datatype (cf. \"datatype_new\")"));
-
-    val Ts = add_nested_types_of fpT1 [];
-    val b_names = map base_name_of_typ Ts;
-    val compat_b_names = map (prefix compatN) b_names;
-    val compat_bs = map Binding.name compat_b_names;
-    val common_name = compatN ^ mk_common_name b_names;
-    val nn_fp = length fpTs;
-    val nn = length Ts;
-    val get_indices = K [];
-    val fp_sugars0 = map (lfp_sugar_of o fst o dest_Type) Ts;
-    val callssss = map (fn fp_sugar0 => indexify_callsss fp_sugar0 []) fp_sugars0;
-
-    val ((fp_sugars, (lfp_sugar_thms, _)), lthy) =
-      if nn > nn_fp then
-        mutualize_fp_sugars Least_FP compat_bs Ts get_indices callssss fp_sugars0 lthy
-      else
-        ((fp_sugars0, (NONE, NONE)), lthy);
-
-    val {ctr_sugars, co_inducts = [induct], co_iterss, co_iter_thmsss = iter_thmsss, ...} :: _ =
-      fp_sugars;
-    val inducts = conj_dests nn induct;
-
-    val mk_dtyp = dtyp_of_typ Ts;
-
-    fun mk_ctr_descr Ts = mk_ctr Ts #> dest_Const ##> (binder_types #> map mk_dtyp);
-    fun mk_typ_descr index (Type (T_name, Ts)) ({ctrs, ...} : ctr_sugar) =
-      (index, (T_name, map mk_dtyp Ts, map (mk_ctr_descr Ts) ctrs));
-
-    val descr = map3 mk_typ_descr (0 upto nn - 1) Ts ctr_sugars;
-    val recs = map (fst o dest_Const o co_rec_of) co_iterss;
-    val rec_thms = flat (map co_rec_of iter_thmsss);
-
-    fun mk_info ({T = Type (T_name0, _), index, ...} : fp_sugar) =
-      let
-        val {casex, exhaust, nchotomy, injects, distincts, case_thms, case_cong, weak_case_cong,
-          split, split_asm, ...} = nth ctr_sugars index;
-      in
-        (T_name0,
-         {index = index, descr = descr, inject = injects, distinct = distincts, induct = induct,
-         inducts = inducts, exhaust = exhaust, nchotomy = nchotomy, rec_names = recs,
-         rec_rewrites = rec_thms, case_name = fst (dest_Const casex), case_rewrites = case_thms,
-         case_cong = case_cong, weak_case_cong = weak_case_cong, split = split,
-         split_asm = split_asm})
-      end;
-
-    val infos = map mk_info (take nn_fp fp_sugars);
-
-    val all_notes =
-      (case lfp_sugar_thms of
-        NONE => []
-      | SOME ((induct_thms, induct_thm, induct_attrs), (fold_thmss, rec_thmss, _)) =>
-        let
-          val common_notes =
-            (if nn > 1 then [(inductN, [induct_thm], induct_attrs)] else [])
-            |> filter_out (null o #2)
-            |> map (fn (thmN, thms, attrs) =>
-              ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
-
-          val notes =
-            [(foldN, fold_thmss, []),
-             (inductN, map single induct_thms, induct_attrs),
-             (recN, rec_thmss, [])]
-            |> filter_out (null o #2)
-            |> maps (fn (thmN, thmss, attrs) =>
-              if forall null thmss then
-                []
-              else
-                map2 (fn b_name => fn thms =>
-                    ((Binding.qualify true b_name (Binding.name thmN), attrs), [(thms, [])]))
-                  compat_b_names thmss);
-        in
-          common_notes @ notes
-        end);
-
-    val register_interpret =
-      Datatype_Data.register infos
-      #> Datatype_Data.interpretation_data (Datatype_Aux.default_config, map fst infos)
-  in
-    lthy
-    |> Local_Theory.raw_theory register_interpret
-    |> Local_Theory.notes all_notes |> snd
-  end;
-
-val _ =
-  Outer_Syntax.local_theory @{command_spec "datatype_new_compat"}
-    "register new-style datatypes as old-style datatypes"
-    (Scan.repeat1 Parse.type_const >> datatype_new_compat_cmd);
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_lfp_rec_sugar.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,606 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_lfp_rec_sugar.ML
-    Author:     Lorenz Panny, TU Muenchen
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2013
-
-Recursor sugar.
-*)
-
-signature BNF_LFP_REC_SUGAR =
-sig
-  val add_primrec: (binding * typ option * mixfix) list ->
-    (Attrib.binding * term) list -> local_theory -> (term list * thm list list) * local_theory
-  val add_primrec_cmd: (binding * string option * mixfix) list ->
-    (Attrib.binding * string) list -> local_theory -> (term list * thm list list) * local_theory
-  val add_primrec_global: (binding * typ option * mixfix) list ->
-    (Attrib.binding * term) list -> theory -> (term list * thm list list) * theory
-  val add_primrec_overloaded: (string * (string * typ) * bool) list ->
-    (binding * typ option * mixfix) list ->
-    (Attrib.binding * term) list -> theory -> (term list * thm list list) * theory
-  val add_primrec_simple: ((binding * typ) * mixfix) list -> term list ->
-    local_theory -> (string list * (term list * (int list list * thm list list))) * local_theory
-end;
-
-structure BNF_LFP_Rec_Sugar : BNF_LFP_REC_SUGAR =
-struct
-
-open Ctr_Sugar
-open BNF_Util
-open BNF_Tactics
-open BNF_Def
-open BNF_FP_Util
-open BNF_FP_Def_Sugar
-open BNF_FP_N2M_Sugar
-open BNF_FP_Rec_Sugar_Util
-
-val nitpicksimp_attrs = @{attributes [nitpick_simp]};
-val simp_attrs = @{attributes [simp]};
-val code_nitpicksimp_simp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs @ simp_attrs;
-
-exception Primrec_Error of string * term list;
-
-fun primrec_error str = raise Primrec_Error (str, []);
-fun primrec_error_eqn str eqn = raise Primrec_Error (str, [eqn]);
-fun primrec_error_eqns str eqns = raise Primrec_Error (str, eqns);
-
-datatype rec_call =
-  No_Rec of int * typ |
-  Mutual_Rec of (int * typ) * (int * typ) |
-  Nested_Rec of int * typ;
-
-type rec_ctr_spec =
-  {ctr: term,
-   offset: int,
-   calls: rec_call list,
-   rec_thm: thm};
-
-type rec_spec =
-  {recx: term,
-   nested_map_idents: thm list,
-   nested_map_comps: thm list,
-   ctr_specs: rec_ctr_spec list};
-
-exception AINT_NO_MAP of term;
-
-fun ill_formed_rec_call ctxt t =
-  error ("Ill-formed recursive call: " ^ quote (Syntax.string_of_term ctxt t));
-fun invalid_map ctxt t =
-  error ("Invalid map function in " ^ quote (Syntax.string_of_term ctxt t));
-fun unexpected_rec_call ctxt t =
-  error ("Unexpected recursive call: " ^ quote (Syntax.string_of_term ctxt t));
-
-fun massage_nested_rec_call ctxt has_call raw_massage_fun bound_Ts y y' =
-  let
-    fun check_no_call t = if has_call t then unexpected_rec_call ctxt t else ();
-
-    val typof = curry fastype_of1 bound_Ts;
-    val build_map_fst = build_map ctxt (fst_const o fst);
-
-    val yT = typof y;
-    val yU = typof y';
-
-    fun y_of_y' () = build_map_fst (yU, yT) $ y';
-    val elim_y = Term.map_aterms (fn t => if t = y then y_of_y' () else t);
-
-    fun massage_mutual_fun U T t =
-      (case t of
-        Const (@{const_name comp}, _) $ t1 $ t2 =>
-        mk_comp bound_Ts (tap check_no_call t1, massage_mutual_fun U T t2)
-      | _ =>
-        if has_call t then
-          (case try HOLogic.dest_prodT U of
-            SOME (U1, U2) => if U1 = T then raw_massage_fun T U2 t else invalid_map ctxt t
-          | NONE => invalid_map ctxt t)
-        else
-          mk_comp bound_Ts (t, build_map_fst (U, T)));
-
-    fun massage_map (Type (_, Us)) (Type (s, Ts)) t =
-        (case try (dest_map ctxt s) t of
-          SOME (map0, fs) =>
-          let
-            val Type (_, ran_Ts) = range_type (typof t);
-            val map' = mk_map (length fs) Us ran_Ts map0;
-            val fs' = map_flattened_map_args ctxt s (map3 massage_map_or_map_arg Us Ts) fs;
-          in
-            Term.list_comb (map', fs')
-          end
-        | NONE => raise AINT_NO_MAP t)
-      | massage_map _ _ t = raise AINT_NO_MAP t
-    and massage_map_or_map_arg U T t =
-      if T = U then
-        tap check_no_call t
-      else
-        massage_map U T t
-        handle AINT_NO_MAP _ => massage_mutual_fun U T t;
-
-    fun massage_call (t as t1 $ t2) =
-        if has_call t then
-          if t2 = y then
-            massage_map yU yT (elim_y t1) $ y'
-            handle AINT_NO_MAP t' => invalid_map ctxt t'
-          else
-            let val (g, xs) = Term.strip_comb t2 in
-              if g = y then
-                if exists has_call xs then unexpected_rec_call ctxt t2
-                else Term.list_comb (massage_call (mk_compN (length xs) bound_Ts (t1, y)), xs)
-              else
-                ill_formed_rec_call ctxt t
-            end
-        else
-          elim_y t
-      | massage_call t = if t = y then y_of_y' () else ill_formed_rec_call ctxt t;
-  in
-    massage_call
-  end;
-
-fun rec_specs_of bs arg_Ts res_Ts get_indices callssss0 lthy0 =
-  let
-    val thy = Proof_Context.theory_of lthy0;
-
-    val ((missing_arg_Ts, perm0_kks,
-          fp_sugars as {nested_bnfs, fp_res = {xtor_co_iterss = ctor_iters1 :: _, ...},
-            co_inducts = [induct_thm], ...} :: _, (lfp_sugar_thms, _)), lthy) =
-      nested_to_mutual_fps Least_FP bs arg_Ts get_indices callssss0 lthy0;
-
-    val perm_fp_sugars = sort (int_ord o pairself #index) fp_sugars;
-
-    val indices = map #index fp_sugars;
-    val perm_indices = map #index perm_fp_sugars;
-
-    val perm_ctrss = map (#ctrs o of_fp_sugar #ctr_sugars) perm_fp_sugars;
-    val perm_ctr_Tsss = map (map (binder_types o fastype_of)) perm_ctrss;
-    val perm_lfpTs = map (body_type o fastype_of o hd) perm_ctrss;
-
-    val nn0 = length arg_Ts;
-    val nn = length perm_lfpTs;
-    val kks = 0 upto nn - 1;
-    val perm_ns = map length perm_ctr_Tsss;
-    val perm_mss = map (map length) perm_ctr_Tsss;
-
-    val perm_Cs = map (body_type o fastype_of o co_rec_of o of_fp_sugar (#xtor_co_iterss o #fp_res))
-      perm_fp_sugars;
-    val perm_fun_arg_Tssss =
-      mk_iter_fun_arg_types perm_ctr_Tsss perm_ns perm_mss (co_rec_of ctor_iters1);
-
-    fun unpermute0 perm0_xs = permute_like (op =) perm0_kks kks perm0_xs;
-    fun unpermute perm_xs = permute_like (op =) perm_indices indices perm_xs;
-
-    val induct_thms = unpermute0 (conj_dests nn induct_thm);
-
-    val lfpTs = unpermute perm_lfpTs;
-    val Cs = unpermute perm_Cs;
-
-    val As_rho = tvar_subst thy (take nn0 lfpTs) arg_Ts;
-    val Cs_rho = map (fst o dest_TVar) Cs ~~ pad_list HOLogic.unitT nn res_Ts;
-
-    val substA = Term.subst_TVars As_rho;
-    val substAT = Term.typ_subst_TVars As_rho;
-    val substCT = Term.typ_subst_TVars Cs_rho;
-    val substACT = substAT o substCT;
-
-    val perm_Cs' = map substCT perm_Cs;
-
-    fun offset_of_ctr 0 _ = 0
-      | offset_of_ctr n (({ctrs, ...} : ctr_sugar) :: ctr_sugars) =
-        length ctrs + offset_of_ctr (n - 1) ctr_sugars;
-
-    fun call_of [i] [T] = (if exists_subtype_in Cs T then Nested_Rec else No_Rec) (i, substACT T)
-      | call_of [i, i'] [T, T'] = Mutual_Rec ((i, substACT T), (i', substACT T'));
-
-    fun mk_ctr_spec ctr offset fun_arg_Tss rec_thm =
-      let
-        val (fun_arg_hss, _) = indexedd fun_arg_Tss 0;
-        val fun_arg_hs = flat_rec_arg_args fun_arg_hss;
-        val fun_arg_iss = map (map (find_index_eq fun_arg_hs)) fun_arg_hss;
-      in
-        {ctr = substA ctr, offset = offset, calls = map2 call_of fun_arg_iss fun_arg_Tss,
-         rec_thm = rec_thm}
-      end;
-
-    fun mk_ctr_specs index (ctr_sugars : ctr_sugar list) iter_thmsss =
-      let
-        val ctrs = #ctrs (nth ctr_sugars index);
-        val rec_thms = co_rec_of (nth iter_thmsss index);
-        val k = offset_of_ctr index ctr_sugars;
-        val n = length ctrs;
-      in
-        map4 mk_ctr_spec ctrs (k upto k + n - 1) (nth perm_fun_arg_Tssss index) rec_thms
-      end;
-
-    fun mk_spec ({T, index, ctr_sugars, co_iterss = iterss, co_iter_thmsss = iter_thmsss, ...}
-      : fp_sugar) =
-      {recx = mk_co_iter thy Least_FP (substAT T) perm_Cs' (co_rec_of (nth iterss index)),
-       nested_map_idents = map (unfold_thms lthy @{thms id_def} o map_id0_of_bnf) nested_bnfs,
-       nested_map_comps = map map_comp_of_bnf nested_bnfs,
-       ctr_specs = mk_ctr_specs index ctr_sugars iter_thmsss};
-  in
-    ((is_some lfp_sugar_thms, map mk_spec fp_sugars, missing_arg_Ts, induct_thm, induct_thms), lthy)
-  end;
-
-val undef_const = Const (@{const_name undefined}, dummyT);
-
-fun permute_args n t =
-  list_comb (t, map Bound (0 :: (n downto 1))) |> fold (K (Term.abs (Name.uu, dummyT))) (0 upto n);
-
-type eqn_data = {
-  fun_name: string,
-  rec_type: typ,
-  ctr: term,
-  ctr_args: term list,
-  left_args: term list,
-  right_args: term list,
-  res_type: typ,
-  rhs_term: term,
-  user_eqn: term
-};
-
-fun dissect_eqn lthy fun_names eqn' =
-  let
-    val eqn = drop_all eqn' |> HOLogic.dest_Trueprop
-      handle TERM _ =>
-        primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
-    val (lhs, rhs) = HOLogic.dest_eq eqn
-        handle TERM _ =>
-          primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
-    val (fun_name, args) = strip_comb lhs
-      |>> (fn x => if is_Free x then fst (dest_Free x)
-          else primrec_error_eqn "malformed function equation (does not start with free)" eqn);
-    val (left_args, rest) = take_prefix is_Free args;
-    val (nonfrees, right_args) = take_suffix is_Free rest;
-    val num_nonfrees = length nonfrees;
-    val _ = num_nonfrees = 1 orelse if num_nonfrees = 0 then
-      primrec_error_eqn "constructor pattern missing in left-hand side" eqn else
-      primrec_error_eqn "more than one non-variable argument in left-hand side" eqn;
-    val _ = member (op =) fun_names fun_name orelse
-      primrec_error_eqn "malformed function equation (does not start with function name)" eqn
-
-    val (ctr, ctr_args) = strip_comb (the_single nonfrees);
-    val _ = try (num_binder_types o fastype_of) ctr = SOME (length ctr_args) orelse
-      primrec_error_eqn "partially applied constructor in pattern" eqn;
-    val _ = let val d = duplicates (op =) (left_args @ ctr_args @ right_args) in null d orelse
-      primrec_error_eqn ("duplicate variable \"" ^ Syntax.string_of_term lthy (hd d) ^
-        "\" in left-hand side") eqn end;
-    val _ = forall is_Free ctr_args orelse
-      primrec_error_eqn "non-primitive pattern in left-hand side" eqn;
-    val _ =
-      let val b = fold_aterms (fn x as Free (v, _) =>
-        if (not (member (op =) (left_args @ ctr_args @ right_args) x) andalso
-        not (member (op =) fun_names v) andalso
-        not (Variable.is_fixed lthy v)) then cons x else I | _ => I) rhs []
-      in
-        null b orelse
-        primrec_error_eqn ("extra variable(s) in right-hand side: " ^
-          commas (map (Syntax.string_of_term lthy) b)) eqn
-      end;
-  in
-    {fun_name = fun_name,
-     rec_type = body_type (type_of ctr),
-     ctr = ctr,
-     ctr_args = ctr_args,
-     left_args = left_args,
-     right_args = right_args,
-     res_type = map fastype_of (left_args @ right_args) ---> fastype_of rhs,
-     rhs_term = rhs,
-     user_eqn = eqn'}
-  end;
-
-fun rewrite_map_arg get_ctr_pos rec_type res_type =
-  let
-    val pT = HOLogic.mk_prodT (rec_type, res_type);
-
-    fun subst d (t as Bound d') = t |> d = SOME d' ? curry (op $) (fst_const pT)
-      | subst d (Abs (v, T, b)) =
-        Abs (v, if d = SOME ~1 then pT else T, subst (Option.map (Integer.add 1) d) b)
-      | subst d t =
-        let
-          val (u, vs) = strip_comb t;
-          val ctr_pos = try (get_ctr_pos o fst o dest_Free) u |> the_default ~1;
-        in
-          if ctr_pos >= 0 then
-            if d = SOME ~1 andalso length vs = ctr_pos then
-              list_comb (permute_args ctr_pos (snd_const pT), vs)
-            else if length vs > ctr_pos andalso is_some d
-                andalso d = try (fn Bound n => n) (nth vs ctr_pos) then
-              list_comb (snd_const pT $ nth vs ctr_pos, map (subst d) (nth_drop ctr_pos vs))
-            else
-              primrec_error_eqn ("recursive call not directly applied to constructor argument") t
-          else
-            list_comb (u, map (subst (d |> d = SOME ~1 ? K NONE)) vs)
-        end
-  in
-    subst (SOME ~1)
-  end;
-
-fun subst_rec_calls lthy get_ctr_pos has_call ctr_args mutual_calls nested_calls =
-  let
-    fun try_nested_rec bound_Ts y t =
-      AList.lookup (op =) nested_calls y
-      |> Option.map (fn y' =>
-        massage_nested_rec_call lthy has_call (rewrite_map_arg get_ctr_pos) bound_Ts y y' t);
-
-    fun subst bound_Ts (t as g' $ y) =
-        let
-          fun subst_rec () = subst bound_Ts g' $ subst bound_Ts y;
-          val y_head = head_of y;
-        in
-          if not (member (op =) ctr_args y_head) then
-            subst_rec ()
-          else
-            (case try_nested_rec bound_Ts y_head t of
-              SOME t' => t'
-            | NONE =>
-              let val (g, g_args) = strip_comb g' in
-                (case try (get_ctr_pos o fst o dest_Free) g of
-                  SOME ctr_pos =>
-                  (length g_args >= ctr_pos orelse
-                   primrec_error_eqn "too few arguments in recursive call" t;
-                   (case AList.lookup (op =) mutual_calls y of
-                     SOME y' => list_comb (y', g_args)
-                   | NONE => subst_rec ()))
-                | NONE => subst_rec ())
-              end)
-        end
-      | subst bound_Ts (Abs (v, T, b)) = Abs (v, T, subst (T :: bound_Ts) b)
-      | subst _ t = t
-
-    fun subst' t =
-      if has_call t then
-        (* FIXME detect this case earlier? *)
-        primrec_error_eqn "recursive call not directly applied to constructor argument" t
-      else
-        try_nested_rec [] (head_of t) t |> the_default t
-  in
-    subst' o subst []
-  end;
-
-fun build_rec_arg lthy (funs_data : eqn_data list list) has_call (ctr_spec : rec_ctr_spec)
-    (eqn_data_opt : eqn_data option) =
-  (case eqn_data_opt of
-    NONE => undef_const
-  | SOME {ctr_args, left_args, right_args, rhs_term = t, ...} =>
-    let
-      val calls = #calls ctr_spec;
-      val n_args = fold (Integer.add o (fn Mutual_Rec _ => 2 | _ => 1)) calls 0;
-
-      val no_calls' = tag_list 0 calls
-        |> map_filter (try (apsnd (fn No_Rec p => p | Mutual_Rec (p, _) => p)));
-      val mutual_calls' = tag_list 0 calls
-        |> map_filter (try (apsnd (fn Mutual_Rec (_, p) => p)));
-      val nested_calls' = tag_list 0 calls
-        |> map_filter (try (apsnd (fn Nested_Rec p => p)));
-
-      fun ensure_unique frees t =
-        if member (op =) frees t then Free (the_single (Term.variant_frees t [dest_Free t])) else t;
-
-      val args = replicate n_args ("", dummyT)
-        |> Term.rename_wrt_term t
-        |> map Free
-        |> fold (fn (ctr_arg_idx, (arg_idx, _)) =>
-            nth_map arg_idx (K (nth ctr_args ctr_arg_idx)))
-          no_calls'
-        |> fold (fn (ctr_arg_idx, (arg_idx, T)) => fn xs =>
-            nth_map arg_idx (K (ensure_unique xs (retype_free T (nth ctr_args ctr_arg_idx)))) xs)
-          mutual_calls'
-        |> fold (fn (ctr_arg_idx, (arg_idx, T)) =>
-            nth_map arg_idx (K (retype_free T (nth ctr_args ctr_arg_idx))))
-          nested_calls';
-
-      val fun_name_ctr_pos_list =
-        map (fn (x :: _) => (#fun_name x, length (#left_args x))) funs_data;
-      val get_ctr_pos = try (the o AList.lookup (op =) fun_name_ctr_pos_list) #> the_default ~1;
-      val mutual_calls = map (apfst (nth ctr_args) o apsnd (nth args o fst)) mutual_calls';
-      val nested_calls = map (apfst (nth ctr_args) o apsnd (nth args o fst)) nested_calls';
-    in
-      t
-      |> subst_rec_calls lthy get_ctr_pos has_call ctr_args mutual_calls nested_calls
-      |> fold_rev lambda (args @ left_args @ right_args)
-    end);
-
-fun build_defs lthy bs mxs (funs_data : eqn_data list list) (rec_specs : rec_spec list) has_call =
-  let
-    val n_funs = length funs_data;
-
-    val ctr_spec_eqn_data_list' =
-      (take n_funs rec_specs |> map #ctr_specs) ~~ funs_data
-      |> maps (uncurry (finds (fn (x, y) => #ctr x = #ctr y))
-          ##> (fn x => null x orelse
-            primrec_error_eqns "excess equations in definition" (map #rhs_term x)) #> fst);
-    val _ = ctr_spec_eqn_data_list' |> map (fn (_, x) => length x <= 1 orelse
-      primrec_error_eqns ("multiple equations for constructor") (map #user_eqn x));
-
-    val ctr_spec_eqn_data_list =
-      ctr_spec_eqn_data_list' @ (drop n_funs rec_specs |> maps #ctr_specs |> map (rpair []));
-
-    val recs = take n_funs rec_specs |> map #recx;
-    val rec_args = ctr_spec_eqn_data_list
-      |> sort ((op <) o pairself (#offset o fst) |> make_ord)
-      |> map (uncurry (build_rec_arg lthy funs_data has_call) o apsnd (try the_single));
-    val ctr_poss = map (fn x =>
-      if length (distinct ((op =) o pairself (length o #left_args)) x) <> 1 then
-        primrec_error ("inconstant constructor pattern position for function " ^
-          quote (#fun_name (hd x)))
-      else
-        hd x |> #left_args |> length) funs_data;
-  in
-    (recs, ctr_poss)
-    |-> map2 (fn recx => fn ctr_pos => list_comb (recx, rec_args) |> permute_args ctr_pos)
-    |> Syntax.check_terms lthy
-    |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.conceal (Thm.def_binding b), []), t)))
-      bs mxs
-  end;
-
-fun find_rec_calls has_call ({ctr, ctr_args, rhs_term, ...} : eqn_data) =
-  let
-    fun find bound_Ts (Abs (_, T, b)) ctr_arg = find (T :: bound_Ts) b ctr_arg
-      | find bound_Ts (t as _ $ _) ctr_arg =
-        let
-          val typof = curry fastype_of1 bound_Ts;
-          val (f', args') = strip_comb t;
-          val n = find_index (equal ctr_arg o head_of) args';
-        in
-          if n < 0 then
-            find bound_Ts f' ctr_arg @ maps (fn x => find bound_Ts x ctr_arg) args'
-          else
-            let
-              val (f, args as arg :: _) = chop n args' |>> curry list_comb f'
-              val (arg_head, arg_args) = Term.strip_comb arg;
-            in
-              if has_call f then
-                mk_partial_compN (length arg_args) (typof arg_head) f ::
-                maps (fn x => find bound_Ts x ctr_arg) args
-              else
-                find bound_Ts f ctr_arg @ maps (fn x => find bound_Ts x ctr_arg) args
-            end
-        end
-      | find _ _ _ = [];
-  in
-    map (find [] rhs_term) ctr_args
-    |> (fn [] => NONE | callss => SOME (ctr, callss))
-  end;
-
-fun mk_primrec_tac ctxt num_extra_args map_idents map_comps fun_defs recx =
-  unfold_thms_tac ctxt fun_defs THEN
-  HEADGOAL (rtac (funpow num_extra_args (fn thm => thm RS fun_cong) recx RS trans)) THEN
-  unfold_thms_tac ctxt (@{thms id_def split o_def fst_conv snd_conv} @ map_comps @ map_idents) THEN
-  HEADGOAL (rtac refl);
-
-fun prepare_primrec fixes specs lthy =
-  let
-    val thy = Proof_Context.theory_of lthy;
-
-    val (bs, mxs) = map_split (apfst fst) fixes;
-    val fun_names = map Binding.name_of bs;
-    val eqns_data = map (dissect_eqn lthy fun_names) specs;
-    val funs_data = eqns_data
-      |> partition_eq ((op =) o pairself #fun_name)
-      |> finds (fn (x, y) => x = #fun_name (hd y)) fun_names |> fst
-      |> map (fn (x, y) => the_single y handle List.Empty =>
-          primrec_error ("missing equations for function " ^ quote x));
-
-    val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
-    val arg_Ts = map (#rec_type o hd) funs_data;
-    val res_Ts = map (#res_type o hd) funs_data;
-    val callssss = funs_data
-      |> map (partition_eq ((op =) o pairself #ctr))
-      |> map (maps (map_filter (find_rec_calls has_call)));
-
-    val _ = (case filter_out (fn (_, T) => Sign.of_sort thy (T, HOLogic.typeS)) (bs ~~ res_Ts) of
-        [] => ()
-      | (b, _) :: _ => primrec_error ("type of " ^ Binding.print b ^ " contains top sort"));
-
-    val ((n2m, rec_specs, _, induct_thm, induct_thms), lthy') =
-      rec_specs_of bs arg_Ts res_Ts (get_indices fixes) callssss lthy;
-
-    val actual_nn = length funs_data;
-
-    val _ = let val ctrs = (maps (map #ctr o #ctr_specs) rec_specs) in
-      map (fn {ctr, user_eqn, ...} => member (op =) ctrs ctr orelse
-        primrec_error_eqn ("argument " ^ quote (Syntax.string_of_term lthy' ctr) ^
-          " is not a constructor in left-hand side") user_eqn) eqns_data end;
-
-    val defs = build_defs lthy' bs mxs funs_data rec_specs has_call;
-
-    fun prove lthy def_thms' ({ctr_specs, nested_map_idents, nested_map_comps, ...} : rec_spec)
-        (fun_data : eqn_data list) =
-      let
-        val def_thms = map (snd o snd) def_thms';
-        val simp_thmss = finds (fn (x, y) => #ctr x = #ctr y) fun_data ctr_specs
-          |> fst
-          |> map_filter (try (fn (x, [y]) =>
-            (#user_eqn x, length (#left_args x) + length (#right_args x), #rec_thm y)))
-          |> map (fn (user_eqn, num_extra_args, rec_thm) =>
-            mk_primrec_tac lthy num_extra_args nested_map_idents nested_map_comps def_thms rec_thm
-            |> K |> Goal.prove_sorry lthy [] [] user_eqn
-            |> Thm.close_derivation);
-        val poss = find_indices (op = o pairself #ctr) fun_data eqns_data;
-      in
-        (poss, simp_thmss)
-      end;
-
-    val notes =
-      (if n2m then map2 (fn name => fn thm =>
-        (name, inductN, [thm], [])) fun_names (take actual_nn induct_thms) else [])
-      |> map (fn (prefix, thmN, thms, attrs) =>
-        ((Binding.qualify true prefix (Binding.name thmN), attrs), [(thms, [])]));
-
-    val common_name = mk_common_name fun_names;
-
-    val common_notes =
-      (if n2m then [(inductN, [induct_thm], [])] else [])
-      |> map (fn (thmN, thms, attrs) =>
-        ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
-  in
-    (((fun_names, defs),
-      fn lthy => fn defs =>
-        split_list (map2 (prove lthy defs) (take actual_nn rec_specs) funs_data)),
-      lthy' |> Local_Theory.notes (notes @ common_notes) |> snd)
-  end;
-
-(* primrec definition *)
-
-fun add_primrec_simple fixes ts lthy =
-  let
-    val (((names, defs), prove), lthy) = prepare_primrec fixes ts lthy
-      handle ERROR str => primrec_error str;
-  in
-    lthy
-    |> fold_map Local_Theory.define defs
-    |-> (fn defs => `(fn lthy => (names, (map fst defs, prove lthy defs))))
-  end
-  handle Primrec_Error (str, eqns) =>
-    if null eqns
-    then error ("primrec_new error:\n  " ^ str)
-    else error ("primrec_new error:\n  " ^ str ^ "\nin\n  " ^
-      space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns));
-
-local
-
-fun gen_primrec prep_spec (raw_fixes : (binding * 'a option * mixfix) list) raw_spec lthy =
-  let
-    val d = duplicates (op =) (map (Binding.name_of o #1) raw_fixes)
-    val _ = null d orelse primrec_error ("duplicate function name(s): " ^ commas d);
-
-    val (fixes, specs) = fst (prep_spec raw_fixes raw_spec lthy);
-
-    val mk_notes =
-      flat ooo map3 (fn poss => fn prefix => fn thms =>
-        let
-          val (bs, attrss) = map_split (fst o nth specs) poss;
-          val notes =
-            map3 (fn b => fn attrs => fn thm =>
-              ((Binding.qualify false prefix b, code_nitpicksimp_simp_attrs @ attrs), [([thm], [])]))
-            bs attrss thms;
-        in
-          ((Binding.qualify true prefix (Binding.name simpsN), []), [(thms, [])]) :: notes
-        end);
-  in
-    lthy
-    |> add_primrec_simple fixes (map snd specs)
-    |-> (fn (names, (ts, (posss, simpss))) =>
-      Spec_Rules.add Spec_Rules.Equational (ts, flat simpss)
-      #> Local_Theory.notes (mk_notes posss names simpss)
-      #>> pair ts o map snd)
-  end;
-
-in
-
-val add_primrec = gen_primrec Specification.check_spec;
-val add_primrec_cmd = gen_primrec Specification.read_spec;
-
-end;
-
-fun add_primrec_global fixes specs thy =
-  let
-    val lthy = Named_Target.theory_init thy;
-    val ((ts, simps), lthy') = add_primrec fixes specs lthy;
-    val simps' = burrow (Proof_Context.export lthy' lthy) simps;
-  in ((ts, simps'), Local_Theory.exit_global lthy') end;
-
-fun add_primrec_overloaded ops fixes specs thy =
-  let
-    val lthy = Overloading.overloading ops thy;
-    val ((ts, simps), lthy') = add_primrec fixes specs lthy;
-    val simps' = burrow (Proof_Context.export lthy' lthy) simps;
-  in ((ts, simps'), Local_Theory.exit_global lthy') end;
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_lfp_tactics.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,817 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_lfp_tactics.ML
-    Author:     Dmitriy Traytel, TU Muenchen
-    Author:     Andrei Popescu, TU Muenchen
-    Copyright   2012
-
-Tactics for the datatype construction.
-*)
-
-signature BNF_LFP_TACTICS =
-sig
-  val mk_alg_min_alg_tac: int -> thm -> thm list -> thm -> thm -> thm list list -> thm list ->
-    thm list -> tactic
-  val mk_alg_not_empty_tac: Proof.context -> thm -> thm list -> thm list -> tactic
-  val mk_alg_select_tac: thm -> {prems: 'a, context: Proof.context} -> tactic
-  val mk_alg_set_tac: thm -> tactic
-  val mk_bd_card_order_tac: thm list -> tactic
-  val mk_bd_limit_tac: int -> thm -> tactic
-  val mk_card_of_min_alg_tac: thm -> thm -> thm -> thm -> thm -> tactic
-  val mk_copy_alg_tac: thm list list -> thm list -> thm -> thm -> thm -> tactic
-  val mk_copy_str_tac: thm list list -> thm -> thm list -> tactic
-  val mk_ctor_induct_tac: Proof.context -> int -> thm list list -> thm -> thm list -> thm ->
-    thm list -> thm list -> thm list -> tactic
-  val mk_ctor_induct2_tac: ctyp option list -> cterm option list -> thm -> thm list ->
-    {prems: 'a, context: Proof.context} -> tactic
-  val mk_ctor_set_tac: thm -> thm -> thm list -> tactic
-  val mk_ctor_rel_tac: Proof.context -> thm list -> int -> thm -> thm -> thm -> thm -> thm list ->
-    thm -> thm -> thm list -> thm list -> thm list list -> tactic
-  val mk_dtor_o_ctor_tac: thm -> thm -> thm -> thm -> thm list -> tactic
-  val mk_ex_copy_alg_tac: int -> thm -> thm -> tactic
-  val mk_init_ex_mor_tac: thm -> thm -> thm -> thm list -> thm -> thm -> thm ->
-    {prems: 'a, context: Proof.context} -> tactic
-  val mk_init_induct_tac: int -> thm -> thm -> thm list -> thm list -> tactic
-  val mk_init_unique_mor_tac: int -> thm -> thm -> thm list -> thm list -> thm list -> thm list ->
-    thm list -> tactic
-  val mk_iso_alt_tac: thm list -> thm -> tactic
-  val mk_fold_unique_mor_tac: thm list -> thm list -> thm list -> thm -> thm -> thm -> tactic
-  val mk_fold_transfer_tac: int -> thm -> thm list -> thm list ->
-    {prems: thm list, context: Proof.context} -> tactic
-  val mk_least_min_alg_tac: thm -> thm -> tactic
-  val mk_le_rel_OO_tac: int -> thm -> thm list -> thm list -> thm list -> thm list ->
-    {prems: 'a, context: Proof.context} -> tactic
-  val mk_map_comp0_tac: thm list -> thm list -> thm -> int -> tactic
-  val mk_map_id0_tac: thm list -> thm -> tactic
-  val mk_map_tac: int -> int -> thm -> thm -> thm -> tactic
-  val mk_ctor_map_unique_tac: thm -> thm list -> Proof.context -> tactic
-  val mk_mcong_tac: (int -> tactic) -> thm list list list -> thm list -> thm list ->
-    {prems: 'a, context: Proof.context} -> tactic
-  val mk_min_algs_card_of_tac: ctyp -> cterm -> int -> thm -> thm list -> thm list -> thm -> thm ->
-    thm -> thm -> thm -> thm -> thm -> tactic
-  val mk_min_algs_least_tac: ctyp -> cterm -> thm -> thm list -> thm list -> tactic
-  val mk_min_algs_mono_tac: Proof.context -> thm -> tactic
-  val mk_min_algs_tac: thm -> thm list -> tactic
-  val mk_mor_Abs_tac: thm -> thm list -> thm list -> tactic
-  val mk_mor_Rep_tac: thm list -> thm -> thm list -> thm list -> thm list ->
-    {prems: 'a, context: Proof.context} -> tactic
-  val mk_mor_UNIV_tac: int -> thm list -> thm -> tactic
-  val mk_mor_comp_tac: thm -> thm list list -> thm list -> tactic
-  val mk_mor_convol_tac: 'a list -> thm -> tactic
-  val mk_mor_elim_tac: thm -> tactic
-  val mk_mor_incl_tac: thm -> thm list -> tactic
-  val mk_mor_inv_tac: thm -> thm -> thm list list -> thm list -> thm list -> thm list -> tactic
-  val mk_mor_fold_tac: ctyp -> cterm -> thm list -> thm -> thm -> tactic
-  val mk_mor_select_tac: thm -> thm -> thm -> thm -> thm -> thm -> thm list -> thm list list ->
-    thm list -> tactic
-  val mk_mor_str_tac: 'a list -> thm -> tactic
-  val mk_rel_induct_tac: int -> thm -> int list -> thm list -> thm list ->
-    {prems: thm list, context: Proof.context} -> tactic
-  val mk_rec_tac: thm list -> thm -> thm list -> {prems: 'a, context: Proof.context} -> tactic
-  val mk_rec_unique_mor_tac: thm list -> thm list -> thm -> {prems: 'a, context: Proof.context} ->
-    tactic
-  val mk_set_bd_tac: int -> (int -> tactic) -> thm -> thm list list -> thm list -> int ->
-    Proof.context -> tactic
-  val mk_set_nat_tac: int -> (int -> tactic) -> thm list list -> thm list -> cterm list ->
-    thm list -> int -> {prems: 'a, context: Proof.context} -> tactic
-  val mk_set_map0_tac: thm -> tactic
-  val mk_set_tac: thm -> tactic
-  val mk_wit_tac: Proof.context -> int -> thm list -> thm list -> tactic
-end;
-
-structure BNF_LFP_Tactics : BNF_LFP_TACTICS =
-struct
-
-open BNF_Tactics
-open BNF_LFP_Util
-open BNF_Util
-
-val fst_snd_convs = @{thms fst_conv snd_conv};
-val ord_eq_le_trans = @{thm ord_eq_le_trans};
-val subset_trans = @{thm subset_trans};
-val trans_fun_cong_image_id_id_apply = @{thm trans[OF fun_cong[OF image_id] id_apply]};
-val rev_bspec = Drule.rotate_prems 1 bspec;
-val Un_cong = @{thm arg_cong2[of _ _ _ _ "op \<union>"]}
-
-fun mk_alg_set_tac alg_def =
-  dtac (alg_def RS iffD1) 1 THEN
-  REPEAT_DETERM (etac conjE 1) THEN
-  EVERY' [etac bspec, rtac CollectI] 1 THEN
-  REPEAT_DETERM (etac conjI 1) THEN atac 1;
-
-fun mk_alg_not_empty_tac ctxt alg_set alg_sets wits =
-  (EVERY' [rtac notI, hyp_subst_tac ctxt, ftac alg_set] THEN'
-  REPEAT_DETERM o FIRST'
-    [rtac subset_UNIV,
-    EVERY' [rtac @{thm subset_emptyI}, eresolve_tac wits],
-    EVERY' [rtac subsetI, rtac FalseE, eresolve_tac wits],
-    EVERY' [rtac subsetI, dresolve_tac wits, hyp_subst_tac ctxt,
-      FIRST' (map (fn thm => rtac thm THEN' atac) alg_sets)]] THEN'
-  etac @{thm emptyE}) 1;
-
-fun mk_mor_elim_tac mor_def =
-  (dtac (subst OF [mor_def]) THEN'
-  REPEAT o etac conjE THEN'
-  TRY o rtac @{thm image_subsetI} THEN'
-  etac bspec THEN'
-  atac) 1;
-
-fun mk_mor_incl_tac mor_def map_ids =
-  (stac mor_def THEN'
-  rtac conjI THEN'
-  CONJ_WRAP' (K (EVERY' [rtac ballI, etac set_mp, stac id_apply, atac])) map_ids THEN'
-  CONJ_WRAP' (fn thm =>
-   (EVERY' [rtac ballI, rtac trans, rtac id_apply, stac thm, rtac refl])) map_ids) 1;
-
-fun mk_mor_comp_tac mor_def set_maps map_comp_ids =
-  let
-    val fbetw_tac = EVERY' [rtac ballI, stac o_apply, etac bspec, etac bspec, atac];
-    fun mor_tac (set_map, map_comp_id) =
-      EVERY' [rtac ballI, stac o_apply, rtac trans,
-        rtac trans, dtac rev_bspec, atac, etac arg_cong,
-         REPEAT o eresolve_tac [CollectE, conjE], etac bspec, rtac CollectI] THEN'
-      CONJ_WRAP' (fn thm =>
-        FIRST' [rtac subset_UNIV,
-          (EVERY' [rtac ord_eq_le_trans, rtac thm, rtac @{thm image_subsetI},
-            etac bspec, etac set_mp, atac])]) set_map THEN'
-      rtac (map_comp_id RS arg_cong);
-  in
-    (dtac (mor_def RS subst) THEN' dtac (mor_def RS subst) THEN' stac mor_def THEN'
-    REPEAT o etac conjE THEN'
-    rtac conjI THEN'
-    CONJ_WRAP' (K fbetw_tac) set_maps THEN'
-    CONJ_WRAP' mor_tac (set_maps ~~ map_comp_ids)) 1
-  end;
-
-fun mk_mor_inv_tac alg_def mor_def set_maps morEs map_comp_ids map_cong0Ls =
-  let
-    val fbetw_tac = EVERY' [rtac ballI, etac set_mp, etac imageI];
-    fun Collect_tac set_map =
-      CONJ_WRAP' (fn thm =>
-        FIRST' [rtac subset_UNIV,
-          (EVERY' [rtac ord_eq_le_trans, rtac thm, rtac subset_trans,
-            etac @{thm image_mono}, atac])]) set_map;
-    fun mor_tac (set_map, ((morE, map_comp_id), map_cong0L)) =
-      EVERY' [rtac ballI, ftac rev_bspec, atac,
-         REPEAT o eresolve_tac [CollectE, conjE], rtac sym, rtac trans, rtac sym,
-         etac @{thm inverE}, etac bspec, rtac CollectI, Collect_tac set_map,
-         rtac trans, etac (morE RS arg_cong), rtac CollectI, Collect_tac set_map,
-         rtac trans, rtac (map_comp_id RS arg_cong), rtac (map_cong0L RS arg_cong),
-         REPEAT_DETERM_N (length morEs) o
-           (EVERY' [rtac subst, rtac @{thm inver_pointfree}, etac @{thm inver_mono}, atac])];
-  in
-    (stac mor_def THEN'
-    dtac (alg_def RS iffD1) THEN'
-    dtac (alg_def RS iffD1) THEN'
-    REPEAT o etac conjE THEN'
-    rtac conjI THEN'
-    CONJ_WRAP' (K fbetw_tac) set_maps THEN'
-    CONJ_WRAP' mor_tac (set_maps ~~ (morEs ~~ map_comp_ids ~~ map_cong0Ls))) 1
-  end;
-
-fun mk_mor_str_tac ks mor_def =
-  (stac mor_def THEN' rtac conjI THEN'
-  CONJ_WRAP' (K (EVERY' [rtac ballI, rtac UNIV_I])) ks THEN'
-  CONJ_WRAP' (K (EVERY' [rtac ballI, rtac refl])) ks) 1;
-
-fun mk_mor_convol_tac ks mor_def =
-  (stac mor_def THEN' rtac conjI THEN'
-  CONJ_WRAP' (K (EVERY' [rtac ballI, rtac UNIV_I])) ks THEN'
-  CONJ_WRAP' (K (EVERY' [rtac ballI, rtac trans, rtac @{thm fst_convol'}, rtac o_apply])) ks) 1;
-
-fun mk_mor_UNIV_tac m morEs mor_def =
-  let
-    val n = length morEs;
-    fun mor_tac morE = EVERY' [rtac ext, rtac trans, rtac o_apply, rtac trans, etac morE,
-      rtac CollectI, CONJ_WRAP' (K (rtac subset_UNIV)) (1 upto m + n),
-      rtac sym, rtac o_apply];
-  in
-    EVERY' [rtac iffI, CONJ_WRAP' mor_tac morEs,
-    stac mor_def, rtac conjI, CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) morEs,
-    REPEAT_DETERM o etac conjE, REPEAT_DETERM_N n o dtac (@{thm fun_eq_iff} RS subst),
-    CONJ_WRAP' (K (EVERY' [rtac ballI, REPEAT_DETERM o etac allE, rtac trans,
-      etac (o_apply RS subst), rtac o_apply])) morEs] 1
-  end;
-
-fun mk_iso_alt_tac mor_images mor_inv =
-  let
-    val n = length mor_images;
-    fun if_wrap_tac thm =
-      EVERY' [rtac ssubst, rtac @{thm bij_betw_iff_ex}, rtac exI, rtac conjI,
-        rtac @{thm inver_surj}, etac thm, etac thm, atac, etac conjI, atac]
-    val if_tac =
-      EVERY' [etac thin_rl, etac thin_rl, REPEAT o eresolve_tac [conjE, exE],
-        rtac conjI, atac, CONJ_WRAP' if_wrap_tac mor_images];
-    val only_if_tac =
-      EVERY' [rtac conjI, etac conjunct1, EVERY' (map (fn thm =>
-        EVERY' [rtac exE, rtac @{thm bij_betw_ex_weakE}, etac (conjunct2 RS thm)])
-        (map (mk_conjunctN n) (1 upto n))), REPEAT o rtac exI, rtac conjI, rtac mor_inv,
-        etac conjunct1, atac, atac, REPEAT_DETERM_N n o atac,
-        CONJ_WRAP' (K (etac conjunct2)) mor_images];
-  in
-    (rtac iffI THEN' if_tac THEN' only_if_tac) 1
-  end;
-
-fun mk_copy_str_tac set_maps alg_def alg_sets =
-  let
-    val n = length alg_sets;
-    val bij_betw_inv_tac =
-      EVERY' [etac thin_rl, REPEAT_DETERM_N n o EVERY' [dtac @{thm bij_betwI}, atac, atac],
-        REPEAT_DETERM_N (2 * n) o etac thin_rl, REPEAT_DETERM_N (n - 1) o etac conjI, atac];
-    fun set_tac thms =
-      EVERY' [rtac ord_eq_le_trans, resolve_tac thms, rtac subset_trans,
-          etac @{thm image_mono}, rtac equalityD1, etac @{thm bij_betw_imageE}];
-    val copy_str_tac =
-      CONJ_WRAP' (fn (thms, thm) =>
-        EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac set_mp,
-          rtac equalityD1, etac @{thm bij_betw_imageE}, rtac imageI, etac thm,
-          REPEAT_DETERM o rtac subset_UNIV, REPEAT_DETERM_N n o (set_tac thms)])
-      (set_maps ~~ alg_sets);
-  in
-    (rtac rev_mp THEN' DETERM o bij_betw_inv_tac THEN' rtac impI THEN'
-    stac alg_def THEN' copy_str_tac) 1
-  end;
-
-fun mk_copy_alg_tac set_maps alg_sets mor_def iso_alt copy_str =
-  let
-    val n = length alg_sets;
-    val fbetw_tac = CONJ_WRAP' (K (etac @{thm bij_betwE})) alg_sets;
-    fun set_tac thms =
-      EVERY' [rtac ord_eq_le_trans, resolve_tac thms, rtac subset_trans,
-        REPEAT_DETERM o etac conjE, etac @{thm image_mono},
-        rtac equalityD1, etac @{thm bij_betw_imageE}];
-    val mor_tac =
-      CONJ_WRAP' (fn (thms, thm) =>
-        EVERY' [rtac ballI, etac CollectE, etac @{thm inverE}, etac thm,
-          REPEAT_DETERM o rtac subset_UNIV, REPEAT_DETERM_N n o (set_tac thms)])
-      (set_maps ~~ alg_sets);
-  in
-    (rtac (iso_alt RS iffD2) THEN'
-    etac copy_str THEN' REPEAT_DETERM o atac THEN'
-    rtac conjI THEN' stac mor_def THEN' rtac conjI THEN' fbetw_tac THEN' mor_tac THEN'
-    CONJ_WRAP' (K atac) alg_sets) 1
-  end;
-
-fun mk_ex_copy_alg_tac n copy_str copy_alg =
-  EVERY' [REPEAT_DETERM_N n o rtac exI, rtac conjI, etac copy_str,
-    REPEAT_DETERM_N n o atac,
-    REPEAT_DETERM_N n o etac @{thm bij_betw_inver2},
-    REPEAT_DETERM_N n o etac @{thm bij_betw_inver1}, etac copy_alg,
-    REPEAT_DETERM_N n o atac,
-    REPEAT_DETERM_N n o etac @{thm bij_betw_inver2},
-    REPEAT_DETERM_N n o etac @{thm bij_betw_inver1}] 1;
-
-fun mk_bd_limit_tac n bd_Cinfinite =
-  EVERY' [REPEAT_DETERM o etac conjE, rtac rev_mp, rtac @{thm Cinfinite_limit_finite},
-    REPEAT_DETERM_N n o rtac @{thm finite.insertI}, rtac @{thm finite.emptyI},
-    REPEAT_DETERM_N n o etac @{thm insert_subsetI}, rtac @{thm empty_subsetI},
-    rtac bd_Cinfinite, rtac impI, etac bexE, rtac bexI,
-    CONJ_WRAP' (fn i =>
-      EVERY' [etac bspec, REPEAT_DETERM_N i o rtac @{thm insertI2}, rtac @{thm insertI1}])
-      (0 upto n - 1),
-    atac] 1;
-
-fun mk_min_algs_tac worel in_congs =
-  let
-    val minG_tac = EVERY' [rtac @{thm UN_cong}, rtac refl, dtac bspec, atac, etac arg_cong];
-    fun minH_tac thm =
-      EVERY' [rtac Un_cong, minG_tac, rtac @{thm image_cong}, rtac thm,
-        REPEAT_DETERM_N (length in_congs) o minG_tac, rtac refl];
-  in
-    (rtac (worel RS (@{thm wo_rel.worec_fixpoint} RS fun_cong)) THEN' rtac ssubst THEN'
-    rtac meta_eq_to_obj_eq THEN' rtac (worel RS @{thm wo_rel.adm_wo_def}) THEN'
-    REPEAT_DETERM_N 3 o rtac allI THEN' rtac impI THEN'
-    CONJ_WRAP_GEN' (EVERY' [rtac Pair_eqI, rtac conjI]) minH_tac in_congs) 1
-  end;
-
-fun mk_min_algs_mono_tac ctxt min_algs = EVERY' [stac @{thm relChain_def}, rtac allI, rtac allI,
-  rtac impI, rtac @{thm case_split}, rtac @{thm xt1(3)}, rtac min_algs, etac @{thm FieldI2},
-  rtac subsetI, rtac UnI1, rtac @{thm UN_I}, etac @{thm underS_I}, atac, atac,
-  rtac equalityD1, dtac @{thm notnotD}, hyp_subst_tac ctxt, rtac refl] 1;
-
-fun mk_min_algs_card_of_tac cT ct m worel min_algs in_bds bd_Card_order bd_Cnotzero
-  suc_Card_order suc_Cinfinite suc_Cnotzero suc_Asuc Asuc_Cinfinite =
-  let
-    val induct = worel RS
-      Drule.instantiate' [SOME cT] [NONE, SOME ct] @{thm well_order_induct_imp};
-    val src = 1 upto m + 1;
-    val dest = (m + 1) :: (1 upto m);
-    val absorbAs_tac = if m = 0 then K (all_tac)
-      else EVERY' [rtac @{thm ordIso_transitive}, rtac @{thm csum_cong1},
-        rtac @{thm ordIso_transitive},
-        BNF_Tactics.mk_rotate_eq_tac (rtac @{thm ordIso_refl} THEN'
-          FIRST' [rtac @{thm card_of_Card_order}, rtac @{thm Card_order_csum},
-            rtac @{thm Card_order_cexp}])
-        @{thm ordIso_transitive} @{thm csum_assoc} @{thm csum_com} @{thm csum_cong}
-        src dest,
-        rtac @{thm csum_absorb1}, rtac Asuc_Cinfinite, rtac ctrans, rtac @{thm ordLeq_csum1},
-        FIRST' [rtac @{thm Card_order_csum}, rtac @{thm card_of_Card_order}],
-        rtac @{thm ordLeq_cexp1}, rtac suc_Cnotzero, rtac @{thm Card_order_csum}];
-
-    val minG_tac = EVERY' [rtac @{thm UNION_Cinfinite_bound}, rtac @{thm ordLess_imp_ordLeq},
-      rtac @{thm ordLess_transitive}, rtac @{thm card_of_underS}, rtac suc_Card_order,
-      atac, rtac suc_Asuc, rtac ballI, etac allE, dtac mp, etac @{thm underS_E},
-      dtac mp, etac @{thm underS_Field}, REPEAT o etac conjE, atac, rtac Asuc_Cinfinite]
-
-    fun mk_minH_tac (min_alg, in_bd) = EVERY' [rtac @{thm ordIso_ordLeq_trans},
-      rtac @{thm card_of_ordIso_subst}, etac min_alg, rtac @{thm Un_Cinfinite_bound},
-      minG_tac, rtac ctrans, rtac @{thm card_of_image}, rtac ctrans, rtac in_bd, rtac ctrans,
-      rtac @{thm cexp_mono1}, rtac @{thm csum_mono1},
-      REPEAT_DETERM_N m o rtac @{thm csum_mono2},
-      CONJ_WRAP_GEN' (rtac @{thm csum_cinfinite_bound}) (K minG_tac) min_algs,
-      REPEAT_DETERM o FIRST'
-        [rtac @{thm card_of_Card_order}, rtac @{thm Card_order_csum},
-        rtac Asuc_Cinfinite, rtac bd_Card_order],
-      rtac @{thm ordIso_ordLeq_trans}, rtac @{thm cexp_cong1}, absorbAs_tac,
-      rtac @{thm csum_absorb1}, rtac Asuc_Cinfinite, rtac @{thm ctwo_ordLeq_Cinfinite},
-      rtac Asuc_Cinfinite, rtac bd_Card_order,
-      rtac @{thm ordIso_imp_ordLeq}, rtac @{thm cexp_cprod_ordLeq},
-      resolve_tac @{thms Card_order_csum Card_order_ctwo}, rtac suc_Cinfinite,
-      rtac bd_Cnotzero, rtac @{thm cardSuc_ordLeq}, rtac bd_Card_order, rtac Asuc_Cinfinite];
-  in
-    (rtac induct THEN'
-    rtac impI THEN'
-    CONJ_WRAP' mk_minH_tac (min_algs ~~ in_bds)) 1
-  end;
-
-fun mk_min_algs_least_tac cT ct worel min_algs alg_sets =
-  let
-    val induct = worel RS
-      Drule.instantiate' [SOME cT] [NONE, SOME ct] @{thm well_order_induct_imp};
-
-    val minG_tac = EVERY' [rtac @{thm UN_least}, etac allE, dtac mp, etac @{thm underS_E},
-      dtac mp, etac @{thm underS_Field}, REPEAT_DETERM o etac conjE, atac];
-
-    fun mk_minH_tac (min_alg, alg_set) = EVERY' [rtac ord_eq_le_trans, etac min_alg,
-      rtac @{thm Un_least}, minG_tac, rtac @{thm image_subsetI},
-      REPEAT_DETERM o eresolve_tac [CollectE, conjE], etac alg_set,
-      REPEAT_DETERM o FIRST' [atac, etac subset_trans THEN' minG_tac]];
-  in
-    (rtac induct THEN'
-    rtac impI THEN'
-    CONJ_WRAP' mk_minH_tac (min_algs ~~ alg_sets)) 1
-  end;
-
-fun mk_alg_min_alg_tac m alg_def min_alg_defs bd_limit bd_Cinfinite
-    set_bdss min_algs min_alg_monos =
-  let
-    val n = length min_algs;
-    fun mk_cardSuc_UNION_tac set_bds (mono, def) = EVERY'
-      [rtac bexE, rtac @{thm cardSuc_UNION_Cinfinite}, rtac bd_Cinfinite, rtac mono,
-       etac (def RSN (2, @{thm subset_trans[OF _ equalityD1]})), resolve_tac set_bds];
-    fun mk_conjunct_tac (set_bds, (min_alg, min_alg_def)) =
-      EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE],
-        EVERY' (map (mk_cardSuc_UNION_tac set_bds) (min_alg_monos ~~ min_alg_defs)), rtac bexE,
-        rtac bd_limit, REPEAT_DETERM_N (n - 1) o etac conjI, atac,
-        rtac (min_alg_def RS @{thm set_mp[OF equalityD2]}),
-        rtac @{thm UN_I}, REPEAT_DETERM_N (m + 3 * n) o etac thin_rl, atac, rtac set_mp,
-        rtac equalityD2, rtac min_alg, atac, rtac UnI2, rtac @{thm image_eqI}, rtac refl,
-        rtac CollectI, REPEAT_DETERM_N m o dtac asm_rl, REPEAT_DETERM_N n o etac thin_rl,
-        REPEAT_DETERM o etac conjE,
-        CONJ_WRAP' (K (FIRST' [atac,
-          EVERY' [etac subset_trans, rtac subsetI, rtac @{thm UN_I},
-            etac @{thm underS_I}, atac, atac]]))
-          set_bds];
-  in
-    (rtac (alg_def RS iffD2) THEN'
-    CONJ_WRAP' mk_conjunct_tac (set_bdss ~~ (min_algs ~~ min_alg_defs))) 1
-  end;
-
-fun mk_card_of_min_alg_tac min_alg_def card_of suc_Card_order suc_Asuc Asuc_Cinfinite =
-  EVERY' [stac min_alg_def, rtac @{thm UNION_Cinfinite_bound},
-    rtac @{thm ordIso_ordLeq_trans}, rtac @{thm card_of_Field_ordIso}, rtac suc_Card_order,
-    rtac @{thm ordLess_imp_ordLeq}, rtac suc_Asuc, rtac ballI, dtac rev_mp, rtac card_of,
-    REPEAT_DETERM o etac conjE, atac, rtac Asuc_Cinfinite] 1;
-
-fun mk_least_min_alg_tac min_alg_def least =
-  EVERY' [stac min_alg_def, rtac @{thm UN_least}, dtac least, dtac mp, atac,
-    REPEAT_DETERM o etac conjE, atac] 1;
-
-fun mk_alg_select_tac Abs_inverse {context = ctxt, prems = _} =
-  EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac ctxt] 1 THEN
-  unfold_thms_tac ctxt (Abs_inverse :: fst_snd_convs) THEN atac 1;
-
-fun mk_mor_select_tac mor_def mor_cong mor_comp mor_incl_min_alg alg_def alg_select alg_sets
-    set_maps str_init_defs =
-  let
-    val n = length alg_sets;
-    val fbetw_tac =
-      CONJ_WRAP' (K (EVERY' [rtac ballI, etac rev_bspec, etac CollectE, atac])) alg_sets;
-    val mor_tac =
-      CONJ_WRAP' (fn thm => EVERY' [rtac ballI, rtac thm]) str_init_defs;
-    fun alg_epi_tac ((alg_set, str_init_def), set_map) =
-      EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac CollectI,
-        rtac ballI, ftac (alg_select RS bspec), stac str_init_def, etac alg_set,
-        REPEAT_DETERM o FIRST' [rtac subset_UNIV,
-          EVERY' [rtac ord_eq_le_trans, resolve_tac set_map, rtac subset_trans,
-            etac @{thm image_mono}, rtac @{thm image_Collect_subsetI}, etac bspec, atac]]];
-  in
-    (rtac mor_cong THEN' REPEAT_DETERM_N n o (rtac sym THEN' rtac @{thm o_id}) THEN'
-    rtac (Thm.permute_prems 0 1 mor_comp) THEN' etac (Thm.permute_prems 0 1 mor_comp) THEN'
-    stac mor_def THEN' rtac conjI THEN' fbetw_tac THEN' mor_tac THEN' rtac mor_incl_min_alg THEN'
-    stac alg_def THEN' CONJ_WRAP' alg_epi_tac ((alg_sets ~~ str_init_defs) ~~ set_maps)) 1
-  end;
-
-fun mk_init_ex_mor_tac Abs_inverse copy_alg_ex alg_min_alg card_of_min_algs
-    mor_comp mor_select mor_incl_min_alg {context = ctxt, prems = _} =
-  let
-    val n = length card_of_min_algs;
-    val card_of_ordIso_tac = EVERY' [rtac ssubst, rtac @{thm card_of_ordIso},
-      rtac @{thm ordIso_symmetric}, rtac conjunct1, rtac conjunct2, atac];
-    fun internalize_tac card_of = EVERY' [rtac subst, rtac @{thm internalize_card_of_ordLeq2},
-      rtac @{thm ordLeq_ordIso_trans}, rtac card_of, rtac subst,
-      rtac @{thm Card_order_iff_ordIso_card_of}, rtac @{thm Card_order_cexp}];
-  in
-    (rtac rev_mp THEN'
-    REPEAT_DETERM_N (2 * n) o (rtac mp THEN' rtac @{thm ex_mono} THEN' rtac impI) THEN'
-    REPEAT_DETERM_N (n + 1) o etac thin_rl THEN' rtac (alg_min_alg RS copy_alg_ex) THEN'
-    REPEAT_DETERM_N n o atac THEN'
-    REPEAT_DETERM_N n o card_of_ordIso_tac THEN'
-    EVERY' (map internalize_tac card_of_min_algs) THEN'
-    rtac impI THEN'
-    REPEAT_DETERM o eresolve_tac [exE, conjE] THEN'
-    REPEAT_DETERM o rtac exI THEN'
-    rtac mor_select THEN' atac THEN' rtac CollectI THEN'
-    REPEAT_DETERM o rtac exI THEN'
-    rtac conjI THEN' rtac refl THEN' atac THEN'
-    K (unfold_thms_tac ctxt (Abs_inverse :: fst_snd_convs)) THEN'
-    etac mor_comp THEN' etac mor_incl_min_alg) 1
-  end;
-
-fun mk_init_unique_mor_tac m
-    alg_def alg_min_alg least_min_algs in_monos alg_sets morEs map_cong0s =
-  let
-    val n = length least_min_algs;
-    val ks = (1 upto n);
-
-    fun mor_tac morE in_mono = EVERY' [etac morE, rtac set_mp, rtac in_mono,
-      REPEAT_DETERM_N n o rtac @{thm Collect_restrict}, rtac CollectI,
-      REPEAT_DETERM_N (m + n) o (TRY o rtac conjI THEN' atac)];
-    fun cong_tac map_cong0 = EVERY' [rtac (map_cong0 RS arg_cong),
-      REPEAT_DETERM_N m o rtac refl,
-      REPEAT_DETERM_N n o (etac @{thm prop_restrict} THEN' atac)];
-
-    fun mk_alg_tac (alg_set, (in_mono, (morE, map_cong0))) = EVERY' [rtac ballI, rtac CollectI,
-      REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac conjI, rtac (alg_min_alg RS alg_set),
-      REPEAT_DETERM_N m o rtac subset_UNIV,
-      REPEAT_DETERM_N n o (etac subset_trans THEN' rtac @{thm Collect_restrict}),
-      rtac trans, mor_tac morE in_mono,
-      rtac trans, cong_tac map_cong0,
-      rtac sym, mor_tac morE in_mono];
-
-    fun mk_unique_tac (k, least_min_alg) =
-      select_prem_tac n (etac @{thm prop_restrict}) k THEN' rtac least_min_alg THEN'
-      stac alg_def THEN'
-      CONJ_WRAP' mk_alg_tac (alg_sets ~~ (in_monos ~~ (morEs ~~ map_cong0s)));
-  in
-    CONJ_WRAP' mk_unique_tac (ks ~~ least_min_algs) 1
-  end;
-
-fun mk_init_induct_tac m alg_def alg_min_alg least_min_algs alg_sets =
-  let
-    val n = length least_min_algs;
-
-    fun mk_alg_tac alg_set = EVERY' [rtac ballI, rtac CollectI,
-      REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac conjI, rtac (alg_min_alg RS alg_set),
-      REPEAT_DETERM_N m o rtac subset_UNIV,
-      REPEAT_DETERM_N n o (etac subset_trans THEN' rtac @{thm Collect_restrict}),
-      rtac mp, etac bspec, rtac CollectI,
-      REPEAT_DETERM_N m o (rtac conjI THEN' atac),
-      CONJ_WRAP' (K (etac subset_trans THEN' rtac @{thm Collect_restrict})) alg_sets,
-      CONJ_WRAP' (K (rtac ballI THEN' etac @{thm prop_restrict} THEN' atac)) alg_sets];
-
-    fun mk_induct_tac least_min_alg =
-      rtac ballI THEN' etac @{thm prop_restrict} THEN' rtac least_min_alg THEN'
-      stac alg_def THEN'
-      CONJ_WRAP' mk_alg_tac alg_sets;
-  in
-    CONJ_WRAP' mk_induct_tac least_min_algs 1
-  end;
-
-fun mk_mor_Rep_tac ctor_defs copy bijs inver_Abss inver_Reps {context = ctxt, prems = _} =
-  (K (unfold_thms_tac ctxt ctor_defs) THEN' rtac conjunct1 THEN' rtac copy THEN'
-  EVERY' (map (fn bij => EVERY' [rtac bij, atac, etac bexI, rtac UNIV_I]) bijs) THEN'
-  EVERY' (map rtac inver_Abss) THEN'
-  EVERY' (map rtac inver_Reps)) 1;
-
-fun mk_mor_Abs_tac inv inver_Abss inver_Reps =
-  (rtac inv THEN'
-  EVERY' (map2 (fn inver_Abs => fn inver_Rep =>
-    EVERY' [rtac conjI, rtac subset_UNIV, rtac conjI, rtac inver_Rep, rtac inver_Abs])
-    inver_Abss inver_Reps)) 1;
-
-fun mk_mor_fold_tac cT ct fold_defs ex_mor mor =
-  (EVERY' (map stac fold_defs) THEN' EVERY' [rtac rev_mp, rtac ex_mor, rtac impI] THEN'
-  REPEAT_DETERM_N (length fold_defs) o etac exE THEN'
-  rtac (Drule.instantiate' [SOME cT] [SOME ct] @{thm someI}) THEN' etac mor) 1;
-
-fun mk_fold_unique_mor_tac type_defs init_unique_mors Reps mor_comp mor_Abs mor_fold =
-  let
-    fun mk_unique type_def =
-      EVERY' [rtac @{thm surj_fun_eq}, rtac (type_def RS @{thm type_definition.Abs_image}),
-        rtac ballI, resolve_tac init_unique_mors,
-        EVERY' (map (fn thm => atac ORELSE' rtac thm) Reps),
-        rtac mor_comp, rtac mor_Abs, atac,
-        rtac mor_comp, rtac mor_Abs, rtac mor_fold];
-  in
-    CONJ_WRAP' mk_unique type_defs 1
-  end;
-
-fun mk_dtor_o_ctor_tac dtor_def foldx map_comp_id map_cong0L ctor_o_folds =
-  EVERY' [stac dtor_def, rtac ext, rtac trans, rtac o_apply, rtac trans, rtac foldx,
-    rtac trans, rtac map_comp_id, rtac trans, rtac map_cong0L,
-    EVERY' (map (fn thm => rtac ballI THEN' rtac (trans OF [thm RS fun_cong, id_apply]))
-      ctor_o_folds),
-    rtac sym, rtac id_apply] 1;
-
-fun mk_rec_tac rec_defs foldx fst_recs {context = ctxt, prems = _}=
-  unfold_thms_tac ctxt
-    (rec_defs @ map (fn thm => thm RS @{thm convol_expand_snd}) fst_recs) THEN
-  EVERY' [rtac trans, rtac o_apply, rtac trans, rtac (foldx RS @{thm arg_cong[of _ _ snd]}),
-    rtac @{thm snd_convol'}] 1;
-
-fun mk_rec_unique_mor_tac rec_defs fst_recs fold_unique_mor {context = ctxt, prems = _} =
-  unfold_thms_tac ctxt
-    (rec_defs @ map (fn thm => thm RS @{thm convol_expand_snd'}) fst_recs) THEN
-  etac fold_unique_mor 1;
-
-fun mk_ctor_induct_tac ctxt m set_mapss init_induct morEs mor_Abs Rep_invs Abs_invs Reps =
-  let
-    val n = length set_mapss;
-    val ks = 1 upto n;
-
-    fun mk_IH_tac Rep_inv Abs_inv set_map =
-      DETERM o EVERY' [dtac meta_mp, rtac (Rep_inv RS arg_cong RS subst), etac bspec,
-        dtac set_rev_mp, rtac equalityD1, rtac set_map, etac imageE,
-        hyp_subst_tac ctxt, rtac (Abs_inv RS ssubst), etac set_mp, atac, atac];
-
-    fun mk_closed_tac (k, (morE, set_maps)) =
-      EVERY' [select_prem_tac n (dtac asm_rl) k, rtac ballI, rtac impI,
-        rtac (mor_Abs RS morE RS arg_cong RS ssubst), atac,
-        REPEAT_DETERM o eresolve_tac [CollectE, conjE], dtac @{thm meta_spec},
-        EVERY' (map3 mk_IH_tac Rep_invs Abs_invs (drop m set_maps)), atac];
-
-    fun mk_induct_tac (Rep, Rep_inv) =
-      EVERY' [rtac (Rep_inv RS arg_cong RS subst), etac (Rep RSN (2, bspec))];
-  in
-    (rtac mp THEN' rtac impI THEN'
-    DETERM o CONJ_WRAP_GEN' (etac conjE THEN' rtac conjI) mk_induct_tac (Reps ~~ Rep_invs) THEN'
-    rtac init_induct THEN'
-    DETERM o CONJ_WRAP' mk_closed_tac (ks ~~ (morEs ~~ set_mapss))) 1
-  end;
-
-fun mk_ctor_induct2_tac cTs cts ctor_induct weak_ctor_inducts {context = ctxt, prems = _} =
-  let
-    val n = length weak_ctor_inducts;
-    val ks = 1 upto n;
-    fun mk_inner_induct_tac induct i =
-      EVERY' [rtac allI, fo_rtac induct ctxt,
-        select_prem_tac n (dtac @{thm meta_spec2}) i,
-        REPEAT_DETERM_N n o
-          EVERY' [dtac meta_mp THEN_ALL_NEW Goal.norm_hhf_tac ctxt,
-            REPEAT_DETERM o dtac @{thm meta_spec}, etac (spec RS meta_mp), atac],
-        atac];
-  in
-    EVERY' [rtac rev_mp, rtac (Drule.instantiate' cTs cts ctor_induct),
-      EVERY' (map2 mk_inner_induct_tac weak_ctor_inducts ks), rtac impI,
-      REPEAT_DETERM o eresolve_tac [conjE, allE],
-      CONJ_WRAP' (K atac) ks] 1
-  end;
-
-fun mk_map_tac m n foldx map_comp_id map_cong0 =
-  EVERY' [rtac ext, rtac trans, rtac o_apply, rtac trans, rtac foldx, rtac trans, rtac o_apply,
-    rtac trans, rtac (map_comp_id RS arg_cong), rtac trans, rtac (map_cong0 RS arg_cong),
-    REPEAT_DETERM_N m o rtac refl,
-    REPEAT_DETERM_N n o (EVERY' (map rtac [trans, o_apply, id_apply])),
-    rtac sym, rtac o_apply] 1;
-
-fun mk_ctor_map_unique_tac fold_unique sym_map_comps ctxt =
-  rtac fold_unique 1 THEN
-  unfold_thms_tac ctxt (sym_map_comps @ @{thms o_assoc[symmetric] id_o o_id}) THEN
-  ALLGOALS atac;
-
-fun mk_set_tac foldx = EVERY' [rtac ext, rtac trans, rtac o_apply,
-  rtac trans, rtac foldx, rtac sym, rtac o_apply] 1;
-
-fun mk_ctor_set_tac set set_map set_maps =
-  let
-    val n = length set_maps;
-    fun mk_UN thm = rtac (thm RS @{thm arg_cong[of _ _ Union]} RS trans) THEN'
-      rtac @{thm Union_image_eq};
-  in
-    EVERY' [rtac (set RS @{thm comp_eq_dest} RS trans), rtac Un_cong,
-      rtac (trans OF [set_map, trans_fun_cong_image_id_id_apply]),
-      REPEAT_DETERM_N (n - 1) o rtac Un_cong,
-      EVERY' (map mk_UN set_maps)] 1
-  end;
-
-fun mk_set_nat_tac m induct_tac set_mapss
-    ctor_maps csets ctor_sets i {context = ctxt, prems = _} =
-  let
-    val n = length ctor_maps;
-
-    fun useIH set_nat = EVERY' [rtac trans, rtac @{thm image_UN}, rtac trans, rtac @{thm UN_cong},
-      rtac refl, Goal.assume_rule_tac ctxt, rtac sym, rtac trans, rtac @{thm UN_cong},
-      rtac set_nat, rtac refl, rtac @{thm UN_simps(10)}];
-
-    fun mk_set_nat cset ctor_map ctor_set set_nats =
-      EVERY' [rtac trans, rtac @{thm image_cong}, rtac ctor_set, rtac refl,
-        rtac sym, rtac (trans OF [ctor_map RS HOL_arg_cong cset, ctor_set RS trans]),
-        rtac sym, EVERY' (map rtac [trans, @{thm image_Un}, Un_cong]),
-        rtac sym, rtac (nth set_nats (i - 1)),
-        REPEAT_DETERM_N (n - 1) o EVERY' (map rtac [trans, @{thm image_Un}, Un_cong]),
-        EVERY' (map useIH (drop m set_nats))];
-  in
-    (induct_tac THEN' EVERY' (map4 mk_set_nat csets ctor_maps ctor_sets set_mapss)) 1
-  end;
-
-fun mk_set_bd_tac m induct_tac bd_Cinfinite set_bdss ctor_sets i ctxt =
-  let
-    val n = length ctor_sets;
-
-    fun useIH set_bd = EVERY' [rtac @{thm UNION_Cinfinite_bound}, rtac set_bd, rtac ballI,
-      Goal.assume_rule_tac ctxt, rtac bd_Cinfinite];
-
-    fun mk_set_nat ctor_set set_bds =
-      EVERY' [rtac @{thm ordIso_ordLeq_trans}, rtac @{thm card_of_ordIso_subst}, rtac ctor_set,
-        rtac (bd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})), rtac (nth set_bds (i - 1)),
-        REPEAT_DETERM_N (n - 1) o rtac (bd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})),
-        EVERY' (map useIH (drop m set_bds))];
-  in
-    (induct_tac THEN' EVERY' (map2 mk_set_nat ctor_sets set_bdss)) 1
-  end;
-
-fun mk_mcong_tac induct_tac set_setsss map_cong0s ctor_maps {context = ctxt, prems = _} =
-  let
-    fun use_asm thm = EVERY' [etac bspec, etac set_rev_mp, rtac thm];
-
-    fun useIH set_sets = EVERY' [rtac mp, Goal.assume_rule_tac ctxt,
-      CONJ_WRAP' (fn thm =>
-        EVERY' [rtac ballI, etac bspec, etac set_rev_mp, etac thm]) set_sets];
-
-    fun mk_map_cong0 ctor_map map_cong0 set_setss =
-      EVERY' [rtac impI, REPEAT_DETERM o etac conjE,
-        rtac trans, rtac ctor_map, rtac trans, rtac (map_cong0 RS arg_cong),
-        EVERY' (map use_asm (map hd set_setss)),
-        EVERY' (map useIH (transpose (map tl set_setss))),
-        rtac sym, rtac ctor_map];
-  in
-    (induct_tac THEN' EVERY' (map3 mk_map_cong0 ctor_maps map_cong0s set_setsss)) 1
-  end;
-
-fun mk_le_rel_OO_tac m induct ctor_nchotomys ctor_Irels rel_mono_strongs rel_OOs 
-    {context = ctxt, prems = _} =
-  EVERY' (rtac induct ::
-  map4 (fn nchotomy => fn Irel => fn rel_mono => fn rel_OO =>
-    EVERY' [rtac impI, etac (nchotomy RS @{thm nchotomy_relcomppE}),
-      REPEAT_DETERM_N 2 o dtac (Irel RS iffD1), rtac (Irel RS iffD2),
-      rtac rel_mono, rtac (rel_OO RS @{thm predicate2_eqD} RS iffD2),
-      rtac @{thm relcomppI}, atac, atac,
-      REPEAT_DETERM_N m o EVERY' [rtac ballI, rtac ballI, rtac impI, atac],
-      REPEAT_DETERM_N (length rel_OOs) o
-        EVERY' [rtac ballI, rtac ballI, Goal.assume_rule_tac ctxt]])
-  ctor_nchotomys ctor_Irels rel_mono_strongs rel_OOs) 1;
-
-(* BNF tactics *)
-
-fun mk_map_id0_tac map_id0s unique =
-  (rtac sym THEN' rtac unique THEN'
-  EVERY' (map (fn thm =>
-    EVERY' [rtac trans, rtac @{thm id_o}, rtac trans, rtac sym, rtac @{thm o_id},
-      rtac (thm RS sym RS arg_cong)]) map_id0s)) 1;
-
-fun mk_map_comp0_tac map_comp0s ctor_maps unique iplus1 =
-  let
-    val i = iplus1 - 1;
-    val unique' = Thm.permute_prems 0 i unique;
-    val map_comp0s' = drop i map_comp0s @ take i map_comp0s;
-    val ctor_maps' = drop i ctor_maps @ take i ctor_maps;
-    fun mk_comp comp simp =
-      EVERY' [rtac ext, rtac trans, rtac o_apply, rtac trans, rtac o_apply,
-        rtac trans, rtac (simp RS arg_cong), rtac trans, rtac simp,
-        rtac trans, rtac (comp RS arg_cong), rtac sym, rtac o_apply];
-  in
-    (rtac sym THEN' rtac unique' THEN' EVERY' (map2 mk_comp map_comp0s' ctor_maps')) 1
-  end;
-
-fun mk_set_map0_tac set_nat =
-  EVERY' (map rtac [ext, trans, o_apply, sym, trans, o_apply, set_nat]) 1;
-
-fun mk_bd_card_order_tac bd_card_orders =
-  CONJ_WRAP_GEN' (rtac @{thm card_order_csum}) rtac bd_card_orders 1;
-
-fun mk_wit_tac ctxt n ctor_set wit =
-  REPEAT_DETERM (atac 1 ORELSE
-    EVERY' [dtac set_rev_mp, rtac equalityD1, resolve_tac ctor_set,
-    REPEAT_DETERM o
-      (TRY o REPEAT_DETERM o etac UnE THEN' TRY o etac @{thm UN_E} THEN'
-        (eresolve_tac wit ORELSE'
-        (dresolve_tac wit THEN'
-          (etac FalseE ORELSE'
-          EVERY' [hyp_subst_tac ctxt, dtac set_rev_mp, rtac equalityD1, resolve_tac ctor_set,
-            REPEAT_DETERM_N n o etac UnE]))))] 1);
-
-fun mk_ctor_rel_tac ctxt in_Irels i in_rel map_comp0 map_cong0 ctor_map ctor_sets ctor_inject
-  ctor_dtor set_map0s ctor_set_incls ctor_set_set_inclss =
-  let
-    val m = length ctor_set_incls;
-    val n = length ctor_set_set_inclss;
-
-    val (passive_set_map0s, active_set_map0s) = chop m set_map0s;
-    val in_Irel = nth in_Irels (i - 1);
-    val le_arg_cong_ctor_dtor = ctor_dtor RS arg_cong RS ord_eq_le_trans;
-    val eq_arg_cong_ctor_dtor = ctor_dtor RS arg_cong RS trans;
-    val if_tac =
-      EVERY' [dtac (in_Irel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
-        rtac (in_rel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
-        EVERY' (map2 (fn set_map0 => fn ctor_set_incl =>
-          EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac set_map0,
-            rtac ord_eq_le_trans, rtac trans_fun_cong_image_id_id_apply,
-            rtac (ctor_set_incl RS subset_trans), etac le_arg_cong_ctor_dtor])
-        passive_set_map0s ctor_set_incls),
-        CONJ_WRAP' (fn (in_Irel, (set_map0, ctor_set_set_incls)) =>
-          EVERY' [rtac ord_eq_le_trans, rtac set_map0, rtac @{thm image_subsetI}, rtac CollectI,
-            rtac @{thm prod_caseI}, rtac (in_Irel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
-            CONJ_WRAP' (fn thm =>
-              EVERY' (map etac [thm RS subset_trans, le_arg_cong_ctor_dtor]))
-            ctor_set_set_incls,
-            rtac conjI, rtac refl, rtac refl])
-        (in_Irels ~~ (active_set_map0s ~~ ctor_set_set_inclss)),
-        CONJ_WRAP' (fn conv =>
-          EVERY' [rtac trans, rtac map_comp0, rtac trans, rtac map_cong0,
-          REPEAT_DETERM_N m o rtac @{thm fun_cong[OF o_id]},
-          REPEAT_DETERM_N n o EVERY' (map rtac [trans, o_apply, conv]),
-          rtac (ctor_inject RS iffD1), rtac trans, rtac sym, rtac ctor_map,
-          etac eq_arg_cong_ctor_dtor])
-        fst_snd_convs];
-    val only_if_tac =
-      EVERY' [dtac (in_rel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
-        rtac (in_Irel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
-        CONJ_WRAP' (fn (ctor_set, passive_set_map0) =>
-          EVERY' [rtac ord_eq_le_trans, rtac ctor_set, rtac @{thm Un_least},
-            rtac ord_eq_le_trans, rtac @{thm box_equals[OF _ refl]},
-            rtac passive_set_map0, rtac trans_fun_cong_image_id_id_apply, atac,
-            CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least}))
-              (fn (active_set_map0, in_Irel) => EVERY' [rtac ord_eq_le_trans,
-                rtac @{thm UN_cong[OF _ refl]}, rtac active_set_map0, rtac @{thm UN_least},
-                dtac set_rev_mp, etac @{thm image_mono}, etac imageE,
-                dtac @{thm ssubst_mem[OF pair_collapse]},
-                REPEAT_DETERM o eresolve_tac (CollectE :: conjE ::
-                  @{thms prod_caseE iffD1[OF Pair_eq, elim_format]}),
-                hyp_subst_tac ctxt,
-                dtac (in_Irel RS iffD1), dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE,
-                TRY o
-                  EVERY' [dtac (Thm.permute_prems 0 1 @{thm ssubst_mem}), atac, hyp_subst_tac ctxt],
-                REPEAT_DETERM o eresolve_tac [CollectE, conjE], atac])
-            (rev (active_set_map0s ~~ in_Irels))])
-        (ctor_sets ~~ passive_set_map0s),
-        rtac conjI,
-        REPEAT_DETERM_N 2 o EVERY' [rtac trans, rtac ctor_map, rtac (ctor_inject RS iffD2),
-          rtac trans, rtac map_comp0, rtac trans, rtac map_cong0,
-          REPEAT_DETERM_N m o rtac @{thm fun_cong[OF o_id]},
-          EVERY' (map (fn in_Irel => EVERY' [rtac trans, rtac o_apply, dtac set_rev_mp, atac,
-            dtac @{thm ssubst_mem[OF pair_collapse]},
-            REPEAT_DETERM o
-              eresolve_tac (CollectE :: conjE :: @{thms prod_caseE iffD1[OF Pair_eq, elim_format]}),
-            hyp_subst_tac ctxt,
-            dtac (in_Irel RS iffD1), dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE, atac])
-          in_Irels),
-          atac]]
-  in
-    EVERY' [rtac iffI, if_tac, only_if_tac] 1
-  end;
-
-fun mk_rel_induct_tac m ctor_induct2 ks ctor_rels rel_mono_strongs {context = ctxt, prems = IHs} =
-  let val n = length ks;
-  in
-    unfold_tac ctxt @{thms le_fun_def le_bool_def all_simps(1,2)[symmetric]} THEN
-    EVERY' [REPEAT_DETERM o rtac allI, rtac ctor_induct2,
-      EVERY' (map3 (fn IH => fn ctor_rel => fn rel_mono_strong =>
-        EVERY' [rtac impI, dtac (ctor_rel RS iffD1), rtac (IH RS @{thm spec2} RS mp),
-          etac rel_mono_strong,
-          REPEAT_DETERM_N m o rtac @{thm ballI[OF ballI[OF imp_refl]]},
-          EVERY' (map (fn j =>
-            EVERY' [select_prem_tac n (dtac asm_rl) j, rtac @{thm ballI[OF ballI]},
-              Goal.assume_rule_tac ctxt]) ks)])
-      IHs ctor_rels rel_mono_strongs)] 1
-  end;
-
-fun mk_fold_transfer_tac m rel_induct map_transfers folds {context = ctxt, prems = _} =
-  let
-    val n = length map_transfers;
-  in
-    unfold_thms_tac ctxt
-      @{thms fun_rel_def_butlast all_conj_distrib[symmetric] imp_conjR[symmetric]} THEN
-    unfold_thms_tac ctxt @{thms fun_rel_iff_leq_vimage2p} THEN
-    HEADGOAL (EVERY'
-      [REPEAT_DETERM o resolve_tac [allI, impI], rtac rel_induct,
-      EVERY' (map (fn map_transfer => EVERY'
-        [REPEAT_DETERM o resolve_tac [allI, impI, @{thm vimage2pI}],
-        SELECT_GOAL (unfold_thms_tac ctxt folds),
-        etac @{thm predicate2D_vimage2p},
-        rtac (funpow (m + n + 1) (fn thm => thm RS @{thm fun_relD}) map_transfer),
-        REPEAT_DETERM_N m o rtac @{thm id_transfer},
-        REPEAT_DETERM_N n o rtac @{thm vimage2p_fun_rel},
-        atac])
-      map_transfers)])
-  end;
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_lfp_util.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,63 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_lfp_util.ML
-    Author:     Dmitriy Traytel, TU Muenchen
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2012
-
-Library for the datatype construction.
-*)
-
-signature BNF_LFP_UTIL =
-sig
-  val HOL_arg_cong: cterm -> thm
-
-  val mk_bij_betw: term -> term -> term -> term
-  val mk_cardSuc: term -> term
-  val mk_inver: term -> term -> term -> term
-  val mk_not_empty: term -> term
-  val mk_not_eq: term -> term -> term
-  val mk_rapp: term -> typ -> term
-  val mk_relChain: term -> term -> term
-  val mk_underS: term -> term
-  val mk_worec: term -> term -> term
-end;
-
-structure BNF_LFP_Util : BNF_LFP_UTIL =
-struct
-
-open BNF_Util
-
-fun HOL_arg_cong ct = Drule.instantiate'
-  (map SOME (Thm.dest_ctyp (Thm.ctyp_of_term ct))) [NONE, NONE, SOME ct] arg_cong;
-
-(*reverse application*)
-fun mk_rapp arg T = Term.absdummy (fastype_of arg --> T) (Bound 0 $ arg);
-
-fun mk_underS r =
-  let val T = fst (dest_relT (fastype_of r));
-  in Const (@{const_name underS}, mk_relT (T, T) --> T --> HOLogic.mk_setT T) $ r end;
-
-fun mk_worec r f =
-  let val (A, AB) = apfst domain_type (dest_funT (fastype_of f));
-  in Const (@{const_name wo_rel.worec}, mk_relT (A, A) --> (AB --> AB) --> AB) $ r $ f end;
-
-fun mk_relChain r f =
-  let val (A, AB) = `domain_type (fastype_of f);
-  in Const (@{const_name relChain}, mk_relT (A, A) --> AB --> HOLogic.boolT) $ r $ f end;
-
-fun mk_cardSuc r =
-  let val T = fst (dest_relT (fastype_of r));
-  in Const (@{const_name cardSuc}, mk_relT (T, T) --> mk_relT (`I (HOLogic.mk_setT T))) $ r end;
-
-fun mk_bij_betw f A B =
- Const (@{const_name bij_betw},
-   fastype_of f --> fastype_of A --> fastype_of B --> HOLogic.boolT) $ f $ A $ B;
-
-fun mk_inver f g A =
- Const (@{const_name inver}, fastype_of f --> fastype_of g --> fastype_of A --> HOLogic.boolT) $
-   f $ g $ A;
-
-fun mk_not_eq x y = HOLogic.mk_not (HOLogic.mk_eq (x, y));
-
-fun mk_not_empty B = mk_not_eq B (HOLogic.mk_set (HOLogic.dest_setT (fastype_of B)) []);
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_tactics.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,109 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_tactics.ML
-    Author:     Dmitriy Traytel, TU Muenchen
-    Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2012
-
-General tactics for bounded natural functors.
-*)
-
-signature BNF_TACTICS =
-sig
-  include CTR_SUGAR_GENERAL_TACTICS
-
-  val fo_rtac: thm -> Proof.context -> int -> tactic
-
-  val mk_flatten_assoc_tac: (int -> tactic) -> thm -> thm -> thm -> tactic
-  val mk_rotate_eq_tac: (int -> tactic) -> thm -> thm -> thm -> thm -> ''a list -> ''a list ->
-    int -> tactic
-
-  val mk_pointfree: Proof.context -> thm -> thm
-
-  val mk_Abs_bij_thm: Proof.context -> thm -> thm -> thm
-  val mk_Abs_inj_thm: thm -> thm
-
-  val mk_ctor_or_dtor_rel_tac:
-    thm -> thm list -> thm list -> thm -> {prems: 'a, context: Proof.context} -> tactic
-
-  val mk_map_comp_id_tac: thm -> tactic
-  val mk_map_cong0_tac: int -> thm -> {prems: 'a, context: Proof.context} -> tactic
-  val mk_map_cong0L_tac: int -> thm -> thm -> tactic
-end;
-
-structure BNF_Tactics : BNF_TACTICS =
-struct
-
-open Ctr_Sugar_General_Tactics
-open BNF_Util
-
-(*stolen from Christian Urban's Cookbook*)
-fun fo_rtac thm = Subgoal.FOCUS (fn {concl, ...} =>
-  let
-    val concl_pat = Drule.strip_imp_concl (cprop_of thm)
-    val insts = Thm.first_order_match (concl_pat, concl)
-  in
-    rtac (Drule.instantiate_normalize insts thm) 1
-  end);
-
-(*transforms f (g x) = h (k x) into f o g = h o k using first order matches for f, g, h, and k*)
-fun mk_pointfree ctxt thm = thm
-  |> Thm.prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq
-  |> pairself (dest_comb #> apsnd (dest_comb #> fst) #> HOLogic.mk_comp)
-  |> mk_Trueprop_eq
-  |> (fn goal => Goal.prove_sorry ctxt [] [] goal
-    (K (rtac ext 1 THEN unfold_thms_tac ctxt [o_apply, mk_sym thm] THEN rtac refl 1)))
-  |> Thm.close_derivation;
-
-
-(* Theorems for open typedefs with UNIV as representing set *)
-
-fun mk_Abs_inj_thm inj = inj OF (replicate 2 UNIV_I);
-fun mk_Abs_bij_thm ctxt Abs_inj_thm surj = rule_by_tactic ctxt ((rtac surj THEN' etac exI) 1)
-  (Abs_inj_thm RS @{thm bijI});
-
-
-
-(* General tactic generators *)
-
-(*applies assoc rule to the lhs of an equation as long as possible*)
-fun mk_flatten_assoc_tac refl_tac trans assoc cong = rtac trans 1 THEN
-  REPEAT_DETERM (CHANGED ((FIRST' [rtac trans THEN' rtac assoc, rtac cong THEN' refl_tac]) 1)) THEN
-  refl_tac 1;
-
-(*proves two sides of an equation to be equal assuming both are flattened and rhs can be obtained
-from lhs by the given permutation of monoms*)
-fun mk_rotate_eq_tac refl_tac trans assoc com cong =
-  let
-    fun gen_tac [] [] = K all_tac
-      | gen_tac [x] [y] = if x = y then refl_tac else error "mk_rotate_eq_tac: different lists"
-      | gen_tac (x :: xs) (y :: ys) = if x = y
-        then rtac cong THEN' refl_tac THEN' gen_tac xs ys
-        else rtac trans THEN' rtac com THEN'
-          K (mk_flatten_assoc_tac refl_tac trans assoc cong) THEN'
-          gen_tac (xs @ [x]) (y :: ys)
-      | gen_tac _ _ = error "mk_rotate_eq_tac: different lists";
-  in
-    gen_tac
-  end;
-
-fun mk_ctor_or_dtor_rel_tac srel_def IJrel_defs IJsrel_defs dtor_srel {context = ctxt, prems = _} =
-  unfold_thms_tac ctxt IJrel_defs THEN
-  rtac (unfold_thms ctxt (IJrel_defs @ IJsrel_defs @
-    @{thms Collect_mem_eq mem_Collect_eq fst_conv snd_conv}) dtor_srel RS trans) 1 THEN
-  unfold_thms_tac ctxt (srel_def ::
-    @{thms pair_collapse Collect_mem_eq mem_Collect_eq prod.cases fst_conv snd_conv
-      split_conv}) THEN
-  rtac refl 1;
-
-fun mk_map_comp_id_tac map_comp0 =
-  (rtac trans THEN' rtac map_comp0 THEN' REPEAT_DETERM o stac @{thm o_id} THEN' rtac refl) 1;
-
-fun mk_map_cong0_tac m map_cong0 {context = ctxt, prems = _} =
-  EVERY' [rtac mp, rtac map_cong0,
-    CONJ_WRAP' (K (rtac ballI THEN' Goal.assume_rule_tac ctxt)) (1 upto m)] 1;
-
-fun mk_map_cong0L_tac passive map_cong0 map_id =
-  (rtac trans THEN' rtac map_cong0 THEN' EVERY' (replicate passive (rtac refl))) 1 THEN
-  REPEAT_DETERM (EVERY' [rtac trans, etac bspec, atac, rtac sym, rtac @{thm id_apply}] 1) THEN
-  rtac map_id 1;
-
-end;
--- a/src/HOL/Tools/BNF/Tools/bnf_util.ML	Mon Jan 20 18:24:56 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,583 +0,0 @@
-(*  Title:      HOL/BNF/Tools/bnf_util.ML
-    Author:     Dmitriy Traytel, TU Muenchen
-    Copyright   2012
-
-Library for bounded natural functors.
-*)
-
-signature BNF_UTIL =
-sig
-  include CTR_SUGAR_UTIL
-
-  val map6: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g) ->
-    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list
-  val map7: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h) ->
-    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list
-  val map8: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i) ->
-    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list -> 'i list
-  val map9: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j) ->
-    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
-    'i list -> 'j list
-  val map10: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j -> 'k) ->
-    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
-    'i list -> 'j list -> 'k list
-  val map11: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j -> 'k -> 'l) ->
-    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
-    'i list -> 'j list -> 'k list -> 'l list
-  val map12: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j -> 'k -> 'l -> 'm) ->
-    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
-    'i list -> 'j list -> 'k list -> 'l list -> 'm list
-  val map13: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j -> 'k -> 'l -> 'm -> 'n) ->
-    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
-    'i list -> 'j list -> 'k list -> 'l list -> 'm list -> 'n list
-  val map14:
-    ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j -> 'k -> 'l -> 'm -> 'n -> 'o) ->
-    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
-    'i list -> 'j list -> 'k list -> 'l list -> 'm list -> 'n list -> 'o list
-  val fold_map4: ('a -> 'b -> 'c -> 'd -> 'e -> 'f * 'e) ->
-    'a list -> 'b list -> 'c list -> 'd list -> 'e -> 'f list * 'e
-  val fold_map5: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g * 'f) ->
-    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f -> 'g list * 'f
-  val fold_map6: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h * 'g) ->
-    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g -> 'h list * 'g
-  val fold_map7: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i * 'h) ->
-    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h -> 'i list * 'h
-  val fold_map8: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j * 'i) ->
-    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list -> 'i ->
-    'j list * 'i
-  val fold_map9: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j -> 'k * 'j) ->
-    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
-    'i list -> 'j -> 'k list * 'j
-  val split_list4: ('a * 'b * 'c * 'd) list -> 'a list * 'b list * 'c list * 'd list
-  val split_list5: ('a * 'b * 'c * 'd * 'e) list -> 'a list * 'b list * 'c list * 'd list * 'e list
-  val find_indices: ('b * 'a -> bool) -> 'a list -> 'b list -> int list
-
-  val mk_TFreess: int list -> Proof.context -> typ list list * Proof.context
-  val mk_Freesss: string -> typ list list list -> Proof.context ->
-    term list list list * Proof.context
-  val mk_Freessss: string -> typ list list list list -> Proof.context ->
-    term list list list list * Proof.context
-  val nonzero_string_of_int: int -> string
-  val retype_free: typ -> term -> term
-
-  val binder_fun_types: typ -> typ list
-  val body_fun_type: typ -> typ
-  val num_binder_types: typ -> int
-  val strip_fun_type: typ -> typ list * typ
-  val strip_typeN: int -> typ -> typ list * typ
-
-  val mk_pred2T: typ -> typ -> typ
-  val mk_relT: typ * typ -> typ
-  val dest_relT: typ -> typ * typ
-  val dest_pred2T: typ -> typ * typ
-  val mk_sumT: typ * typ -> typ
-
-  val ctwo: term
-  val fst_const: typ -> term
-  val snd_const: typ -> term
-  val Id_const: typ -> term
-
-  val mk_Ball: term -> term -> term
-  val mk_Bex: term -> term -> term
-  val mk_Card_order: term -> term
-  val mk_Field: term -> term
-  val mk_Gr: term -> term -> term
-  val mk_Grp: term -> term -> term
-  val mk_UNION: term -> term -> term
-  val mk_Union: typ -> term
-  val mk_card_binop: string -> (typ * typ -> typ) -> term -> term -> term
-  val mk_card_of: term -> term
-  val mk_card_order: term -> term
-  val mk_cexp: term -> term -> term
-  val mk_cinfinite: term -> term
-  val mk_collect: term list -> typ -> term
-  val mk_converse: term -> term
-  val mk_conversep: term -> term
-  val mk_cprod: term -> term -> term
-  val mk_csum: term -> term -> term
-  val mk_dir_image: term -> term -> term
-  val mk_fun_rel: term -> term -> term
-  val mk_image: term -> term
-  val mk_in: term list -> term list -> typ -> term
-  val mk_leq: term -> term -> term
-  val mk_ordLeq: term -> term -> term
-  val mk_rel_comp: term * term -> term
-  val mk_rel_compp: term * term -> term
-
-  (*parameterized terms*)
-  val mk_nthN: int -> term -> int -> term
-
-  (*parameterized thms*)
-  val mk_Un_upper: int -> int -> thm
-  val mk_conjIN: int -> thm
-  val mk_conjunctN: int -> int -> thm
-  val conj_dests: int -> thm -> thm list
-  val mk_nthI: int -> int -> thm
-  val mk_nth_conv: int -> int -> thm
-  val mk_ordLeq_csum: int -> int -> thm -> thm
-  val mk_UnIN: int -> int -> thm
-
-  val Pair_eqD: thm
-  val Pair_eqI: thm
-  val ctrans: thm
-  val id_apply: thm
-  val meta_mp: thm
-  val meta_spec: thm
-  val o_apply: thm
-  val set_mp: thm
-  val set_rev_mp: thm
-  val subset_UNIV: thm
-  val mk_sym: thm -> thm
-  val mk_trans: thm -> thm -> thm
-
-  val is_refl: thm -> bool
-  val is_concl_refl: thm -> bool
-  val no_refl: thm list -> thm list
-  val no_reflexive: thm list -> thm list
-
-  val fold_thms: Proof.context -> thm list -> thm -> thm
-
-  val parse_binding_colon: binding parser
-  val parse_opt_binding_colon: binding parser
-  val parse_type_args_named_constrained: (binding option * (string * string option)) list parser
-  val parse_map_rel_bindings: (binding * binding) parser
-
-  val typedef: binding * (string * sort) list * mixfix -> term ->
-    (binding * binding) option -> tactic -> local_theory -> (string * Typedef.info) * local_theory
-end;
-
-structure BNF_Util : BNF_UTIL =
-struct
-
-open Ctr_Sugar_Util
-
-(* Library proper *)
-
-fun map6 _ [] [] [] [] [] [] = []
-  | map6 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) =
-    f x1 x2 x3 x4 x5 x6 :: map6 f x1s x2s x3s x4s x5s x6s
-  | map6 _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun map7 _ [] [] [] [] [] [] [] = []
-  | map7 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) =
-    f x1 x2 x3 x4 x5 x6 x7 :: map7 f x1s x2s x3s x4s x5s x6s x7s
-  | map7 _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun map8 _ [] [] [] [] [] [] [] [] = []
-  | map8 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) (x8::x8s) =
-    f x1 x2 x3 x4 x5 x6 x7 x8 :: map8 f x1s x2s x3s x4s x5s x6s x7s x8s
-  | map8 _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun map9 _ [] [] [] [] [] [] [] [] [] = []
-  | map9 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) (x8::x8s)
-      (x9::x9s) =
-    f x1 x2 x3 x4 x5 x6 x7 x8 x9 :: map9 f x1s x2s x3s x4s x5s x6s x7s x8s x9s
-  | map9 _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun map10 _ [] [] [] [] [] [] [] [] [] [] = []
-  | map10 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) (x8::x8s)
-      (x9::x9s) (x10::x10s) =
-    f x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 :: map10 f x1s x2s x3s x4s x5s x6s x7s x8s x9s x10s
-  | map10 _ _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun map11 _ [] [] [] [] [] [] [] [] [] [] [] = []
-  | map11 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) (x8::x8s)
-      (x9::x9s) (x10::x10s) (x11::x11s) =
-    f x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 :: map11 f x1s x2s x3s x4s x5s x6s x7s x8s x9s x10s x11s
-  | map11 _ _ _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun map12 _ [] [] [] [] [] [] [] [] [] [] [] [] = []
-  | map12 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) (x8::x8s)
-      (x9::x9s) (x10::x10s) (x11::x11s) (x12::x12s) =
-    f x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 ::
-      map12 f x1s x2s x3s x4s x5s x6s x7s x8s x9s x10s x11s x12s
-  | map12 _ _ _ _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun map13 _ [] [] [] [] [] [] [] [] [] [] [] [] [] = []
-  | map13 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) (x8::x8s)
-      (x9::x9s) (x10::x10s) (x11::x11s) (x12::x12s) (x13::x13s) =
-    f x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13 ::
-      map13 f x1s x2s x3s x4s x5s x6s x7s x8s x9s x10s x11s x12s x13s
-  | map13 _ _ _ _ _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun map14 _ [] [] [] [] [] [] [] [] [] [] [] [] [] [] = []
-  | map14 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) (x8::x8s)
-      (x9::x9s) (x10::x10s) (x11::x11s) (x12::x12s) (x13::x13s) (x14::x14s) =
-    f x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13 x14 ::
-      map14 f x1s x2s x3s x4s x5s x6s x7s x8s x9s x10s x11s x12s x13s x14s
-  | map14 _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun fold_map4 _ [] [] [] [] acc = ([], acc)
-  | fold_map4 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) acc =
-    let
-      val (x, acc') = f x1 x2 x3 x4 acc;
-      val (xs, acc'') = fold_map4 f x1s x2s x3s x4s acc';
-    in (x :: xs, acc'') end
-  | fold_map4 _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun fold_map5 _ [] [] [] [] [] acc = ([], acc)
-  | fold_map5 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) acc =
-    let
-      val (x, acc') = f x1 x2 x3 x4 x5 acc;
-      val (xs, acc'') = fold_map5 f x1s x2s x3s x4s x5s acc';
-    in (x :: xs, acc'') end
-  | fold_map5 _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun fold_map6 _ [] [] [] [] [] [] acc = ([], acc)
-  | fold_map6 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) acc =
-    let
-      val (x, acc') = f x1 x2 x3 x4 x5 x6 acc;
-      val (xs, acc'') = fold_map6 f x1s x2s x3s x4s x5s x6s acc';
-    in (x :: xs, acc'') end
-  | fold_map6 _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun fold_map7 _ [] [] [] [] [] [] [] acc = ([], acc)
-  | fold_map7 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) acc =
-    let
-      val (x, acc') = f x1 x2 x3 x4 x5 x6 x7 acc;
-      val (xs, acc'') = fold_map7 f x1s x2s x3s x4s x5s x6s x7s acc';
-    in (x :: xs, acc'') end
-  | fold_map7 _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun fold_map8 _ [] [] [] [] [] [] [] [] acc = ([], acc)
-  | fold_map8 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) (x8::x8s)
-      acc =
-    let
-      val (x, acc') = f x1 x2 x3 x4 x5 x6 x7 x8 acc;
-      val (xs, acc'') = fold_map8 f x1s x2s x3s x4s x5s x6s x7s x8s acc';
-    in (x :: xs, acc'') end
-  | fold_map8 _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun fold_map9 _ [] [] [] [] [] [] [] [] [] acc = ([], acc)
-  | fold_map9 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) (x8::x8s)
-      (x9::x9s) acc =
-    let
-      val (x, acc') = f x1 x2 x3 x4 x5 x6 x7 x8 x9 acc;
-      val (xs, acc'') = fold_map9 f x1s x2s x3s x4s x5s x6s x7s x8s x9s acc';
-    in (x :: xs, acc'') end
-  | fold_map9 _ _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun split_list4 [] = ([], [], [], [])
-  | split_list4 ((x1, x2, x3, x4) :: xs) =
-    let val (xs1, xs2, xs3, xs4) = split_list4 xs;
-    in (x1 :: xs1, x2 :: xs2, x3 :: xs3, x4 :: xs4) end;
-
-fun split_list5 [] = ([], [], [], [], [])
-  | split_list5 ((x1, x2, x3, x4, x5) :: xs) =
-    let val (xs1, xs2, xs3, xs4, xs5) = split_list5 xs;
-    in (x1 :: xs1, x2 :: xs2, x3 :: xs3, x4 :: xs4, x5 :: xs5) end;
-
-val parse_binding_colon = parse_binding --| @{keyword ":"};
-val parse_opt_binding_colon = Scan.optional parse_binding_colon Binding.empty;
-
-val parse_type_arg_constrained =
-  Parse.type_ident -- Scan.option (@{keyword "::"} |-- Parse.!!! Parse.sort);
-
-val parse_type_arg_named_constrained =
-   (Parse.minus --| @{keyword ":"} >> K NONE || parse_opt_binding_colon >> SOME) --
-   parse_type_arg_constrained;
-
-val parse_type_args_named_constrained =
-  parse_type_arg_constrained >> (single o pair (SOME Binding.empty)) ||
-  @{keyword "("} |-- Parse.!!! (Parse.list1 parse_type_arg_named_constrained --| @{keyword ")"}) ||
-  Scan.succeed [];
-
-val parse_map_rel_binding = Parse.short_ident --| @{keyword ":"} -- parse_binding;
-
-val no_map_rel = (Binding.empty, Binding.empty);
-
-fun extract_map_rel ("map", b) = apfst (K b)
-  | extract_map_rel ("rel", b) = apsnd (K b)
-  | extract_map_rel (s, _) = error ("Unknown label " ^ quote s ^ " (expected \"map\" or \"rel\")");
-
-val parse_map_rel_bindings =
-  @{keyword "("} |-- Scan.repeat parse_map_rel_binding --| @{keyword ")"}
-    >> (fn ps => fold extract_map_rel ps no_map_rel) ||
-  Scan.succeed no_map_rel;
-
-
-(*TODO: is this really different from Typedef.add_typedef_global?*)
-fun typedef (b, Ts, mx) set opt_morphs tac lthy =
-  let
-    (*Work around loss of qualification in "typedef" axioms by replicating it in the name*)
-    val b' = fold_rev Binding.prefix_name (map (suffix "_" o fst) (#2 (Binding.dest b))) b;
-    val ((name, info), (lthy, lthy_old)) =
-      lthy
-      |> Typedef.add_typedef (b', Ts, mx) set opt_morphs tac
-      ||> `Local_Theory.restore;
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-  in
-    ((name, Typedef.transform_info phi info), lthy)
-  end;
-
-
-
-(* Term construction *)
-
-(** Fresh variables **)
-
-fun nonzero_string_of_int 0 = ""
-  | nonzero_string_of_int n = string_of_int n;
-
-val mk_TFreess = fold_map mk_TFrees;
-
-fun mk_Freesss x Tsss = fold_map2 mk_Freess (mk_names (length Tsss) x) Tsss;
-fun mk_Freessss x Tssss = fold_map2 mk_Freesss (mk_names (length Tssss) x) Tssss;
-
-fun retype_free T (Free (s, _)) = Free (s, T)
-  | retype_free _ t = raise TERM ("retype_free", [t]);
-
-
-(** Types **)
-
-(*stolen from ~~/src/HOL/Tools/Nitpick/nitpick_hol.ML*)
-fun num_binder_types (Type (@{type_name fun}, [_, T2])) =
-    1 + num_binder_types T2
-  | num_binder_types _ = 0
-
-(*maps [T1,...,Tn]--->T to ([T1,T2,...,Tn], T)*)
-fun strip_typeN 0 T = ([], T)
-  | strip_typeN n (Type (@{type_name fun}, [T, T'])) = strip_typeN (n - 1) T' |>> cons T
-  | strip_typeN _ T = raise TYPE ("strip_typeN", [T], []);
-
-(*maps [T1,...,Tn]--->T-->U to ([T1,T2,...,Tn], T-->U), where U is not a function type*)
-fun strip_fun_type T = strip_typeN (num_binder_types T - 1) T;
-
-val binder_fun_types = fst o strip_fun_type;
-val body_fun_type = snd o strip_fun_type;
-
-fun mk_pred2T T U = mk_predT [T, U];
-val mk_relT = HOLogic.mk_setT o HOLogic.mk_prodT;
-val dest_relT = HOLogic.dest_prodT o HOLogic.dest_setT;
-val dest_pred2T = apsnd Term.domain_type o Term.dest_funT;
-fun mk_sumT (LT, RT) = Type (@{type_name Sum_Type.sum}, [LT, RT]);
-
-
-(** Constants **)
-
-fun fst_const T = Const (@{const_name fst}, T --> fst (HOLogic.dest_prodT T));
-fun snd_const T = Const (@{const_name snd}, T --> snd (HOLogic.dest_prodT T));
-fun Id_const T = Const (@{const_name Id}, mk_relT (T, T));
-
-
-(** Operators **)
-
-fun mk_converse R =
-  let
-    val RT = dest_relT (fastype_of R);
-    val RST = mk_relT (snd RT, fst RT);
-  in Const (@{const_name converse}, fastype_of R --> RST) $ R end;
-
-fun mk_rel_comp (R, S) =
-  let
-    val RT = fastype_of R;
-    val ST = fastype_of S;
-    val RST = mk_relT (fst (dest_relT RT), snd (dest_relT ST));
-  in Const (@{const_name relcomp}, RT --> ST --> RST) $ R $ S end;
-
-fun mk_Gr A f =
-  let val ((AT, BT), FT) = `dest_funT (fastype_of f);
-  in Const (@{const_name Gr}, HOLogic.mk_setT AT --> FT --> mk_relT (AT, BT)) $ A $ f end;
-
-fun mk_conversep R =
-  let
-    val RT = dest_pred2T (fastype_of R);
-    val RST = mk_pred2T (snd RT) (fst RT);
-  in Const (@{const_name conversep}, fastype_of R --> RST) $ R end;
-
-fun mk_rel_compp (R, S) =
-  let
-    val RT = fastype_of R;
-    val ST = fastype_of S;
-    val RST = mk_pred2T (fst (dest_pred2T RT)) (snd (dest_pred2T ST));
-  in Const (@{const_name relcompp}, RT --> ST --> RST) $ R $ S end;
-
-fun mk_Grp A f =
-  let val ((AT, BT), FT) = `dest_funT (fastype_of f);
-  in Const (@{const_name Grp}, HOLogic.mk_setT AT --> FT --> mk_pred2T AT BT) $ A $ f end;
-
-fun mk_image f =
-  let val (T, U) = dest_funT (fastype_of f);
-  in Const (@{const_name image},
-    (T --> U) --> (HOLogic.mk_setT T) --> (HOLogic.mk_setT U)) $ f end;
-
-fun mk_Ball X f =
-  Const (@{const_name Ball}, fastype_of X --> fastype_of f --> HOLogic.boolT) $ X $ f;
-
-fun mk_Bex X f =
-  Const (@{const_name Bex}, fastype_of X --> fastype_of f --> HOLogic.boolT) $ X $ f;
-
-fun mk_UNION X f =
-  let val (T, U) = dest_funT (fastype_of f);
-  in Const (@{const_name SUPR}, fastype_of X --> (T --> U) --> U) $ X $ f end;
-
-fun mk_Union T =
-  Const (@{const_name Sup}, HOLogic.mk_setT (HOLogic.mk_setT T) --> HOLogic.mk_setT T);
-
-fun mk_Field r =
-  let val T = fst (dest_relT (fastype_of r));
-  in Const (@{const_name Field}, mk_relT (T, T) --> HOLogic.mk_setT T) $ r end;
-
-fun mk_card_order bd =
-  let
-    val T = fastype_of bd;
-    val AT = fst (dest_relT T);
-  in
-    Const (@{const_name card_order_on}, HOLogic.mk_setT AT --> T --> HOLogic.boolT) $
-      (HOLogic.mk_UNIV AT) $ bd
-  end;
-
-fun mk_Card_order bd =
-  let
-    val T = fastype_of bd;
-    val AT = fst (dest_relT T);
-  in
-    Const (@{const_name card_order_on}, HOLogic.mk_setT AT --> T --> HOLogic.boolT) $
-      mk_Field bd $ bd
-  end;
-
-fun mk_cinfinite bd =
-  Const (@{const_name cinfinite}, fastype_of bd --> HOLogic.boolT) $ bd;
-
-fun mk_ordLeq t1 t2 =
-  HOLogic.mk_mem (HOLogic.mk_prod (t1, t2),
-    Const (@{const_name ordLeq}, mk_relT (fastype_of t1, fastype_of t2)));
-
-fun mk_card_of A =
-  let
-    val AT = fastype_of A;
-    val T = HOLogic.dest_setT AT;
-  in
-    Const (@{const_name card_of}, AT --> mk_relT (T, T)) $ A
-  end;
-
-fun mk_dir_image r f =
-  let val (T, U) = dest_funT (fastype_of f);
-  in Const (@{const_name dir_image}, mk_relT (T, T) --> (T --> U) --> mk_relT (U, U)) $ r $ f end;
-
-fun mk_fun_rel R S =
-  let
-    val ((RA, RB), RT) = `dest_pred2T (fastype_of R);
-    val ((SA, SB), ST) = `dest_pred2T (fastype_of S);
-  in Const (@{const_name fun_rel}, RT --> ST --> mk_pred2T (RA --> SA) (RB --> SB)) $ R $ S end;
-
-(*FIXME: "x"?*)
-(*(nth sets i) must be of type "T --> 'ai set"*)
-fun mk_in As sets T =
-  let
-    fun in_single set A =
-      let val AT = fastype_of A;
-      in Const (@{const_name less_eq},
-        AT --> AT --> HOLogic.boolT) $ (set $ Free ("x", T)) $ A end;
-  in
-    if length sets > 0
-    then HOLogic.mk_Collect ("x", T, foldr1 (HOLogic.mk_conj) (map2 in_single sets As))
-    else HOLogic.mk_UNIV T
-  end;
-
-fun mk_leq t1 t2 =
-  Const (@{const_name less_eq}, (fastype_of t1) --> (fastype_of t2) --> HOLogic.boolT) $ t1 $ t2;
-
-fun mk_card_binop binop typop t1 t2 =
-  let
-    val (T1, relT1) = `(fst o dest_relT) (fastype_of t1);
-    val (T2, relT2) = `(fst o dest_relT) (fastype_of t2);
-  in
-    Const (binop, relT1 --> relT2 --> mk_relT (typop (T1, T2), typop (T1, T2))) $ t1 $ t2
-  end;
-
-val mk_csum = mk_card_binop @{const_name csum} mk_sumT;
-val mk_cprod = mk_card_binop @{const_name cprod} HOLogic.mk_prodT;
-val mk_cexp = mk_card_binop @{const_name cexp} (op --> o swap);
-val ctwo = @{term ctwo};
-
-fun mk_collect xs defT =
-  let val T = (case xs of [] => defT | (x::_) => fastype_of x);
-  in Const (@{const_name collect}, HOLogic.mk_setT T --> T) $ (HOLogic.mk_set T xs) end;
-
-fun find_indices eq xs ys = map_filter I
-  (map_index (fn (i, y) => if member eq xs y then SOME i else NONE) ys);
-
-fun mk_trans thm1 thm2 = trans OF [thm1, thm2];
-fun mk_sym thm = thm RS sym;
-
-(*TODO: antiquote heavily used theorems once*)
-val Pair_eqD = @{thm iffD1[OF Pair_eq]};
-val Pair_eqI = @{thm iffD2[OF Pair_eq]};
-val ctrans = @{thm ordLeq_transitive};
-val id_apply = @{thm id_apply};
-val meta_mp = @{thm meta_mp};
-val meta_spec = @{thm meta_spec};
-val o_apply = @{thm o_apply};
-val set_mp = @{thm set_mp};
-val set_rev_mp = @{thm set_rev_mp};
-val subset_UNIV = @{thm subset_UNIV};
-
-fun mk_nthN 1 t 1 = t
-  | mk_nthN _ t 1 = HOLogic.mk_fst t
-  | mk_nthN 2 t 2 = HOLogic.mk_snd t
-  | mk_nthN n t m = mk_nthN (n - 1) (HOLogic.mk_snd t) (m - 1);
-
-fun mk_nth_conv n m =
-  let
-    fun thm b = if b then @{thm fstI} else @{thm sndI}
-    fun mk_nth_conv _ 1 1 = refl
-      | mk_nth_conv _ _ 1 = @{thm fst_conv}
-      | mk_nth_conv _ 2 2 = @{thm snd_conv}
-      | mk_nth_conv b _ 2 = @{thm snd_conv} RS thm b
-      | mk_nth_conv b n m = mk_nth_conv false (n - 1) (m - 1) RS thm b;
-  in mk_nth_conv (not (m = n)) n m end;
-
-fun mk_nthI 1 1 = @{thm TrueE[OF TrueI]}
-  | mk_nthI n m = fold (curry op RS) (replicate (m - 1) @{thm sndI})
-    (if m = n then @{thm TrueE[OF TrueI]} else @{thm fstI});
-
-fun mk_conjunctN 1 1 = @{thm TrueE[OF TrueI]}
-  | mk_conjunctN _ 1 = conjunct1
-  | mk_conjunctN 2 2 = conjunct2
-  | mk_conjunctN n m = conjunct2 RS (mk_conjunctN (n - 1) (m - 1));
-
-fun conj_dests n thm = map (fn k => thm RS mk_conjunctN n k) (1 upto n);
-
-fun mk_conjIN 1 = @{thm TrueE[OF TrueI]}
-  | mk_conjIN n = mk_conjIN (n - 1) RSN (2, conjI);
-
-fun mk_ordLeq_csum 1 1 thm = thm
-  | mk_ordLeq_csum _ 1 thm = @{thm ordLeq_transitive} OF [thm, @{thm ordLeq_csum1}]
-  | mk_ordLeq_csum 2 2 thm = @{thm ordLeq_transitive} OF [thm, @{thm ordLeq_csum2}]
-  | mk_ordLeq_csum n m thm = @{thm ordLeq_transitive} OF
-    [mk_ordLeq_csum (n - 1) (m - 1) thm, @{thm ordLeq_csum2[OF Card_order_csum]}];
-
-local
-  fun mk_Un_upper' 0 = subset_refl
-    | mk_Un_upper' 1 = @{thm Un_upper1}
-    | mk_Un_upper' k = Library.foldr (op RS o swap)
-      (replicate (k - 1) @{thm subset_trans[OF Un_upper1]}, @{thm Un_upper1});
-in
-  fun mk_Un_upper 1 1 = subset_refl
-    | mk_Un_upper n 1 = mk_Un_upper' (n - 2) RS @{thm subset_trans[OF Un_upper1]}
-    | mk_Un_upper n m = mk_Un_upper' (n - m) RS @{thm subset_trans[OF Un_upper2]};
-end;
-
-local
-  fun mk_UnIN' 0 = @{thm UnI2}
-    | mk_UnIN' m = mk_UnIN' (m - 1) RS @{thm UnI1};
-in
-  fun mk_UnIN 1 1 = @{thm TrueE[OF TrueI]}
-    | mk_UnIN n 1 = Library.foldr1 (op RS o swap) (replicate (n - 1) @{thm UnI1})
-    | mk_UnIN n m = mk_UnIN' (n - m)
-end;
-
-fun is_refl_prop t =
-  op aconv (HOLogic.dest_eq (HOLogic.dest_Trueprop t))
-  handle TERM _ => false;
-
-val is_refl = is_refl_prop o Thm.prop_of;
-val is_concl_refl = is_refl_prop o Logic.strip_imp_concl o Thm.prop_of;
-
-val no_refl = filter_out is_refl;
-val no_reflexive = filter_out Thm.is_reflexive;
-
-fun fold_thms ctxt thms = Local_Defs.fold ctxt (distinct Thm.eq_thm_prop thms);
-
-end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_comp.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,704 @@
+(*  Title:      HOL/BNF/Tools/bnf_comp.ML
+    Author:     Dmitriy Traytel, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2012
+
+Composition of bounded natural functors.
+*)
+
+signature BNF_COMP =
+sig
+  val ID_bnf: BNF_Def.bnf
+  val DEADID_bnf: BNF_Def.bnf
+
+  type unfold_set
+  val empty_unfolds: unfold_set
+
+  exception BAD_DEAD of typ * typ
+
+  val bnf_of_typ: BNF_Def.const_policy -> (binding -> binding) ->
+    ((string * sort) list list -> (string * sort) list) -> (string * sort) list -> typ ->
+    unfold_set * Proof.context ->
+    (BNF_Def.bnf * (typ list * typ list)) * (unfold_set * Proof.context)
+  val default_comp_sort: (string * sort) list list -> (string * sort) list
+  val normalize_bnfs: (int -> binding -> binding) -> ''a list list -> ''a list ->
+    (''a list list -> ''a list) -> BNF_Def.bnf list -> unfold_set -> Proof.context ->
+    (int list list * ''a list) * (BNF_Def.bnf list * (unfold_set * Proof.context))
+  val seal_bnf: (binding -> binding) -> unfold_set -> binding -> typ list -> BNF_Def.bnf ->
+    Proof.context -> (BNF_Def.bnf * typ list) * local_theory
+end;
+
+structure BNF_Comp : BNF_COMP =
+struct
+
+open BNF_Def
+open BNF_Util
+open BNF_Tactics
+open BNF_Comp_Tactics
+
+val ID_bnf = the (bnf_of @{context} "Basic_BNFs.ID");
+val DEADID_bnf = the (bnf_of @{context} "Basic_BNFs.DEADID");
+
+(* TODO: Replace by "BNF_Defs.defs list" *)
+type unfold_set = {
+  map_unfolds: thm list,
+  set_unfoldss: thm list list,
+  rel_unfolds: thm list
+};
+
+val empty_unfolds = {map_unfolds = [], set_unfoldss = [], rel_unfolds = []};
+
+fun add_to_thms thms new = thms |> not (Thm.is_reflexive new) ? insert Thm.eq_thm new;
+fun adds_to_thms thms news = insert (eq_set Thm.eq_thm) (no_reflexive news) thms;
+
+fun add_to_unfolds map sets rel
+  {map_unfolds, set_unfoldss, rel_unfolds} =
+  {map_unfolds = add_to_thms map_unfolds map,
+    set_unfoldss = adds_to_thms set_unfoldss sets,
+    rel_unfolds = add_to_thms rel_unfolds rel};
+
+fun add_bnf_to_unfolds bnf =
+  add_to_unfolds (map_def_of_bnf bnf) (set_defs_of_bnf bnf) (rel_def_of_bnf bnf);
+
+val bdTN = "bdT";
+
+fun mk_killN n = "_kill" ^ string_of_int n;
+fun mk_liftN n = "_lift" ^ string_of_int n;
+fun mk_permuteN src dest =
+  "_permute_" ^ implode (map string_of_int src) ^ "_" ^ implode (map string_of_int dest);
+
+(*copied from Envir.expand_term_free*)
+fun expand_term_const defs =
+  let
+    val eqs = map ((fn ((x, U), u) => (x, (U, u))) o apfst dest_Const) defs;
+    val get = fn Const (x, _) => AList.lookup (op =) eqs x | _ => NONE;
+  in Envir.expand_term get end;
+
+fun clean_compose_bnf const_policy qualify b outer inners (unfold_set, lthy) =
+  let
+    val olive = live_of_bnf outer;
+    val onwits = nwits_of_bnf outer;
+    val odead = dead_of_bnf outer;
+    val inner = hd inners;
+    val ilive = live_of_bnf inner;
+    val ideads = map dead_of_bnf inners;
+    val inwitss = map nwits_of_bnf inners;
+
+    (* TODO: check olive = length inners > 0,
+                   forall inner from inners. ilive = live,
+                   forall inner from inners. idead = dead  *)
+
+    val (oDs, lthy1) = apfst (map TFree)
+      (Variable.invent_types (replicate odead HOLogic.typeS) lthy);
+    val (Dss, lthy2) = apfst (map (map TFree))
+        (fold_map Variable.invent_types (map (fn n => replicate n HOLogic.typeS) ideads) lthy1);
+    val (Ass, lthy3) = apfst (replicate ilive o map TFree)
+      (Variable.invent_types (replicate ilive HOLogic.typeS) lthy2);
+    val As = if ilive > 0 then hd Ass else [];
+    val Ass_repl = replicate olive As;
+    val (Bs, _(*lthy4*)) = apfst (map TFree)
+      (Variable.invent_types (replicate ilive HOLogic.typeS) lthy3);
+    val Bss_repl = replicate olive Bs;
+
+    val ((((fs', Qs'), Asets), xs), _(*names_lthy*)) = lthy
+      |> apfst snd o mk_Frees' "f" (map2 (curry op -->) As Bs)
+      ||>> apfst snd o mk_Frees' "Q" (map2 mk_pred2T As Bs)
+      ||>> mk_Frees "A" (map HOLogic.mk_setT As)
+      ||>> mk_Frees "x" As;
+
+    val CAs = map3 mk_T_of_bnf Dss Ass_repl inners;
+    val CCA = mk_T_of_bnf oDs CAs outer;
+    val CBs = map3 mk_T_of_bnf Dss Bss_repl inners;
+    val outer_sets = mk_sets_of_bnf (replicate olive oDs) (replicate olive CAs) outer;
+    val inner_setss = map3 mk_sets_of_bnf (map (replicate ilive) Dss) (replicate olive Ass) inners;
+    val inner_bds = map3 mk_bd_of_bnf Dss Ass_repl inners;
+    val outer_bd = mk_bd_of_bnf oDs CAs outer;
+
+    (*%f1 ... fn. outer.map (inner_1.map f1 ... fn) ... (inner_m.map f1 ... fn)*)
+    val mapx = fold_rev Term.abs fs'
+      (Term.list_comb (mk_map_of_bnf oDs CAs CBs outer,
+        map2 (fn Ds => (fn f => Term.list_comb (f, map Bound (ilive - 1 downto 0))) o
+          mk_map_of_bnf Ds As Bs) Dss inners));
+    (*%Q1 ... Qn. outer.rel (inner_1.rel Q1 ... Qn) ... (inner_m.rel Q1 ... Qn)*)
+    val rel = fold_rev Term.abs Qs'
+      (Term.list_comb (mk_rel_of_bnf oDs CAs CBs outer,
+        map2 (fn Ds => (fn f => Term.list_comb (f, map Bound (ilive - 1 downto 0))) o
+          mk_rel_of_bnf Ds As Bs) Dss inners));
+
+    (*Union o collect {outer.set_1 ... outer.set_m} o outer.map inner_1.set_i ... inner_m.set_i*)
+    (*Union o collect {image inner_1.set_i o outer.set_1 ... image inner_m.set_i o outer.set_m}*)
+    fun mk_set i =
+      let
+        val (setTs, T) = `(replicate olive o HOLogic.mk_setT) (nth As i);
+        val outer_set = mk_collect
+          (mk_sets_of_bnf (replicate olive oDs) (replicate olive setTs) outer)
+          (mk_T_of_bnf oDs setTs outer --> HOLogic.mk_setT T);
+        val inner_sets = map (fn sets => nth sets i) inner_setss;
+        val outer_map = mk_map_of_bnf oDs CAs setTs outer;
+        val map_inner_sets = Term.list_comb (outer_map, inner_sets);
+        val collect_image = mk_collect
+          (map2 (fn f => fn set => HOLogic.mk_comp (mk_image f, set)) inner_sets outer_sets)
+          (CCA --> HOLogic.mk_setT T);
+      in
+        (Library.foldl1 HOLogic.mk_comp [mk_Union T, outer_set, map_inner_sets],
+        HOLogic.mk_comp (mk_Union T, collect_image))
+      end;
+
+    val (sets, sets_alt) = map_split mk_set (0 upto ilive - 1);
+
+    (*(inner_1.bd +c ... +c inner_m.bd) *c outer.bd*)
+    val bd = mk_cprod (Library.foldr1 (uncurry mk_csum) inner_bds) outer_bd;
+
+    fun map_id0_tac _ =
+      mk_comp_map_id0_tac (map_id0_of_bnf outer) (map_cong0_of_bnf outer)
+        (map map_id0_of_bnf inners);
+
+    fun map_comp0_tac _ =
+      mk_comp_map_comp0_tac (map_comp0_of_bnf outer) (map_cong0_of_bnf outer)
+        (map map_comp0_of_bnf inners);
+
+    fun mk_single_set_map0_tac i _ =
+      mk_comp_set_map0_tac (map_comp0_of_bnf outer) (map_cong0_of_bnf outer)
+        (collect_set_map_of_bnf outer)
+        (map ((fn thms => nth thms i) o set_map0_of_bnf) inners);
+
+    val set_map0_tacs = map mk_single_set_map0_tac (0 upto ilive - 1);
+
+    fun bd_card_order_tac _ =
+      mk_comp_bd_card_order_tac (map bd_card_order_of_bnf inners) (bd_card_order_of_bnf outer);
+
+    fun bd_cinfinite_tac _ =
+      mk_comp_bd_cinfinite_tac (bd_cinfinite_of_bnf inner) (bd_cinfinite_of_bnf outer);
+
+    val set_alt_thms =
+      if Config.get lthy quick_and_dirty then
+        []
+      else
+        map (fn goal =>
+          Goal.prove_sorry lthy [] [] goal
+            (fn {context = ctxt, prems = _} =>
+              mk_comp_set_alt_tac ctxt (collect_set_map_of_bnf outer))
+          |> Thm.close_derivation)
+        (map2 (curry (HOLogic.mk_Trueprop o HOLogic.mk_eq)) sets sets_alt);
+
+    fun map_cong0_tac _ =
+      mk_comp_map_cong0_tac set_alt_thms (map_cong0_of_bnf outer) (map map_cong0_of_bnf inners);
+
+    val set_bd_tacs =
+      if Config.get lthy quick_and_dirty then
+        replicate ilive (K all_tac)
+      else
+        let
+          val outer_set_bds = set_bd_of_bnf outer;
+          val inner_set_bdss = map set_bd_of_bnf inners;
+          val inner_bd_Card_orders = map bd_Card_order_of_bnf inners;
+          fun single_set_bd_thm i j =
+            @{thm comp_single_set_bd} OF [nth inner_bd_Card_orders j, nth (nth inner_set_bdss j) i,
+              nth outer_set_bds j]
+          val single_set_bd_thmss =
+            map ((fn f => map f (0 upto olive - 1)) o single_set_bd_thm) (0 upto ilive - 1);
+        in
+          map2 (fn set_alt => fn single_set_bds => fn {context = ctxt, prems = _} =>
+            mk_comp_set_bd_tac ctxt set_alt single_set_bds)
+          set_alt_thms single_set_bd_thmss
+        end;
+
+    val in_alt_thm =
+      let
+        val inx = mk_in Asets sets CCA;
+        val in_alt = mk_in (map2 (mk_in Asets) inner_setss CAs) outer_sets CCA;
+        val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));
+      in
+        Goal.prove_sorry lthy [] [] goal
+          (fn {context = ctxt, prems = _} => mk_comp_in_alt_tac ctxt set_alt_thms)
+        |> Thm.close_derivation
+      end;
+
+    fun le_rel_OO_tac _ = mk_le_rel_OO_tac (le_rel_OO_of_bnf outer) (rel_mono_of_bnf outer)
+      (map le_rel_OO_of_bnf inners);
+
+    fun rel_OO_Grp_tac _ =
+      let
+        val outer_rel_Grp = rel_Grp_of_bnf outer RS sym;
+        val outer_rel_cong = rel_cong_of_bnf outer;
+        val thm =
+          (trans OF [in_alt_thm RS @{thm OO_Grp_cong},
+             trans OF [@{thm arg_cong2[of _ _ _ _ relcompp]} OF
+               [trans OF [outer_rel_Grp RS @{thm arg_cong[of _ _ conversep]},
+                 rel_conversep_of_bnf outer RS sym], outer_rel_Grp],
+               trans OF [rel_OO_of_bnf outer RS sym, outer_rel_cong OF
+                 (map (fn bnf => rel_OO_Grp_of_bnf bnf RS sym) inners)]]] RS sym)
+          (*|> unfold_thms lthy (rel_def_of_bnf outer :: map rel_def_of_bnf inners)*);
+      in
+        rtac thm 1
+      end;
+
+    val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac
+      bd_cinfinite_tac set_bd_tacs le_rel_OO_tac rel_OO_Grp_tac;
+
+    val outer_wits = mk_wits_of_bnf (replicate onwits oDs) (replicate onwits CAs) outer;
+
+    val inner_witss = map (map (fn (I, wit) => Term.list_comb (wit, map (nth xs) I)))
+      (map3 (fn Ds => fn n => mk_wits_of_bnf (replicate n Ds) (replicate n As))
+        Dss inwitss inners);
+
+    val inner_witsss = map (map (nth inner_witss) o fst) outer_wits;
+
+    val wits = (inner_witsss, (map (single o snd) outer_wits))
+      |-> map2 (fold (map_product (fn iwit => fn owit => owit $ iwit)))
+      |> flat
+      |> map (`(fn t => Term.add_frees t []))
+      |> minimize_wits
+      |> map (fn (frees, t) => fold absfree frees t);
+
+    fun wit_tac {context = ctxt, prems = _} =
+      mk_comp_wit_tac ctxt (wit_thms_of_bnf outer) (collect_set_map_of_bnf outer)
+        (maps wit_thms_of_bnf inners);
+
+    val (bnf', lthy') =
+      bnf_def const_policy (K Dont_Note) qualify tacs wit_tac (SOME (oDs @ flat Dss)) Binding.empty
+        Binding.empty [] ((((((b, CCA), mapx), sets), bd), wits), SOME rel) lthy;
+  in
+    (bnf', (add_bnf_to_unfolds bnf' unfold_set, lthy'))
+  end;
+
+(* Killing live variables *)
+
+fun kill_bnf qualify n bnf (unfold_set, lthy) = if n = 0 then (bnf, (unfold_set, lthy)) else
+  let
+    val b = Binding.suffix_name (mk_killN n) (name_of_bnf bnf);
+    val live = live_of_bnf bnf;
+    val dead = dead_of_bnf bnf;
+    val nwits = nwits_of_bnf bnf;
+
+    (* TODO: check 0 < n <= live *)
+
+    val (Ds, lthy1) = apfst (map TFree)
+      (Variable.invent_types (replicate dead HOLogic.typeS) lthy);
+    val ((killedAs, As), lthy2) = apfst (`(take n) o map TFree)
+      (Variable.invent_types (replicate live HOLogic.typeS) lthy1);
+    val (Bs, _(*lthy3*)) = apfst (append killedAs o map TFree)
+      (Variable.invent_types (replicate (live - n) HOLogic.typeS) lthy2);
+
+    val ((Asets, lives), _(*names_lthy*)) = lthy
+      |> mk_Frees "A" (map HOLogic.mk_setT (drop n As))
+      ||>> mk_Frees "x" (drop n As);
+    val xs = map (fn T => HOLogic.choice_const T $ absdummy T @{term True}) killedAs @ lives;
+
+    val T = mk_T_of_bnf Ds As bnf;
+
+    (*bnf.map id ... id*)
+    val mapx = Term.list_comb (mk_map_of_bnf Ds As Bs bnf, map HOLogic.id_const killedAs);
+    (*bnf.rel (op =) ... (op =)*)
+    val rel = Term.list_comb (mk_rel_of_bnf Ds As Bs bnf, map HOLogic.eq_const killedAs);
+
+    val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf;
+    val sets = drop n bnf_sets;
+
+    (*(|UNIV :: A1 set| +c ... +c |UNIV :: An set|) *c bnf.bd*)
+    val bnf_bd = mk_bd_of_bnf Ds As bnf;
+    val bd = mk_cprod
+      (Library.foldr1 (uncurry mk_csum) (map (mk_card_of o HOLogic.mk_UNIV) killedAs)) bnf_bd;
+
+    fun map_id0_tac _ = rtac (map_id0_of_bnf bnf) 1;
+    fun map_comp0_tac {context = ctxt, prems = _} =
+      unfold_thms_tac ctxt ((map_comp0_of_bnf bnf RS sym) :: @{thms o_assoc id_o o_id}) THEN
+      rtac refl 1;
+    fun map_cong0_tac {context = ctxt, prems = _} =
+      mk_kill_map_cong0_tac ctxt n (live - n) (map_cong0_of_bnf bnf);
+    val set_map0_tacs = map (fn thm => fn _ => rtac thm 1) (drop n (set_map0_of_bnf bnf));
+    fun bd_card_order_tac _ = mk_kill_bd_card_order_tac n (bd_card_order_of_bnf bnf);
+    fun bd_cinfinite_tac _ = mk_kill_bd_cinfinite_tac (bd_Cinfinite_of_bnf bnf);
+    val set_bd_tacs =
+      map (fn thm => fn _ => mk_kill_set_bd_tac (bd_Card_order_of_bnf bnf) thm)
+        (drop n (set_bd_of_bnf bnf));
+
+    val in_alt_thm =
+      let
+        val inx = mk_in Asets sets T;
+        val in_alt = mk_in (map HOLogic.mk_UNIV killedAs @ Asets) bnf_sets T;
+        val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));
+      in
+        Goal.prove_sorry lthy [] [] goal (K kill_in_alt_tac) |> Thm.close_derivation
+      end;
+
+    fun le_rel_OO_tac {context = ctxt, prems = _} =
+      EVERY' [rtac @{thm ord_le_eq_trans}, rtac (le_rel_OO_of_bnf bnf)] 1 THEN
+      unfold_thms_tac ctxt @{thms eq_OO} THEN rtac refl 1;
+
+    fun rel_OO_Grp_tac _ =
+      let
+        val rel_Grp = rel_Grp_of_bnf bnf RS sym
+        val thm =
+          (trans OF [in_alt_thm RS @{thm OO_Grp_cong},
+            trans OF [@{thm arg_cong2[of _ _ _ _ relcompp]} OF
+              [trans OF [rel_Grp RS @{thm arg_cong[of _ _ conversep]},
+                rel_conversep_of_bnf bnf RS sym], rel_Grp],
+              trans OF [rel_OO_of_bnf bnf RS sym, rel_cong_of_bnf bnf OF
+                (replicate n @{thm trans[OF Grp_UNIV_id[OF refl] eq_alt[symmetric]]} @
+                 replicate (live - n) @{thm Grp_fst_snd})]]] RS sym);
+      in
+        rtac thm 1
+      end;
+
+    val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac
+      bd_cinfinite_tac set_bd_tacs le_rel_OO_tac rel_OO_Grp_tac;
+
+    val bnf_wits = mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf;
+
+    val wits = map (fn t => fold absfree (Term.add_frees t []) t)
+      (map (fn (I, wit) => Term.list_comb (wit, map (nth xs) I)) bnf_wits);
+
+    fun wit_tac _ = mk_simple_wit_tac (wit_thms_of_bnf bnf);
+
+    val (bnf', lthy') =
+      bnf_def Smart_Inline (K Dont_Note) qualify tacs wit_tac (SOME (killedAs @ Ds)) Binding.empty
+        Binding.empty [] ((((((b, T), mapx), sets), bd), wits), SOME rel) lthy;
+  in
+    (bnf', (add_bnf_to_unfolds bnf' unfold_set, lthy'))
+  end;
+
+(* Adding dummy live variables *)
+
+fun lift_bnf qualify n bnf (unfold_set, lthy) = if n = 0 then (bnf, (unfold_set, lthy)) else
+  let
+    val b = Binding.suffix_name (mk_liftN n) (name_of_bnf bnf);
+    val live = live_of_bnf bnf;
+    val dead = dead_of_bnf bnf;
+    val nwits = nwits_of_bnf bnf;
+
+    (* TODO: check 0 < n *)
+
+    val (Ds, lthy1) = apfst (map TFree)
+      (Variable.invent_types (replicate dead HOLogic.typeS) lthy);
+    val ((newAs, As), lthy2) = apfst (chop n o map TFree)
+      (Variable.invent_types (replicate (n + live) HOLogic.typeS) lthy1);
+    val ((newBs, Bs), _(*lthy3*)) = apfst (chop n o map TFree)
+      (Variable.invent_types (replicate (n + live) HOLogic.typeS) lthy2);
+
+    val (Asets, _(*names_lthy*)) = lthy
+      |> mk_Frees "A" (map HOLogic.mk_setT (newAs @ As));
+
+    val T = mk_T_of_bnf Ds As bnf;
+
+    (*%f1 ... fn. bnf.map*)
+    val mapx =
+      fold_rev Term.absdummy (map2 (curry op -->) newAs newBs) (mk_map_of_bnf Ds As Bs bnf);
+    (*%Q1 ... Qn. bnf.rel*)
+    val rel = fold_rev Term.absdummy (map2 mk_pred2T newAs newBs) (mk_rel_of_bnf Ds As Bs bnf);
+
+    val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf;
+    val sets = map (fn A => absdummy T (HOLogic.mk_set A [])) newAs @ bnf_sets;
+
+    val bd = mk_bd_of_bnf Ds As bnf;
+
+    fun map_id0_tac _ = rtac (map_id0_of_bnf bnf) 1;
+    fun map_comp0_tac {context = ctxt, prems = _} =
+      unfold_thms_tac ctxt ((map_comp0_of_bnf bnf RS sym) :: @{thms o_assoc id_o o_id}) THEN
+      rtac refl 1;
+    fun map_cong0_tac {context = ctxt, prems = _} =
+      rtac (map_cong0_of_bnf bnf) 1 THEN REPEAT_DETERM_N live (Goal.assume_rule_tac ctxt 1);
+    val set_map0_tacs =
+      if Config.get lthy quick_and_dirty then
+        replicate (n + live) (K all_tac)
+      else
+        replicate n (K empty_natural_tac) @
+        map (fn thm => fn _ => rtac thm 1) (set_map0_of_bnf bnf);
+    fun bd_card_order_tac _ = rtac (bd_card_order_of_bnf bnf) 1;
+    fun bd_cinfinite_tac _ = rtac (bd_cinfinite_of_bnf bnf) 1;
+    val set_bd_tacs =
+      if Config.get lthy quick_and_dirty then
+        replicate (n + live) (K all_tac)
+      else
+        replicate n (K (mk_lift_set_bd_tac (bd_Card_order_of_bnf bnf))) @
+        (map (fn thm => fn _ => rtac thm 1) (set_bd_of_bnf bnf));
+
+    val in_alt_thm =
+      let
+        val inx = mk_in Asets sets T;
+        val in_alt = mk_in (drop n Asets) bnf_sets T;
+        val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));
+      in
+        Goal.prove_sorry lthy [] [] goal (K lift_in_alt_tac) |> Thm.close_derivation
+      end;
+
+    fun le_rel_OO_tac _ = rtac (le_rel_OO_of_bnf bnf) 1;
+
+    fun rel_OO_Grp_tac _ = mk_simple_rel_OO_Grp_tac (rel_OO_Grp_of_bnf bnf) in_alt_thm;
+
+    val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac
+      bd_cinfinite_tac set_bd_tacs le_rel_OO_tac rel_OO_Grp_tac;
+
+    val wits = map snd (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf);
+
+    fun wit_tac _ = mk_simple_wit_tac (wit_thms_of_bnf bnf);
+
+    val (bnf', lthy') =
+      bnf_def Smart_Inline (K Dont_Note) qualify tacs wit_tac (SOME Ds) Binding.empty Binding.empty
+        [] ((((((b, T), mapx), sets), bd), wits), SOME rel) lthy;
+  in
+    (bnf', (add_bnf_to_unfolds bnf' unfold_set, lthy'))
+  end;
+
+(* Changing the order of live variables *)
+
+fun permute_bnf qualify src dest bnf (unfold_set, lthy) =
+  if src = dest then (bnf, (unfold_set, lthy)) else
+  let
+    val b = Binding.suffix_name (mk_permuteN src dest) (name_of_bnf bnf);
+    val live = live_of_bnf bnf;
+    val dead = dead_of_bnf bnf;
+    val nwits = nwits_of_bnf bnf;
+    fun permute xs = permute_like (op =) src dest xs;
+    fun unpermute xs = permute_like (op =) dest src xs;
+
+    val (Ds, lthy1) = apfst (map TFree)
+      (Variable.invent_types (replicate dead HOLogic.typeS) lthy);
+    val (As, lthy2) = apfst (map TFree)
+      (Variable.invent_types (replicate live HOLogic.typeS) lthy1);
+    val (Bs, _(*lthy3*)) = apfst (map TFree)
+      (Variable.invent_types (replicate live HOLogic.typeS) lthy2);
+
+    val (Asets, _(*names_lthy*)) = lthy
+      |> mk_Frees "A" (map HOLogic.mk_setT (permute As));
+
+    val T = mk_T_of_bnf Ds As bnf;
+
+    (*%f(1) ... f(n). bnf.map f\<sigma>(1) ... f\<sigma>(n)*)
+    val mapx = fold_rev Term.absdummy (permute (map2 (curry op -->) As Bs))
+      (Term.list_comb (mk_map_of_bnf Ds As Bs bnf, unpermute (map Bound (live - 1 downto 0))));
+    (*%Q(1) ... Q(n). bnf.rel Q\<sigma>(1) ... Q\<sigma>(n)*)
+    val rel = fold_rev Term.absdummy (permute (map2 mk_pred2T As Bs))
+      (Term.list_comb (mk_rel_of_bnf Ds As Bs bnf, unpermute (map Bound (live - 1 downto 0))));
+
+    val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf;
+    val sets = permute bnf_sets;
+
+    val bd = mk_bd_of_bnf Ds As bnf;
+
+    fun map_id0_tac _ = rtac (map_id0_of_bnf bnf) 1;
+    fun map_comp0_tac _ = rtac (map_comp0_of_bnf bnf) 1;
+    fun map_cong0_tac {context = ctxt, prems = _} =
+      rtac (map_cong0_of_bnf bnf) 1 THEN REPEAT_DETERM_N live (Goal.assume_rule_tac ctxt 1);
+    val set_map0_tacs = permute (map (fn thm => fn _ => rtac thm 1) (set_map0_of_bnf bnf));
+    fun bd_card_order_tac _ = rtac (bd_card_order_of_bnf bnf) 1;
+    fun bd_cinfinite_tac _ = rtac (bd_cinfinite_of_bnf bnf) 1;
+    val set_bd_tacs = permute (map (fn thm => fn _ => rtac thm 1) (set_bd_of_bnf bnf));
+
+    val in_alt_thm =
+      let
+        val inx = mk_in Asets sets T;
+        val in_alt = mk_in (unpermute Asets) bnf_sets T;
+        val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));
+      in
+        Goal.prove_sorry lthy [] [] goal (K (mk_permute_in_alt_tac src dest))
+        |> Thm.close_derivation
+      end;
+
+    fun le_rel_OO_tac _ = rtac (le_rel_OO_of_bnf bnf) 1;
+
+    fun rel_OO_Grp_tac _ = mk_simple_rel_OO_Grp_tac (rel_OO_Grp_of_bnf bnf) in_alt_thm;
+
+    val tacs = zip_axioms map_id0_tac map_comp0_tac map_cong0_tac set_map0_tacs bd_card_order_tac
+      bd_cinfinite_tac set_bd_tacs le_rel_OO_tac rel_OO_Grp_tac;
+
+    val wits = map snd (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf);
+
+    fun wit_tac _ = mk_simple_wit_tac (wit_thms_of_bnf bnf);
+
+    val (bnf', lthy') =
+      bnf_def Smart_Inline (K Dont_Note) qualify tacs wit_tac (SOME Ds) Binding.empty Binding.empty
+        [] ((((((b, T), mapx), sets), bd), wits), SOME rel) lthy;
+  in
+    (bnf', (add_bnf_to_unfolds bnf' unfold_set, lthy'))
+  end;
+
+(* Composition pipeline *)
+
+fun permute_and_kill qualify n src dest bnf =
+  bnf
+  |> permute_bnf qualify src dest
+  #> uncurry (kill_bnf qualify n);
+
+fun lift_and_permute qualify n src dest bnf =
+  bnf
+  |> lift_bnf qualify n
+  #> uncurry (permute_bnf qualify src dest);
+
+fun normalize_bnfs qualify Ass Ds sort bnfs unfold_set lthy =
+  let
+    val before_kill_src = map (fn As => 0 upto (length As - 1)) Ass;
+    val kill_poss = map (find_indices op = Ds) Ass;
+    val live_poss = map2 (subtract op =) kill_poss before_kill_src;
+    val before_kill_dest = map2 append kill_poss live_poss;
+    val kill_ns = map length kill_poss;
+    val (inners', (unfold_set', lthy')) =
+      fold_map5 (fn i => permute_and_kill (qualify i))
+        (if length bnfs = 1 then [0] else (1 upto length bnfs))
+        kill_ns before_kill_src before_kill_dest bnfs (unfold_set, lthy);
+
+    val Ass' = map2 (map o nth) Ass live_poss;
+    val As = sort Ass';
+    val after_lift_dest = replicate (length Ass') (0 upto (length As - 1));
+    val old_poss = map (map (fn x => find_index (fn y => x = y) As)) Ass';
+    val new_poss = map2 (subtract op =) old_poss after_lift_dest;
+    val after_lift_src = map2 append new_poss old_poss;
+    val lift_ns = map (fn xs => length As - length xs) Ass';
+  in
+    ((kill_poss, As), fold_map5 (fn i => lift_and_permute (qualify i))
+      (if length bnfs = 1 then [0] else (1 upto length bnfs))
+      lift_ns after_lift_src after_lift_dest inners' (unfold_set', lthy'))
+  end;
+
+fun default_comp_sort Ass =
+  Library.sort (Term_Ord.typ_ord o pairself TFree) (fold (fold (insert (op =))) Ass []);
+
+fun compose_bnf const_policy qualify sort outer inners oDs Dss tfreess (unfold_set, lthy) =
+  let
+    val b = name_of_bnf outer;
+
+    val Ass = map (map Term.dest_TFree) tfreess;
+    val Ds = fold (fold Term.add_tfreesT) (oDs :: Dss) [];
+
+    val ((kill_poss, As), (inners', (unfold_set', lthy'))) =
+      normalize_bnfs qualify Ass Ds sort inners unfold_set lthy;
+
+    val Ds = oDs @ flat (map3 (append oo map o nth) tfreess kill_poss Dss);
+    val As = map TFree As;
+  in
+    apfst (rpair (Ds, As))
+      (clean_compose_bnf const_policy (qualify 0) b outer inners' (unfold_set', lthy'))
+  end;
+
+(* Hide the type of the bound (optimization) and unfold the definitions (nicer to the user) *)
+
+fun seal_bnf qualify (unfold_set : unfold_set) b Ds bnf lthy =
+  let
+    val live = live_of_bnf bnf;
+    val nwits = nwits_of_bnf bnf;
+
+    val (As, lthy1) = apfst (map TFree)
+      (Variable.invent_types (replicate live HOLogic.typeS) (fold Variable.declare_typ Ds lthy));
+    val (Bs, _) = apfst (map TFree)
+      (Variable.invent_types (replicate live HOLogic.typeS) lthy1);
+
+    val map_unfolds = #map_unfolds unfold_set;
+    val set_unfoldss = #set_unfoldss unfold_set;
+    val rel_unfolds = #rel_unfolds unfold_set;
+
+    val expand_maps =
+      fold expand_term_const (map (single o Logic.dest_equals o Thm.prop_of) map_unfolds);
+    val expand_sets =
+      fold expand_term_const (map (map (Logic.dest_equals o Thm.prop_of)) set_unfoldss);
+    val expand_rels =
+      fold expand_term_const (map (single o Logic.dest_equals o Thm.prop_of) rel_unfolds);
+    fun unfold_maps ctxt = fold (unfold_thms ctxt o single) map_unfolds;
+    fun unfold_sets ctxt = fold (unfold_thms ctxt) set_unfoldss;
+    fun unfold_rels ctxt = unfold_thms ctxt rel_unfolds;
+    fun unfold_all ctxt = unfold_sets ctxt o unfold_maps ctxt o unfold_rels ctxt;
+    val bnf_map = expand_maps (mk_map_of_bnf Ds As Bs bnf);
+    val bnf_sets = map (expand_maps o expand_sets)
+      (mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf);
+    val bnf_bd = mk_bd_of_bnf Ds As bnf;
+    val bnf_rel = expand_rels (mk_rel_of_bnf Ds As Bs bnf);
+    val T = mk_T_of_bnf Ds As bnf;
+
+    (*bd should only depend on dead type variables!*)
+    val bd_repT = fst (dest_relT (fastype_of bnf_bd));
+    val bdT_bind = qualify (Binding.suffix_name ("_" ^ bdTN) b);
+    val params = fold Term.add_tfreesT Ds [];
+    val deads = map TFree params;
+
+    val ((bdT_name, (bdT_glob_info, bdT_loc_info)), lthy) =
+      typedef (bdT_bind, params, NoSyn)
+        (HOLogic.mk_UNIV bd_repT) NONE (EVERY' [rtac exI, rtac UNIV_I] 1) lthy;
+
+    val bnf_bd' = mk_dir_image bnf_bd
+      (Const (#Abs_name bdT_glob_info, bd_repT --> Type (bdT_name, deads)))
+
+    val Abs_bdT_inj = mk_Abs_inj_thm (#Abs_inject bdT_loc_info);
+    val Abs_bdT_bij = mk_Abs_bij_thm lthy Abs_bdT_inj (#Abs_cases bdT_loc_info);
+
+    val bd_ordIso = @{thm dir_image} OF [Abs_bdT_inj, bd_Card_order_of_bnf bnf];
+    val bd_card_order =
+      @{thm card_order_dir_image} OF [Abs_bdT_bij, bd_card_order_of_bnf bnf];
+    val bd_cinfinite =
+      (@{thm Cinfinite_cong} OF [bd_ordIso, bd_Cinfinite_of_bnf bnf]) RS conjunct1;
+
+    val set_bds =
+      map (fn thm => @{thm ordLeq_ordIso_trans} OF [thm, bd_ordIso]) (set_bd_of_bnf bnf);
+
+    fun mk_tac thm {context = ctxt, prems = _} =
+      (rtac (unfold_all ctxt thm) THEN'
+      SOLVE o REPEAT_DETERM o (atac ORELSE' Goal.assume_rule_tac ctxt)) 1;
+
+    val tacs = zip_axioms (mk_tac (map_id0_of_bnf bnf)) (mk_tac (map_comp0_of_bnf bnf))
+      (mk_tac (map_cong0_of_bnf bnf)) (map mk_tac (set_map0_of_bnf bnf))
+      (K (rtac bd_card_order 1)) (K (rtac bd_cinfinite 1)) (map mk_tac set_bds)
+      (mk_tac (le_rel_OO_of_bnf bnf))
+      (mk_tac (rel_OO_Grp_of_bnf bnf));
+
+    val bnf_wits = map snd (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf);
+
+    fun wit_tac {context = ctxt, prems = _} =
+      mk_simple_wit_tac (map (unfold_all ctxt) (wit_thms_of_bnf bnf));
+
+    val (bnf', lthy') =
+      bnf_def Hardly_Inline (user_policy Dont_Note) qualify tacs wit_tac (SOME deads)
+        Binding.empty Binding.empty []
+        ((((((b, T), bnf_map), bnf_sets), bnf_bd'), bnf_wits), SOME bnf_rel) lthy;
+  in
+    ((bnf', deads), lthy')
+  end;
+
+exception BAD_DEAD of typ * typ;
+
+fun bnf_of_typ _ _ _ _ (T as TFree _) accum = ((ID_bnf, ([], [T])), accum)
+  | bnf_of_typ _ _ _ _ (TVar _) _ = error "Unexpected schematic variable"
+  | bnf_of_typ const_policy qualify' sort Xs (T as Type (C, Ts)) (unfold_set, lthy) =
+    let
+      fun check_bad_dead ((_, (deads, _)), _) =
+        let val Ds = fold Term.add_tfreesT deads [] in
+          (case Library.inter (op =) Ds Xs of [] => ()
+           | X :: _ => raise BAD_DEAD (TFree X, T))
+        end;
+
+      val tfrees = Term.add_tfreesT T [];
+      val bnf_opt = if null tfrees then NONE else bnf_of lthy C;
+    in
+      (case bnf_opt of
+        NONE => ((DEADID_bnf, ([T], [])), (unfold_set, lthy))
+      | SOME bnf =>
+        if forall (can Term.dest_TFree) Ts andalso length Ts = length tfrees then
+          let
+            val T' = T_of_bnf bnf;
+            val deads = deads_of_bnf bnf;
+            val lives = lives_of_bnf bnf;
+            val tvars' = Term.add_tvarsT T' [];
+            val deads_lives =
+              pairself (map (Term.typ_subst_TVars (map fst tvars' ~~ map TFree tfrees)))
+                (deads, lives);
+          in ((bnf, deads_lives), (unfold_set, lthy)) end
+        else
+          let
+            val name = Long_Name.base_name C;
+            fun qualify i =
+              let val namei = name ^ nonzero_string_of_int i;
+              in qualify' o Binding.qualify true namei end;
+            val odead = dead_of_bnf bnf;
+            val olive = live_of_bnf bnf;
+            val oDs_pos = find_indices op = [TFree ("dead", [])] (snd (Term.dest_Type
+              (mk_T_of_bnf (replicate odead (TFree ("dead", []))) (replicate olive dummyT) bnf)));
+            val oDs = map (nth Ts) oDs_pos;
+            val Ts' = map (nth Ts) (subtract (op =) oDs_pos (0 upto length Ts - 1));
+            val ((inners, (Dss, Ass)), (unfold_set', lthy')) =
+              apfst (apsnd split_list o split_list)
+                (fold_map2 (fn i => bnf_of_typ Smart_Inline (qualify i) sort Xs)
+                (if length Ts' = 1 then [0] else (1 upto length Ts')) Ts' (unfold_set, lthy));
+          in
+            compose_bnf const_policy qualify sort bnf inners oDs Dss Ass (unfold_set', lthy')
+          end)
+      |> tap check_bad_dead
+    end;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_comp_tactics.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,252 @@
+(*  Title:      HOL/BNF/Tools/bnf_comp_tactics.ML
+    Author:     Dmitriy Traytel, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2012
+
+Tactics for composition of bounded natural functors.
+*)
+
+signature BNF_COMP_TACTICS =
+sig
+  val mk_comp_bd_card_order_tac: thm list -> thm -> tactic
+  val mk_comp_bd_cinfinite_tac: thm -> thm -> tactic
+  val mk_comp_in_alt_tac: Proof.context -> thm list -> tactic
+  val mk_comp_map_comp0_tac: thm -> thm -> thm list -> tactic
+  val mk_comp_map_cong0_tac: thm list -> thm -> thm list -> tactic
+  val mk_comp_map_id0_tac: thm -> thm -> thm list -> tactic
+  val mk_comp_set_alt_tac: Proof.context -> thm -> tactic
+  val mk_comp_set_bd_tac: Proof.context -> thm -> thm list -> tactic
+  val mk_comp_set_map0_tac: thm -> thm -> thm -> thm list -> tactic
+  val mk_comp_wit_tac: Proof.context -> thm list -> thm -> thm list -> tactic
+
+  val mk_kill_bd_card_order_tac: int -> thm -> tactic
+  val mk_kill_bd_cinfinite_tac: thm -> tactic
+  val kill_in_alt_tac: tactic
+  val mk_kill_map_cong0_tac: Proof.context -> int -> int -> thm -> tactic
+  val mk_kill_set_bd_tac: thm -> thm -> tactic
+
+  val empty_natural_tac: tactic
+  val lift_in_alt_tac: tactic
+  val mk_lift_set_bd_tac: thm -> tactic
+
+  val mk_permute_in_alt_tac: ''a list -> ''a list -> tactic
+
+  val mk_le_rel_OO_tac: thm -> thm -> thm list -> tactic
+  val mk_simple_rel_OO_Grp_tac: thm -> thm -> tactic
+  val mk_simple_wit_tac: thm list -> tactic
+end;
+
+structure BNF_Comp_Tactics : BNF_COMP_TACTICS =
+struct
+
+open BNF_Util
+open BNF_Tactics
+
+val Cnotzero_UNIV = @{thm Cnotzero_UNIV};
+val arg_cong_Union = @{thm arg_cong[of _ _ Union]};
+val csum_Cnotzero1 = @{thm csum_Cnotzero1};
+val o_eq_dest_lhs = @{thm o_eq_dest_lhs};
+val trans_image_cong_o_apply = @{thm trans[OF image_cong[OF o_apply refl]]};
+val trans_o_apply = @{thm trans[OF o_apply]};
+
+
+
+(* Composition *)
+
+fun mk_comp_set_alt_tac ctxt collect_set_map =
+  unfold_thms_tac ctxt @{thms sym[OF o_assoc]} THEN
+  unfold_thms_tac ctxt [collect_set_map RS sym] THEN
+  rtac refl 1;
+
+fun mk_comp_map_id0_tac Gmap_id0 Gmap_cong0 map_id0s =
+  EVERY' ([rtac ext, rtac (Gmap_cong0 RS trans)] @
+    map (fn thm => rtac (thm RS fun_cong)) map_id0s @ [rtac (Gmap_id0 RS fun_cong)]) 1;
+
+fun mk_comp_map_comp0_tac Gmap_comp0 Gmap_cong0 map_comp0s =
+  EVERY' ([rtac ext, rtac sym, rtac trans_o_apply,
+    rtac (Gmap_comp0 RS sym RS o_eq_dest_lhs RS trans), rtac Gmap_cong0] @
+    map (fn thm => rtac (thm RS sym RS fun_cong)) map_comp0s) 1;
+
+fun mk_comp_set_map0_tac Gmap_comp0 Gmap_cong0 Gset_map0 set_map0s =
+  EVERY' ([rtac ext] @
+    replicate 3 (rtac trans_o_apply) @
+    [rtac (arg_cong_Union RS trans),
+     rtac (@{thm arg_cong2[of _ _ _ _ collect, OF refl]} RS trans),
+     rtac (Gmap_comp0 RS sym RS o_eq_dest_lhs RS trans),
+     rtac Gmap_cong0] @
+     map (fn thm => rtac (thm RS fun_cong)) set_map0s @
+     [rtac (Gset_map0 RS o_eq_dest_lhs), rtac sym, rtac trans_o_apply,
+     rtac trans_image_cong_o_apply, rtac trans_image_cong_o_apply,
+     rtac (@{thm image_cong} OF [Gset_map0 RS o_eq_dest_lhs RS arg_cong_Union, refl] RS trans),
+     rtac @{thm trans[OF comp_eq_dest[OF Union_natural[symmetric]]]}, rtac arg_cong_Union,
+     rtac @{thm trans[OF o_eq_dest_lhs[OF image_o_collect[symmetric]]]},
+     rtac @{thm fun_cong[OF arg_cong[of _ _ collect]]}] @
+     [REPEAT_DETERM_N (length set_map0s) o EVERY' [rtac @{thm trans[OF image_insert]},
+        rtac @{thm arg_cong2[of _ _ _ _ insert]}, rtac ext, rtac trans_o_apply,
+        rtac trans_image_cong_o_apply, rtac @{thm trans[OF image_image]},
+        rtac @{thm sym[OF trans[OF o_apply]]}, rtac @{thm image_cong[OF refl o_apply]}],
+     rtac @{thm image_empty}]) 1;
+
+fun mk_comp_map_cong0_tac comp_set_alts map_cong0 map_cong0s =
+  let
+     val n = length comp_set_alts;
+  in
+    (if n = 0 then rtac refl 1
+    else rtac map_cong0 1 THEN
+      EVERY' (map_index (fn (i, map_cong0) =>
+        rtac map_cong0 THEN' EVERY' (map_index (fn (k, set_alt) =>
+          EVERY' [select_prem_tac n (dtac @{thm meta_spec}) (k + 1), etac meta_mp,
+            rtac (equalityD2 RS set_mp), rtac (set_alt RS fun_cong RS trans),
+            rtac trans_o_apply, rtac (@{thm collect_def} RS arg_cong_Union),
+            rtac @{thm UnionI}, rtac @{thm UN_I}, REPEAT_DETERM_N i o rtac @{thm insertI2},
+            rtac @{thm insertI1}, rtac (o_apply RS equalityD2 RS set_mp),
+            etac @{thm imageI}, atac])
+          comp_set_alts))
+      map_cong0s) 1)
+  end;
+
+fun mk_comp_bd_card_order_tac Fbd_card_orders Gbd_card_order =
+  let
+    val (card_orders, last_card_order) = split_last Fbd_card_orders;
+    fun gen_before thm = rtac @{thm card_order_csum} THEN' rtac thm;
+  in
+    (rtac @{thm card_order_cprod} THEN'
+    WRAP' gen_before (K (K all_tac)) card_orders (rtac last_card_order) THEN'
+    rtac Gbd_card_order) 1
+  end;
+
+fun mk_comp_bd_cinfinite_tac Fbd_cinfinite Gbd_cinfinite =
+  (rtac @{thm cinfinite_cprod} THEN'
+   ((K (TRY ((rtac @{thm cinfinite_csum} THEN' rtac disjI1) 1)) THEN'
+     ((rtac @{thm cinfinite_csum} THEN' rtac disjI1 THEN' rtac Fbd_cinfinite) ORELSE'
+      rtac Fbd_cinfinite)) ORELSE'
+    rtac Fbd_cinfinite) THEN'
+   rtac Gbd_cinfinite) 1;
+
+fun mk_comp_set_bd_tac ctxt comp_set_alt Gset_Fset_bds =
+  let
+    val (bds, last_bd) = split_last Gset_Fset_bds;
+    fun gen_before bd =
+      rtac ctrans THEN' rtac @{thm Un_csum} THEN'
+      rtac ctrans THEN' rtac @{thm csum_mono} THEN'
+      rtac bd;
+    fun gen_after _ = rtac @{thm ordIso_imp_ordLeq} THEN' rtac @{thm cprod_csum_distrib1};
+  in
+    unfold_thms_tac ctxt [comp_set_alt] THEN
+    rtac @{thm comp_set_bd_Union_o_collect} 1 THEN
+    unfold_thms_tac ctxt @{thms Union_image_insert Union_image_empty Union_Un_distrib o_apply} THEN
+    (rtac ctrans THEN'
+     WRAP' gen_before gen_after bds (rtac last_bd) THEN'
+     rtac @{thm ordIso_imp_ordLeq} THEN'
+     rtac @{thm cprod_com}) 1
+  end;
+
+val comp_in_alt_thms = @{thms o_apply collect_def SUP_def image_insert image_empty Union_insert
+  Union_empty Un_empty_right Union_Un_distrib Un_subset_iff conj_subset_def UN_image_subset
+  conj_assoc};
+
+fun mk_comp_in_alt_tac ctxt comp_set_alts =
+  unfold_thms_tac ctxt (comp_set_alts @ comp_in_alt_thms) THEN
+  unfold_thms_tac ctxt @{thms set_eq_subset} THEN
+  rtac conjI 1 THEN
+  REPEAT_DETERM (
+    rtac @{thm subsetI} 1 THEN
+    unfold_thms_tac ctxt @{thms mem_Collect_eq Ball_def} THEN
+    (REPEAT_DETERM (CHANGED (etac conjE 1)) THEN
+     REPEAT_DETERM (CHANGED ((
+       (rtac conjI THEN' (atac ORELSE' rtac subset_UNIV)) ORELSE'
+       atac ORELSE'
+       (rtac subset_UNIV)) 1)) ORELSE rtac subset_UNIV 1));
+
+val comp_wit_thms = @{thms Union_empty_conv o_apply collect_def SUP_def
+  Union_image_insert Union_image_empty};
+
+fun mk_comp_wit_tac ctxt Gwit_thms collect_set_map Fwit_thms =
+  ALLGOALS (dtac @{thm in_Union_o_assoc}) THEN
+  unfold_thms_tac ctxt (collect_set_map :: comp_wit_thms) THEN
+  REPEAT_DETERM ((atac ORELSE'
+    REPEAT_DETERM o eresolve_tac @{thms UnionE UnE} THEN'
+    etac imageE THEN' TRY o dresolve_tac Gwit_thms THEN'
+    (etac FalseE ORELSE'
+    hyp_subst_tac ctxt THEN'
+    dresolve_tac Fwit_thms THEN'
+    (etac FalseE ORELSE' atac))) 1);
+
+
+
+(* Kill operation *)
+
+fun mk_kill_map_cong0_tac ctxt n m map_cong0 =
+  (rtac map_cong0 THEN' EVERY' (replicate n (rtac refl)) THEN'
+    EVERY' (replicate m (Goal.assume_rule_tac ctxt))) 1;
+
+fun mk_kill_bd_card_order_tac n bd_card_order =
+  (rtac @{thm card_order_cprod} THEN'
+  K (REPEAT_DETERM_N (n - 1)
+    ((rtac @{thm card_order_csum} THEN'
+    rtac @{thm card_of_card_order_on}) 1)) THEN'
+  rtac @{thm card_of_card_order_on} THEN'
+  rtac bd_card_order) 1;
+
+fun mk_kill_bd_cinfinite_tac bd_Cinfinite =
+  (rtac @{thm cinfinite_cprod2} THEN'
+  TRY o rtac csum_Cnotzero1 THEN'
+  rtac Cnotzero_UNIV THEN'
+  rtac bd_Cinfinite) 1;
+
+fun mk_kill_set_bd_tac bd_Card_order set_bd =
+  (rtac ctrans THEN'
+  rtac set_bd THEN'
+  rtac @{thm ordLeq_cprod2} THEN'
+  TRY o rtac csum_Cnotzero1 THEN'
+  rtac Cnotzero_UNIV THEN'
+  rtac bd_Card_order) 1
+
+val kill_in_alt_tac =
+  ((rtac @{thm Collect_cong} THEN' rtac iffI) 1 THEN
+  REPEAT_DETERM (CHANGED (etac conjE 1)) THEN
+  REPEAT_DETERM (CHANGED ((etac conjI ORELSE'
+    rtac conjI THEN' rtac subset_UNIV) 1)) THEN
+  (rtac subset_UNIV ORELSE' atac) 1 THEN
+  REPEAT_DETERM (CHANGED (etac conjE 1)) THEN
+  REPEAT_DETERM (CHANGED ((etac conjI ORELSE' atac) 1))) ORELSE
+  ((rtac @{thm UNIV_eq_I} THEN' rtac CollectI) 1 THEN
+    REPEAT_DETERM (TRY (rtac conjI 1) THEN rtac subset_UNIV 1));
+
+
+
+(* Lift operation *)
+
+val empty_natural_tac = rtac @{thm empty_natural} 1;
+
+fun mk_lift_set_bd_tac bd_Card_order = (rtac @{thm Card_order_empty} THEN' rtac bd_Card_order) 1;
+
+val lift_in_alt_tac =
+  ((rtac @{thm Collect_cong} THEN' rtac iffI) 1 THEN
+  REPEAT_DETERM (CHANGED (etac conjE 1)) THEN
+  REPEAT_DETERM (CHANGED ((etac conjI ORELSE' atac) 1)) THEN
+  REPEAT_DETERM (CHANGED (etac conjE 1)) THEN
+  REPEAT_DETERM (CHANGED ((etac conjI ORELSE'
+    rtac conjI THEN' rtac @{thm empty_subsetI}) 1)) THEN
+  (rtac @{thm empty_subsetI} ORELSE' atac) 1) ORELSE
+  ((rtac sym THEN' rtac @{thm UNIV_eq_I} THEN' rtac CollectI) 1 THEN
+    REPEAT_DETERM (TRY (rtac conjI 1) THEN rtac @{thm empty_subsetI} 1));
+
+
+
+(* Permute operation *)
+
+fun mk_permute_in_alt_tac src dest =
+  (rtac @{thm Collect_cong} THEN'
+  mk_rotate_eq_tac (rtac refl) trans @{thm conj_assoc} @{thm conj_commute} @{thm conj_cong}
+    dest src) 1;
+
+fun mk_le_rel_OO_tac outer_le_rel_OO outer_rel_mono inner_le_rel_OOs =
+  EVERY' (map rtac (@{thm order_trans} :: outer_le_rel_OO :: outer_rel_mono :: inner_le_rel_OOs)) 1;
+
+fun mk_simple_rel_OO_Grp_tac rel_OO_Grp in_alt_thm =
+  rtac (trans OF [rel_OO_Grp, in_alt_thm RS @{thm OO_Grp_cong} RS sym]) 1;
+
+fun mk_simple_wit_tac wit_thms = ALLGOALS (atac ORELSE' eresolve_tac (@{thm emptyE} :: wit_thms));
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_def.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,1393 @@
+(*  Title:      HOL/BNF/Tools/bnf_def.ML
+    Author:     Dmitriy Traytel, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2012
+
+Definition of bounded natural functors.
+*)
+
+signature BNF_DEF =
+sig
+  type bnf
+  type nonemptiness_witness = {I: int list, wit: term, prop: thm list}
+
+  val morph_bnf: morphism -> bnf -> bnf
+  val eq_bnf: bnf * bnf -> bool
+  val bnf_of: Proof.context -> string -> bnf option
+  val register_bnf: string -> (bnf * local_theory) -> (bnf * local_theory)
+
+  val name_of_bnf: bnf -> binding
+  val T_of_bnf: bnf -> typ
+  val live_of_bnf: bnf -> int
+  val lives_of_bnf: bnf -> typ list
+  val dead_of_bnf: bnf -> int
+  val deads_of_bnf: bnf -> typ list
+  val nwits_of_bnf: bnf -> int
+
+  val mapN: string
+  val relN: string
+  val setN: string
+  val mk_setN: int -> string
+  val mk_witN: int -> string
+
+  val map_of_bnf: bnf -> term
+  val sets_of_bnf: bnf -> term list
+  val rel_of_bnf: bnf -> term
+
+  val mk_T_of_bnf: typ list -> typ list -> bnf -> typ
+  val mk_bd_of_bnf: typ list -> typ list -> bnf -> term
+  val mk_map_of_bnf: typ list -> typ list -> typ list -> bnf -> term
+  val mk_rel_of_bnf: typ list -> typ list -> typ list -> bnf -> term
+  val mk_sets_of_bnf: typ list list -> typ list list -> bnf -> term list
+  val mk_wits_of_bnf: typ list list -> typ list list -> bnf -> (int list * term) list
+
+  val bd_Card_order_of_bnf: bnf -> thm
+  val bd_Cinfinite_of_bnf: bnf -> thm
+  val bd_Cnotzero_of_bnf: bnf -> thm
+  val bd_card_order_of_bnf: bnf -> thm
+  val bd_cinfinite_of_bnf: bnf -> thm
+  val collect_set_map_of_bnf: bnf -> thm
+  val in_bd_of_bnf: bnf -> thm
+  val in_cong_of_bnf: bnf -> thm
+  val in_mono_of_bnf: bnf -> thm
+  val in_rel_of_bnf: bnf -> thm
+  val map_comp0_of_bnf: bnf -> thm
+  val map_comp_of_bnf: bnf -> thm
+  val map_cong0_of_bnf: bnf -> thm
+  val map_cong_of_bnf: bnf -> thm
+  val map_def_of_bnf: bnf -> thm
+  val map_id0_of_bnf: bnf -> thm
+  val map_id_of_bnf: bnf -> thm
+  val map_transfer_of_bnf: bnf -> thm
+  val le_rel_OO_of_bnf: bnf -> thm
+  val rel_def_of_bnf: bnf -> thm
+  val rel_Grp_of_bnf: bnf -> thm
+  val rel_OO_of_bnf: bnf -> thm
+  val rel_OO_Grp_of_bnf: bnf -> thm
+  val rel_cong_of_bnf: bnf -> thm
+  val rel_conversep_of_bnf: bnf -> thm
+  val rel_mono_of_bnf: bnf -> thm
+  val rel_mono_strong_of_bnf: bnf -> thm
+  val rel_eq_of_bnf: bnf -> thm
+  val rel_flip_of_bnf: bnf -> thm
+  val set_bd_of_bnf: bnf -> thm list
+  val set_defs_of_bnf: bnf -> thm list
+  val set_map0_of_bnf: bnf -> thm list
+  val set_map_of_bnf: bnf -> thm list
+  val wit_thms_of_bnf: bnf -> thm list
+  val wit_thmss_of_bnf: bnf -> thm list list
+
+  val mk_map: int -> typ list -> typ list -> term -> term
+  val mk_rel: int -> typ list -> typ list -> term -> term
+  val build_map: Proof.context -> (typ * typ -> term) -> typ * typ -> term
+  val build_rel: Proof.context -> (typ * typ -> term) -> typ * typ -> term
+  val flatten_type_args_of_bnf: bnf -> 'a -> 'a list -> 'a list
+  val map_flattened_map_args: Proof.context -> string -> (term list -> 'a list) -> term list ->
+    'a list
+
+  val mk_witness: int list * term -> thm list -> nonemptiness_witness
+  val minimize_wits: (''a list * 'b) list -> (''a list * 'b) list
+  val wits_of_bnf: bnf -> nonemptiness_witness list
+
+  val zip_axioms: 'a -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list
+
+  datatype const_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline
+  datatype fact_policy = Dont_Note | Note_Some | Note_All
+
+  val bnf_note_all: bool Config.T
+  val bnf_timing: bool Config.T
+  val user_policy: fact_policy -> Proof.context -> fact_policy
+  val note_bnf_thms: fact_policy -> (binding -> binding) -> binding -> bnf -> Proof.context ->
+    Proof.context
+
+  val print_bnfs: Proof.context -> unit
+  val prepare_def: const_policy -> (Proof.context -> fact_policy) -> (binding -> binding) ->
+    (Proof.context -> 'a -> typ) -> (Proof.context -> 'b -> term) -> typ list option ->
+    binding -> binding -> binding list ->
+    (((((binding * 'a) * 'b) * 'b list) * 'b) * 'b list) * 'b option -> Proof.context ->
+    string * term list *
+    ((thm list -> {context: Proof.context, prems: thm list} -> tactic) option * term list list) *
+    ((thm list -> thm list list) -> thm list list -> Proof.context -> bnf * local_theory) *
+    local_theory * thm list
+
+  val define_bnf_consts: const_policy -> fact_policy -> typ list option ->
+    binding -> binding -> binding list ->
+    (((((binding * typ) * term) * term list) * term) * term list) * term option -> local_theory ->
+      ((typ list * typ list * typ list * typ) *
+       (term * term list * term * (int list * term) list * term) *
+       (thm * thm list * thm * thm list * thm) *
+       ((typ list -> typ list -> typ list -> term) *
+        (typ list -> typ list -> term -> term) *
+        (typ list -> typ list -> typ -> typ) *
+        (typ list -> typ list -> typ list -> term) *
+        (typ list -> typ list -> typ list -> term))) * local_theory
+
+  val bnf_def: const_policy -> (Proof.context -> fact_policy) -> (binding -> binding) ->
+    ({prems: thm list, context: Proof.context} -> tactic) list ->
+    ({prems: thm list, context: Proof.context} -> tactic) -> typ list option -> binding ->
+    binding -> binding list ->
+    (((((binding * typ) * term) * term list) * term) * term list) * term option ->
+    local_theory -> bnf * local_theory
+end;
+
+structure BNF_Def : BNF_DEF =
+struct
+
+open BNF_Util
+open BNF_Tactics
+open BNF_Def_Tactics
+
+val fundefcong_attrs = @{attributes [fundef_cong]};
+
+type axioms = {
+  map_id0: thm,
+  map_comp0: thm,
+  map_cong0: thm,
+  set_map0: thm list,
+  bd_card_order: thm,
+  bd_cinfinite: thm,
+  set_bd: thm list,
+  le_rel_OO: thm,
+  rel_OO_Grp: thm
+};
+
+fun mk_axioms' ((((((((id, comp), cong), map), c_o), cinf), set_bd), le_rel_OO), rel) =
+  {map_id0 = id, map_comp0 = comp, map_cong0 = cong, set_map0 = map, bd_card_order = c_o,
+   bd_cinfinite = cinf, set_bd = set_bd, le_rel_OO = le_rel_OO, rel_OO_Grp = rel};
+
+fun dest_cons [] = raise List.Empty
+  | dest_cons (x :: xs) = (x, xs);
+
+fun mk_axioms n thms = thms
+  |> map the_single
+  |> dest_cons
+  ||>> dest_cons
+  ||>> dest_cons
+  ||>> chop n
+  ||>> dest_cons
+  ||>> dest_cons
+  ||>> chop n
+  ||>> dest_cons
+  ||> the_single
+  |> mk_axioms';
+
+fun zip_axioms mid mcomp mcong smap bdco bdinf sbd le_rel_OO rel =
+  [mid, mcomp, mcong] @ smap @ [bdco, bdinf] @ sbd @ [le_rel_OO, rel];
+
+fun dest_axioms {map_id0, map_comp0, map_cong0, set_map0, bd_card_order, bd_cinfinite, set_bd,
+  le_rel_OO, rel_OO_Grp} =
+  zip_axioms map_id0 map_comp0 map_cong0 set_map0 bd_card_order bd_cinfinite set_bd le_rel_OO
+    rel_OO_Grp;
+
+fun map_axioms f {map_id0, map_comp0, map_cong0, set_map0, bd_card_order, bd_cinfinite, set_bd,
+  le_rel_OO, rel_OO_Grp} =
+  {map_id0 = f map_id0,
+    map_comp0 = f map_comp0,
+    map_cong0 = f map_cong0,
+    set_map0 = map f set_map0,
+    bd_card_order = f bd_card_order,
+    bd_cinfinite = f bd_cinfinite,
+    set_bd = map f set_bd,
+    le_rel_OO = f le_rel_OO,
+    rel_OO_Grp = f rel_OO_Grp};
+
+val morph_axioms = map_axioms o Morphism.thm;
+
+type defs = {
+  map_def: thm,
+  set_defs: thm list,
+  rel_def: thm
+}
+
+fun mk_defs map sets rel = {map_def = map, set_defs = sets, rel_def = rel};
+
+fun map_defs f {map_def, set_defs, rel_def} =
+  {map_def = f map_def, set_defs = map f set_defs, rel_def = f rel_def};
+
+val morph_defs = map_defs o Morphism.thm;
+
+type facts = {
+  bd_Card_order: thm,
+  bd_Cinfinite: thm,
+  bd_Cnotzero: thm,
+  collect_set_map: thm lazy,
+  in_bd: thm lazy,
+  in_cong: thm lazy,
+  in_mono: thm lazy,
+  in_rel: thm lazy,
+  map_comp: thm lazy,
+  map_cong: thm lazy,
+  map_id: thm lazy,
+  map_transfer: thm lazy,
+  rel_eq: thm lazy,
+  rel_flip: thm lazy,
+  set_map: thm lazy list,
+  rel_cong: thm lazy,
+  rel_mono: thm lazy,
+  rel_mono_strong: thm lazy,
+  rel_Grp: thm lazy,
+  rel_conversep: thm lazy,
+  rel_OO: thm lazy
+};
+
+fun mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_map in_bd in_cong in_mono in_rel
+    map_comp map_cong map_id map_transfer rel_eq rel_flip set_map rel_cong rel_mono
+    rel_mono_strong rel_Grp rel_conversep rel_OO = {
+  bd_Card_order = bd_Card_order,
+  bd_Cinfinite = bd_Cinfinite,
+  bd_Cnotzero = bd_Cnotzero,
+  collect_set_map = collect_set_map,
+  in_bd = in_bd,
+  in_cong = in_cong,
+  in_mono = in_mono,
+  in_rel = in_rel,
+  map_comp = map_comp,
+  map_cong = map_cong,
+  map_id = map_id,
+  map_transfer = map_transfer,
+  rel_eq = rel_eq,
+  rel_flip = rel_flip,
+  set_map = set_map,
+  rel_cong = rel_cong,
+  rel_mono = rel_mono,
+  rel_mono_strong = rel_mono_strong,
+  rel_Grp = rel_Grp,
+  rel_conversep = rel_conversep,
+  rel_OO = rel_OO};
+
+fun map_facts f {
+  bd_Card_order,
+  bd_Cinfinite,
+  bd_Cnotzero,
+  collect_set_map,
+  in_bd,
+  in_cong,
+  in_mono,
+  in_rel,
+  map_comp,
+  map_cong,
+  map_id,
+  map_transfer,
+  rel_eq,
+  rel_flip,
+  set_map,
+  rel_cong,
+  rel_mono,
+  rel_mono_strong,
+  rel_Grp,
+  rel_conversep,
+  rel_OO} =
+  {bd_Card_order = f bd_Card_order,
+    bd_Cinfinite = f bd_Cinfinite,
+    bd_Cnotzero = f bd_Cnotzero,
+    collect_set_map = Lazy.map f collect_set_map,
+    in_bd = Lazy.map f in_bd,
+    in_cong = Lazy.map f in_cong,
+    in_mono = Lazy.map f in_mono,
+    in_rel = Lazy.map f in_rel,
+    map_comp = Lazy.map f map_comp,
+    map_cong = Lazy.map f map_cong,
+    map_id = Lazy.map f map_id,
+    map_transfer = Lazy.map f map_transfer,
+    rel_eq = Lazy.map f rel_eq,
+    rel_flip = Lazy.map f rel_flip,
+    set_map = map (Lazy.map f) set_map,
+    rel_cong = Lazy.map f rel_cong,
+    rel_mono = Lazy.map f rel_mono,
+    rel_mono_strong = Lazy.map f rel_mono_strong,
+    rel_Grp = Lazy.map f rel_Grp,
+    rel_conversep = Lazy.map f rel_conversep,
+    rel_OO = Lazy.map f rel_OO};
+
+val morph_facts = map_facts o Morphism.thm;
+
+type nonemptiness_witness = {
+  I: int list,
+  wit: term,
+  prop: thm list
+};
+
+fun mk_witness (I, wit) prop = {I = I, wit = wit, prop = prop};
+fun map_witness f g {I, wit, prop} = {I = I, wit = f wit, prop = map g prop};
+fun morph_witness phi = map_witness (Morphism.term phi) (Morphism.thm phi);
+
+datatype bnf = BNF of {
+  name: binding,
+  T: typ,
+  live: int,
+  lives: typ list, (*source type variables of map*)
+  lives': typ list, (*target type variables of map*)
+  dead: int,
+  deads: typ list,
+  map: term,
+  sets: term list,
+  bd: term,
+  axioms: axioms,
+  defs: defs,
+  facts: facts,
+  nwits: int,
+  wits: nonemptiness_witness list,
+  rel: term
+};
+
+(* getters *)
+
+fun rep_bnf (BNF bnf) = bnf;
+val name_of_bnf = #name o rep_bnf;
+val T_of_bnf = #T o rep_bnf;
+fun mk_T_of_bnf Ds Ts bnf =
+  let val bnf_rep = rep_bnf bnf
+  in Term.typ_subst_atomic ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#T bnf_rep) end;
+val live_of_bnf = #live o rep_bnf;
+val lives_of_bnf = #lives o rep_bnf;
+val dead_of_bnf = #dead o rep_bnf;
+val deads_of_bnf = #deads o rep_bnf;
+val axioms_of_bnf = #axioms o rep_bnf;
+val facts_of_bnf = #facts o rep_bnf;
+val nwits_of_bnf = #nwits o rep_bnf;
+val wits_of_bnf = #wits o rep_bnf;
+
+fun flatten_type_args_of_bnf bnf dead_x xs =
+  let
+    val Type (_, Ts) = T_of_bnf bnf;
+    val lives = lives_of_bnf bnf;
+    val deads = deads_of_bnf bnf;
+  in
+    permute_like (op =) (deads @ lives) Ts (replicate (length deads) dead_x @ xs)
+  end;
+
+(*terms*)
+val map_of_bnf = #map o rep_bnf;
+val sets_of_bnf = #sets o rep_bnf;
+fun mk_map_of_bnf Ds Ts Us bnf =
+  let val bnf_rep = rep_bnf bnf;
+  in
+    Term.subst_atomic_types
+      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#map bnf_rep)
+  end;
+fun mk_sets_of_bnf Dss Tss bnf =
+  let val bnf_rep = rep_bnf bnf;
+  in
+    map2 (fn (Ds, Ts) => Term.subst_atomic_types
+      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts))) (Dss ~~ Tss) (#sets bnf_rep)
+  end;
+val bd_of_bnf = #bd o rep_bnf;
+fun mk_bd_of_bnf Ds Ts bnf =
+  let val bnf_rep = rep_bnf bnf;
+  in Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#bd bnf_rep) end;
+fun mk_wits_of_bnf Dss Tss bnf =
+  let
+    val bnf_rep = rep_bnf bnf;
+    val wits = map (fn x => (#I x, #wit x)) (#wits bnf_rep);
+  in
+    map2 (fn (Ds, Ts) => apsnd (Term.subst_atomic_types
+      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)))) (Dss ~~ Tss) wits
+  end;
+val rel_of_bnf = #rel o rep_bnf;
+fun mk_rel_of_bnf Ds Ts Us bnf =
+  let val bnf_rep = rep_bnf bnf;
+  in
+    Term.subst_atomic_types
+      ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#rel bnf_rep)
+  end;
+
+(*thms*)
+val bd_card_order_of_bnf = #bd_card_order o #axioms o rep_bnf;
+val bd_cinfinite_of_bnf = #bd_cinfinite o #axioms o rep_bnf;
+val bd_Card_order_of_bnf = #bd_Card_order o #facts o rep_bnf;
+val bd_Cinfinite_of_bnf = #bd_Cinfinite o #facts o rep_bnf;
+val bd_Cnotzero_of_bnf = #bd_Cnotzero o #facts o rep_bnf;
+val collect_set_map_of_bnf = Lazy.force o #collect_set_map o #facts o rep_bnf;
+val in_bd_of_bnf = Lazy.force o #in_bd o #facts o rep_bnf;
+val in_cong_of_bnf = Lazy.force o #in_cong o #facts o rep_bnf;
+val in_mono_of_bnf = Lazy.force o #in_mono o #facts o rep_bnf;
+val in_rel_of_bnf = Lazy.force o #in_rel o #facts o rep_bnf;
+val map_def_of_bnf = #map_def o #defs o rep_bnf;
+val map_id0_of_bnf = #map_id0 o #axioms o rep_bnf;
+val map_id_of_bnf = Lazy.force o #map_id o #facts o rep_bnf;
+val map_comp0_of_bnf = #map_comp0 o #axioms o rep_bnf;
+val map_comp_of_bnf = Lazy.force o #map_comp o #facts o rep_bnf;
+val map_cong0_of_bnf = #map_cong0 o #axioms o rep_bnf;
+val map_cong_of_bnf = Lazy.force o #map_cong o #facts o rep_bnf;
+val map_transfer_of_bnf = Lazy.force o #map_transfer o #facts o rep_bnf;
+val le_rel_OO_of_bnf = #le_rel_OO o #axioms o rep_bnf;
+val rel_def_of_bnf = #rel_def o #defs o rep_bnf;
+val rel_eq_of_bnf = Lazy.force o #rel_eq o #facts o rep_bnf;
+val rel_flip_of_bnf = Lazy.force o #rel_flip o #facts o rep_bnf;
+val set_bd_of_bnf = #set_bd o #axioms o rep_bnf;
+val set_defs_of_bnf = #set_defs o #defs o rep_bnf;
+val set_map0_of_bnf = #set_map0 o #axioms o rep_bnf;
+val set_map_of_bnf = map Lazy.force o #set_map o #facts o rep_bnf;
+val rel_cong_of_bnf = Lazy.force o #rel_cong o #facts o rep_bnf;
+val rel_mono_of_bnf = Lazy.force o #rel_mono o #facts o rep_bnf;
+val rel_mono_strong_of_bnf = Lazy.force o #rel_mono_strong o #facts o rep_bnf;
+val rel_Grp_of_bnf = Lazy.force o #rel_Grp o #facts o rep_bnf;
+val rel_conversep_of_bnf = Lazy.force o #rel_conversep o #facts o rep_bnf;
+val rel_OO_of_bnf = Lazy.force o #rel_OO o #facts o rep_bnf;
+val rel_OO_Grp_of_bnf = #rel_OO_Grp o #axioms o rep_bnf;
+val wit_thms_of_bnf = maps #prop o wits_of_bnf;
+val wit_thmss_of_bnf = map #prop o wits_of_bnf;
+
+fun mk_bnf name T live lives lives' dead deads map sets bd axioms defs facts wits rel =
+  BNF {name = name, T = T,
+       live = live, lives = lives, lives' = lives', dead = dead, deads = deads,
+       map = map, sets = sets, bd = bd,
+       axioms = axioms, defs = defs, facts = facts,
+       nwits = length wits, wits = wits, rel = rel};
+
+fun morph_bnf phi (BNF {name = name, T = T, live = live, lives = lives, lives' = lives',
+  dead = dead, deads = deads, map = map, sets = sets, bd = bd,
+  axioms = axioms, defs = defs, facts = facts,
+  nwits = nwits, wits = wits, rel = rel}) =
+  BNF {name = Morphism.binding phi name, T = Morphism.typ phi T,
+    live = live, lives = List.map (Morphism.typ phi) lives,
+    lives' = List.map (Morphism.typ phi) lives',
+    dead = dead, deads = List.map (Morphism.typ phi) deads,
+    map = Morphism.term phi map, sets = List.map (Morphism.term phi) sets,
+    bd = Morphism.term phi bd,
+    axioms = morph_axioms phi axioms,
+    defs = morph_defs phi defs,
+    facts = morph_facts phi facts,
+    nwits = nwits,
+    wits = List.map (morph_witness phi) wits,
+    rel = Morphism.term phi rel};
+
+fun eq_bnf (BNF {T = T1, live = live1, dead = dead1, ...},
+  BNF {T = T2, live = live2, dead = dead2, ...}) =
+  Type.could_unify (T1, T2) andalso live1 = live2 andalso dead1 = dead2;
+
+structure Data = Generic_Data
+(
+  type T = bnf Symtab.table;
+  val empty = Symtab.empty;
+  val extend = I;
+  val merge = Symtab.merge eq_bnf;
+);
+
+fun bnf_of ctxt =
+  Symtab.lookup (Data.get (Context.Proof ctxt))
+  #> Option.map (morph_bnf (Morphism.transfer_morphism (Proof_Context.theory_of ctxt)));
+
+
+(* Utilities *)
+
+fun normalize_set insts instA set =
+  let
+    val (T, T') = dest_funT (fastype_of set);
+    val A = fst (Term.dest_TVar (HOLogic.dest_setT T'));
+    val params = Term.add_tvar_namesT T [];
+  in Term.subst_TVars ((A :: params) ~~ (instA :: insts)) set end;
+
+fun normalize_rel ctxt instTs instA instB rel =
+  let
+    val thy = Proof_Context.theory_of ctxt;
+    val tyenv =
+      Sign.typ_match thy (fastype_of rel, Library.foldr (op -->) (instTs, mk_pred2T instA instB))
+        Vartab.empty;
+  in Envir.subst_term (tyenv, Vartab.empty) rel end
+  handle Type.TYPE_MATCH => error "Bad relator";
+
+fun normalize_wit insts CA As wit =
+  let
+    fun strip_param (Ts, T as Type (@{type_name fun}, [T1, T2])) =
+        if Type.raw_instance (CA, T) then (Ts, T) else strip_param (T1 :: Ts, T2)
+      | strip_param x = x;
+    val (Ts, T) = strip_param ([], fastype_of wit);
+    val subst = Term.add_tvar_namesT T [] ~~ insts;
+    fun find y = find_index (fn x => x = y) As;
+  in
+    (map (find o Term.typ_subst_TVars subst) (rev Ts), Term.subst_TVars subst wit)
+  end;
+
+fun minimize_wits wits =
+ let
+   fun minimize done [] = done
+     | minimize done ((I, wit) :: todo) =
+       if exists (fn (J, _) => subset (op =) (J, I)) (done @ todo)
+       then minimize done todo
+       else minimize ((I, wit) :: done) todo;
+ in minimize [] wits end;
+
+fun mk_map live Ts Us t =
+  let val (Type (_, Ts0), Type (_, Us0)) = strip_typeN (live + 1) (fastype_of t) |>> List.last in
+    Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
+  end;
+
+fun mk_rel live Ts Us t =
+  let val [Type (_, Ts0), Type (_, Us0)] = binder_types (snd (strip_typeN live (fastype_of t))) in
+    Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
+  end;
+
+fun build_map_or_rel mk const of_bnf dest ctxt build_simple =
+  let
+    fun build (TU as (T, U)) =
+      if T = U then
+        const T
+      else
+        (case TU of
+          (Type (s, Ts), Type (s', Us)) =>
+          if s = s' then
+            let
+              val bnf = the (bnf_of ctxt s);
+              val live = live_of_bnf bnf;
+              val mapx = mk live Ts Us (of_bnf bnf);
+              val TUs' = map dest (fst (strip_typeN live (fastype_of mapx)));
+            in Term.list_comb (mapx, map build TUs') end
+          else
+            build_simple TU
+        | _ => build_simple TU);
+  in build end;
+
+val build_map = build_map_or_rel mk_map HOLogic.id_const map_of_bnf dest_funT;
+val build_rel = build_map_or_rel mk_rel HOLogic.eq_const rel_of_bnf dest_pred2T;
+
+fun map_flattened_map_args ctxt s map_args fs =
+  let
+    val flat_fs = flatten_type_args_of_bnf (the (bnf_of ctxt s)) Term.dummy fs;
+    val flat_fs' = map_args flat_fs;
+  in
+    permute_like (op aconv) flat_fs fs flat_fs'
+  end;
+
+
+(* Names *)
+
+val mapN = "map";
+val setN = "set";
+fun mk_setN i = setN ^ nonzero_string_of_int i;
+val bdN = "bd";
+val witN = "wit";
+fun mk_witN i = witN ^ nonzero_string_of_int i;
+val relN = "rel";
+
+val bd_card_orderN = "bd_card_order";
+val bd_cinfiniteN = "bd_cinfinite";
+val bd_Card_orderN = "bd_Card_order";
+val bd_CinfiniteN = "bd_Cinfinite";
+val bd_CnotzeroN = "bd_Cnotzero";
+val collect_set_mapN = "collect_set_map";
+val in_bdN = "in_bd";
+val in_monoN = "in_mono";
+val in_relN = "in_rel";
+val map_id0N = "map_id0";
+val map_idN = "map_id";
+val map_comp0N = "map_comp0";
+val map_compN = "map_comp";
+val map_cong0N = "map_cong0";
+val map_congN = "map_cong";
+val map_transferN = "map_transfer";
+val rel_eqN = "rel_eq";
+val rel_flipN = "rel_flip";
+val set_map0N = "set_map0";
+val set_mapN = "set_map";
+val set_bdN = "set_bd";
+val rel_GrpN = "rel_Grp";
+val rel_conversepN = "rel_conversep";
+val rel_monoN = "rel_mono"
+val rel_mono_strongN = "rel_mono_strong"
+val rel_comppN = "rel_compp";
+val rel_compp_GrpN = "rel_compp_Grp";
+
+datatype const_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline;
+
+datatype fact_policy = Dont_Note | Note_Some | Note_All;
+
+val bnf_note_all = Attrib.setup_config_bool @{binding bnf_note_all} (K false);
+val bnf_timing = Attrib.setup_config_bool @{binding bnf_timing} (K false);
+
+fun user_policy policy ctxt = if Config.get ctxt bnf_note_all then Note_All else policy;
+
+val smart_max_inline_size = 25; (*FUDGE*)
+
+fun note_bnf_thms fact_policy qualify' bnf_b bnf =
+  let
+    val axioms = axioms_of_bnf bnf;
+    val facts = facts_of_bnf bnf;
+    val wits = wits_of_bnf bnf;
+    val qualify =
+      let val (_, qs, _) = Binding.dest bnf_b;
+      in fold_rev (fn (s, mand) => Binding.qualify mand s) qs #> qualify' end;
+  in
+    (if fact_policy = Note_All then
+      let
+        val witNs = if length wits = 1 then [witN] else map mk_witN (1 upto length wits);
+        val notes =
+          [(bd_card_orderN, [#bd_card_order axioms]),
+            (bd_cinfiniteN, [#bd_cinfinite axioms]),
+            (bd_Card_orderN, [#bd_Card_order facts]),
+            (bd_CinfiniteN, [#bd_Cinfinite facts]),
+            (bd_CnotzeroN, [#bd_Cnotzero facts]),
+            (collect_set_mapN, [Lazy.force (#collect_set_map facts)]),
+            (in_bdN, [Lazy.force (#in_bd facts)]),
+            (in_monoN, [Lazy.force (#in_mono facts)]),
+            (in_relN, [Lazy.force (#in_rel facts)]),
+            (map_comp0N, [#map_comp0 axioms]),
+            (map_id0N, [#map_id0 axioms]),
+            (map_transferN, [Lazy.force (#map_transfer facts)]),
+            (rel_mono_strongN, [Lazy.force (#rel_mono_strong facts)]),
+            (set_map0N, #set_map0 axioms),
+            (set_bdN, #set_bd axioms)] @
+            (witNs ~~ wit_thmss_of_bnf bnf)
+            |> map (fn (thmN, thms) =>
+              ((qualify (Binding.qualify true (Binding.name_of bnf_b) (Binding.name thmN)), []),
+              [(thms, [])]));
+        in
+          Local_Theory.notes notes #> snd
+        end
+      else
+        I)
+    #> (if fact_policy <> Dont_Note then
+        let
+          val notes =
+            [(map_compN, [Lazy.force (#map_comp facts)], []),
+            (map_cong0N, [#map_cong0 axioms], []),
+            (map_congN, [Lazy.force (#map_cong facts)], fundefcong_attrs),
+            (map_idN, [Lazy.force (#map_id facts)], []),
+            (rel_comppN, [Lazy.force (#rel_OO facts)], []),
+            (rel_compp_GrpN, no_refl [#rel_OO_Grp axioms], []),
+            (rel_conversepN, [Lazy.force (#rel_conversep facts)], []),
+            (rel_eqN, [Lazy.force (#rel_eq facts)], []),
+            (rel_flipN, [Lazy.force (#rel_flip facts)], []),
+            (rel_GrpN, [Lazy.force (#rel_Grp facts)], []),
+            (rel_monoN, [Lazy.force (#rel_mono facts)], []),
+            (set_mapN, map Lazy.force (#set_map facts), [])]
+            |> filter_out (null o #2)
+            |> map (fn (thmN, thms, attrs) =>
+              ((qualify (Binding.qualify true (Binding.name_of bnf_b) (Binding.name thmN)),
+                attrs), [(thms, [])]));
+        in
+          Local_Theory.notes notes #> snd
+        end
+      else
+        I)
+  end;
+
+
+(* Define new BNFs *)
+
+fun define_bnf_consts const_policy fact_policy Ds_opt map_b rel_b set_bs
+  ((((((bnf_b, T_rhs), map_rhs), set_rhss), bd_rhs), wit_rhss), rel_rhs_opt) no_defs_lthy =
+  let
+    val live = length set_rhss;
+
+    val def_qualify = Binding.conceal o Binding.qualify false (Binding.name_of bnf_b);
+
+    fun mk_prefix_binding pre = Binding.prefix_name (pre ^ "_") bnf_b;
+
+    fun maybe_define user_specified (b, rhs) lthy =
+      let
+        val inline =
+          (user_specified orelse fact_policy = Dont_Note) andalso
+          (case const_policy of
+            Dont_Inline => false
+          | Hardly_Inline => Term.is_Free rhs orelse Term.is_Const rhs
+          | Smart_Inline => Term.size_of_term rhs <= smart_max_inline_size
+          | Do_Inline => true)
+      in
+        if inline then
+          ((rhs, Drule.reflexive_thm), lthy)
+        else
+          let val b = b () in
+            apfst (apsnd snd) (Local_Theory.define ((b, NoSyn), ((Thm.def_binding b, []), rhs))
+              lthy)
+          end
+      end;
+
+    fun maybe_restore lthy_old lthy =
+      lthy |> not (pointer_eq (lthy_old, lthy)) ? Local_Theory.restore;
+
+    val map_bind_def =
+      (fn () => def_qualify (if Binding.is_empty map_b then mk_prefix_binding mapN else map_b),
+         map_rhs);
+    val set_binds_defs =
+      let
+        fun set_name i get_b =
+          (case try (nth set_bs) (i - 1) of
+            SOME b => if Binding.is_empty b then get_b else K b
+          | NONE => get_b) #> def_qualify;
+        val bs = if live = 1 then [set_name 1 (fn () => mk_prefix_binding setN)]
+          else map (fn i => set_name i (fn () => mk_prefix_binding (mk_setN i))) (1 upto live);
+      in bs ~~ set_rhss end;
+    val bd_bind_def = (fn () => def_qualify (mk_prefix_binding bdN), bd_rhs);
+
+    val ((((bnf_map_term, raw_map_def),
+      (bnf_set_terms, raw_set_defs)),
+      (bnf_bd_term, raw_bd_def)), (lthy, lthy_old)) =
+        no_defs_lthy
+        |> maybe_define true map_bind_def
+        ||>> apfst split_list o fold_map (maybe_define true) set_binds_defs
+        ||>> maybe_define true bd_bind_def
+        ||> `(maybe_restore no_defs_lthy);
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+
+
+    val bnf_map_def = Morphism.thm phi raw_map_def;
+    val bnf_set_defs = map (Morphism.thm phi) raw_set_defs;
+    val bnf_bd_def = Morphism.thm phi raw_bd_def;
+
+    val bnf_map = Morphism.term phi bnf_map_term;
+
+    (*TODO: handle errors*)
+    (*simple shape analysis of a map function*)
+    val ((alphas, betas), (Calpha, _)) =
+      fastype_of bnf_map
+      |> strip_typeN live
+      |>> map_split dest_funT
+      ||> dest_funT
+      handle TYPE _ => error "Bad map function";
+
+    val Calpha_params = map TVar (Term.add_tvarsT Calpha []);
+
+    val bnf_T = Morphism.typ phi T_rhs;
+    val bad_args = Term.add_tfreesT bnf_T [];
+    val _ = if null bad_args then () else error ("Locally fixed type arguments " ^
+      commas_quote (map (Syntax.string_of_typ no_defs_lthy o TFree) bad_args));
+
+    val bnf_sets =
+      map2 (normalize_set Calpha_params) alphas (map (Morphism.term phi) bnf_set_terms);
+    val bnf_bd =
+      Term.subst_TVars (Term.add_tvar_namesT bnf_T [] ~~ Calpha_params)
+        (Morphism.term phi bnf_bd_term);
+
+    (*TODO: assert Ds = (TVars of bnf_map) \ (alphas @ betas) as sets*)
+    val deads = (case Ds_opt of
+      NONE => subtract (op =) (alphas @ betas) (map TVar (Term.add_tvars bnf_map []))
+    | SOME Ds => map (Morphism.typ phi) Ds);
+
+    (*TODO: further checks of type of bnf_map*)
+    (*TODO: check types of bnf_sets*)
+    (*TODO: check type of bnf_bd*)
+    (*TODO: check type of bnf_rel*)
+
+    fun mk_bnf_map Ds As' Bs' =
+      Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As') @ (betas ~~ Bs')) bnf_map;
+    fun mk_bnf_t Ds As' = Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As'));
+    fun mk_bnf_T Ds As' = Term.typ_subst_atomic ((deads ~~ Ds) @ (alphas ~~ As'));
+
+    val (((As, Bs), Ds), names_lthy) = lthy
+      |> mk_TFrees live
+      ||>> mk_TFrees live
+      ||>> mk_TFrees (length deads);
+    val RTs = map2 (curry HOLogic.mk_prodT) As Bs;
+    val pred2RTs = map2 mk_pred2T As Bs;
+    val (Rs, Rs') = names_lthy |> mk_Frees' "R" pred2RTs |> fst
+    val CA = mk_bnf_T Ds As Calpha;
+    val CR = mk_bnf_T Ds RTs Calpha;
+    val setRs =
+      map3 (fn R => fn T => fn U =>
+          HOLogic.Collect_const (HOLogic.mk_prodT (T, U)) $ HOLogic.mk_split R) Rs As Bs;
+
+    (*Grp (in (Collect (split R1) .. Collect (split Rn))) (map fst .. fst)^--1 OO
+      Grp (in (Collect (split R1) .. Collect (split Rn))) (map snd .. snd)*)
+    val OO_Grp =
+      let
+        val map1 = Term.list_comb (mk_bnf_map Ds RTs As, map fst_const RTs);
+        val map2 = Term.list_comb (mk_bnf_map Ds RTs Bs, map snd_const RTs);
+        val bnf_in = mk_in setRs (map (mk_bnf_t Ds RTs) bnf_sets) CR;
+      in
+        mk_rel_compp (mk_conversep (mk_Grp bnf_in map1), mk_Grp bnf_in map2)
+        |> fold_rev Term.absfree Rs'
+      end;
+
+    val rel_rhs = the_default OO_Grp rel_rhs_opt;
+
+    val rel_bind_def =
+      (fn () => def_qualify (if Binding.is_empty rel_b then mk_prefix_binding relN else rel_b),
+         rel_rhs);
+
+    val wit_rhss =
+      if null wit_rhss then
+        [fold_rev Term.absdummy As (Term.list_comb (mk_bnf_map Ds As As,
+          map2 (fn T => fn i => Term.absdummy T (Bound i)) As (live downto 1)) $
+          Const (@{const_name undefined}, CA))]
+      else wit_rhss;
+    val nwits = length wit_rhss;
+    val wit_binds_defs =
+      let
+        val bs = if nwits = 1 then [fn () => def_qualify (mk_prefix_binding witN)]
+          else map (fn i => fn () => def_qualify (mk_prefix_binding (mk_witN i))) (1 upto nwits);
+      in bs ~~ wit_rhss end;
+
+    val (((bnf_rel_term, raw_rel_def), (bnf_wit_terms, raw_wit_defs)), (lthy, lthy_old)) =
+      lthy
+      |> maybe_define (is_some rel_rhs_opt) rel_bind_def
+      ||>> apfst split_list o fold_map (maybe_define (not (null wit_rhss))) wit_binds_defs
+      ||> `(maybe_restore lthy);
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+    val bnf_rel_def = Morphism.thm phi raw_rel_def;
+    val bnf_rel = Morphism.term phi bnf_rel_term;
+    fun mk_bnf_rel Ds As' Bs' =
+      normalize_rel lthy (map2 mk_pred2T As' Bs') (mk_bnf_T Ds As' Calpha) (mk_bnf_T Ds Bs' Calpha)
+        bnf_rel;
+
+    val bnf_wit_defs = map (Morphism.thm phi) raw_wit_defs;
+    val bnf_wits =
+      map (normalize_wit Calpha_params Calpha alphas o Morphism.term phi) bnf_wit_terms;
+
+    fun mk_OO_Grp Ds' As' Bs' =
+      Term.subst_atomic_types ((Ds ~~ Ds') @ (As ~~ As') @ (Bs ~~ Bs')) OO_Grp;
+  in
+    (((alphas, betas, deads, Calpha),
+     (bnf_map, bnf_sets, bnf_bd, bnf_wits, bnf_rel),
+     (bnf_map_def, bnf_set_defs, bnf_bd_def, bnf_wit_defs, bnf_rel_def),
+     (mk_bnf_map, mk_bnf_t, mk_bnf_T, mk_bnf_rel, mk_OO_Grp)), lthy)
+  end;
+
+fun prepare_def const_policy mk_fact_policy qualify prep_typ prep_term Ds_opt map_b rel_b set_bs
+  ((((((raw_bnf_b, raw_bnf_T), raw_map), raw_sets), raw_bd), raw_wits), raw_rel_opt)
+  no_defs_lthy =
+  let
+    val fact_policy = mk_fact_policy no_defs_lthy;
+    val bnf_b = qualify raw_bnf_b;
+    val live = length raw_sets;
+
+    val T_rhs = prep_typ no_defs_lthy raw_bnf_T;
+    val map_rhs = prep_term no_defs_lthy raw_map;
+    val set_rhss = map (prep_term no_defs_lthy) raw_sets;
+    val bd_rhs = prep_term no_defs_lthy raw_bd;
+    val wit_rhss = map (prep_term no_defs_lthy) raw_wits;
+    val rel_rhs_opt = Option.map (prep_term no_defs_lthy) raw_rel_opt;
+
+    fun err T =
+      error ("Trying to register the type " ^ quote (Syntax.string_of_typ no_defs_lthy T) ^
+        " as unnamed BNF");
+
+    val (bnf_b, key) =
+      if Binding.eq_name (bnf_b, Binding.empty) then
+        (case T_rhs of
+          Type (C, Ts) => if forall (can dest_TFree) Ts
+            then (Binding.qualified_name C, C) else err T_rhs
+        | T => err T)
+      else (bnf_b, Local_Theory.full_name no_defs_lthy bnf_b);
+
+    val (((alphas, betas, deads, Calpha),
+     (bnf_map, bnf_sets, bnf_bd, bnf_wits, bnf_rel),
+     (bnf_map_def, bnf_set_defs, bnf_bd_def, bnf_wit_defs, bnf_rel_def),
+     (mk_bnf_map_Ds, mk_bnf_t_Ds, mk_bnf_T_Ds, _, mk_OO_Grp)), lthy) =
+       define_bnf_consts const_policy fact_policy Ds_opt map_b rel_b set_bs
+         ((((((bnf_b, T_rhs), map_rhs), set_rhss), bd_rhs), wit_rhss), rel_rhs_opt) no_defs_lthy;
+
+    val dead = length deads;
+
+    val ((((((As', Bs'), Cs), Ds), B1Ts), B2Ts), (Ts, T)) = lthy
+      |> mk_TFrees live
+      ||>> mk_TFrees live
+      ||>> mk_TFrees live
+      ||>> mk_TFrees dead
+      ||>> mk_TFrees live
+      ||>> mk_TFrees live
+      ||> fst o mk_TFrees 1
+      ||> the_single
+      ||> `(replicate live);
+
+    val mk_bnf_map = mk_bnf_map_Ds Ds;
+    val mk_bnf_t = mk_bnf_t_Ds Ds;
+    val mk_bnf_T = mk_bnf_T_Ds Ds;
+
+    val pred2RTs = map2 mk_pred2T As' Bs';
+    val pred2RTsAsCs = map2 mk_pred2T As' Cs;
+    val pred2RTsBsCs = map2 mk_pred2T Bs' Cs;
+    val pred2RT's = map2 mk_pred2T Bs' As';
+    val self_pred2RTs = map2 mk_pred2T As' As';
+    val transfer_domRTs = map2 mk_pred2T As' B1Ts;
+    val transfer_ranRTs = map2 mk_pred2T Bs' B2Ts;
+
+    val CA' = mk_bnf_T As' Calpha;
+    val CB' = mk_bnf_T Bs' Calpha;
+    val CC' = mk_bnf_T Cs Calpha;
+    val CB1 = mk_bnf_T B1Ts Calpha;
+    val CB2 = mk_bnf_T B2Ts Calpha;
+
+    val bnf_map_AsAs = mk_bnf_map As' As';
+    val bnf_map_AsBs = mk_bnf_map As' Bs';
+    val bnf_map_AsCs = mk_bnf_map As' Cs;
+    val bnf_map_BsCs = mk_bnf_map Bs' Cs;
+    val bnf_sets_As = map (mk_bnf_t As') bnf_sets;
+    val bnf_sets_Bs = map (mk_bnf_t Bs') bnf_sets;
+    val bnf_bd_As = mk_bnf_t As' bnf_bd;
+    fun mk_bnf_rel RTs CA CB = normalize_rel lthy RTs CA CB bnf_rel;
+
+    val pre_names_lthy = lthy;
+    val (((((((((((((((fs, gs), hs), x), y), zs), ys), As),
+      As_copy), bs), Rs), Rs_copy), Ss),
+      transfer_domRs), transfer_ranRs), names_lthy) = pre_names_lthy
+      |> mk_Frees "f" (map2 (curry op -->) As' Bs')
+      ||>> mk_Frees "g" (map2 (curry op -->) Bs' Cs)
+      ||>> mk_Frees "h" (map2 (curry op -->) As' Ts)
+      ||>> yield_singleton (mk_Frees "x") CA'
+      ||>> yield_singleton (mk_Frees "y") CB'
+      ||>> mk_Frees "z" As'
+      ||>> mk_Frees "y" Bs'
+      ||>> mk_Frees "A" (map HOLogic.mk_setT As')
+      ||>> mk_Frees "A" (map HOLogic.mk_setT As')
+      ||>> mk_Frees "b" As'
+      ||>> mk_Frees "R" pred2RTs
+      ||>> mk_Frees "R" pred2RTs
+      ||>> mk_Frees "S" pred2RTsBsCs
+      ||>> mk_Frees "R" transfer_domRTs
+      ||>> mk_Frees "S" transfer_ranRTs;
+
+    val fs_copy = map2 (retype_free o fastype_of) fs gs;
+    val x_copy = retype_free CA' y;
+
+    val rel = mk_bnf_rel pred2RTs CA' CB';
+    val relAsAs = mk_bnf_rel self_pred2RTs CA' CA';
+    val bnf_wit_As = map (apsnd (mk_bnf_t As')) bnf_wits;
+
+    val map_id0_goal =
+      let val bnf_map_app_id = Term.list_comb (bnf_map_AsAs, map HOLogic.id_const As') in
+        mk_Trueprop_eq (bnf_map_app_id, HOLogic.id_const CA')
+      end;
+
+    val map_comp0_goal =
+      let
+        val bnf_map_app_comp = Term.list_comb (bnf_map_AsCs, map2 (curry HOLogic.mk_comp) gs fs);
+        val comp_bnf_map_app = HOLogic.mk_comp
+          (Term.list_comb (bnf_map_BsCs, gs), Term.list_comb (bnf_map_AsBs, fs));
+      in
+        fold_rev Logic.all (fs @ gs) (mk_Trueprop_eq (bnf_map_app_comp, comp_bnf_map_app))
+      end;
+
+    fun mk_map_cong_prem x z set f f_copy =
+      Logic.all z (Logic.mk_implies
+        (HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set $ x)),
+        mk_Trueprop_eq (f $ z, f_copy $ z)));
+
+    val map_cong0_goal =
+      let
+        val prems = map4 (mk_map_cong_prem x) zs bnf_sets_As fs fs_copy;
+        val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x,
+          Term.list_comb (bnf_map_AsBs, fs_copy) $ x);
+      in
+        fold_rev Logic.all (x :: fs @ fs_copy) (Logic.list_implies (prems, eq))
+      end;
+
+    val set_map0s_goal =
+      let
+        fun mk_goal setA setB f =
+          let
+            val set_comp_map =
+              HOLogic.mk_comp (setB, Term.list_comb (bnf_map_AsBs, fs));
+            val image_comp_set = HOLogic.mk_comp (mk_image f, setA);
+          in
+            fold_rev Logic.all fs (mk_Trueprop_eq (set_comp_map, image_comp_set))
+          end;
+      in
+        map3 mk_goal bnf_sets_As bnf_sets_Bs fs
+      end;
+
+    val card_order_bd_goal = HOLogic.mk_Trueprop (mk_card_order bnf_bd_As);
+
+    val cinfinite_bd_goal = HOLogic.mk_Trueprop (mk_cinfinite bnf_bd_As);
+
+    val set_bds_goal =
+      let
+        fun mk_goal set =
+          Logic.all x (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (set $ x)) bnf_bd_As));
+      in
+        map mk_goal bnf_sets_As
+      end;
+
+    val relAsCs = mk_bnf_rel pred2RTsAsCs CA' CC';
+    val relBsCs = mk_bnf_rel pred2RTsBsCs CB' CC';
+    val rel_OO_lhs = Term.list_comb (relAsCs, map2 (curry mk_rel_compp) Rs Ss);
+    val rel_OO_rhs = mk_rel_compp (Term.list_comb (rel, Rs), Term.list_comb (relBsCs, Ss));
+    val le_rel_OO_goal =
+      fold_rev Logic.all (Rs @ Ss) (HOLogic.mk_Trueprop (mk_leq rel_OO_rhs rel_OO_lhs));
+
+    val rel_OO_Grp_goal = fold_rev Logic.all Rs (mk_Trueprop_eq (Term.list_comb (rel, Rs),
+      Term.list_comb (mk_OO_Grp Ds As' Bs', Rs)));
+
+    val goals = zip_axioms map_id0_goal map_comp0_goal map_cong0_goal set_map0s_goal
+      card_order_bd_goal cinfinite_bd_goal set_bds_goal le_rel_OO_goal rel_OO_Grp_goal;
+
+    fun mk_wit_goals (I, wit) =
+      let
+        val xs = map (nth bs) I;
+        fun wit_goal i =
+          let
+            val z = nth zs i;
+            val set_wit = nth bnf_sets_As i $ Term.list_comb (wit, xs);
+            val concl = HOLogic.mk_Trueprop
+              (if member (op =) I i then HOLogic.mk_eq (z, nth bs i)
+              else @{term False});
+          in
+            fold_rev Logic.all (z :: xs)
+              (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set_wit)), concl))
+          end;
+      in
+        map wit_goal (0 upto live - 1)
+      end;
+
+    val triv_wit_tac = mk_trivial_wit_tac bnf_wit_defs;
+
+    val wit_goalss =
+      (if null raw_wits then SOME triv_wit_tac else NONE, map mk_wit_goals bnf_wit_As);
+
+    fun after_qed mk_wit_thms thms lthy =
+      let
+        val (axioms, nontriv_wit_thms) = apfst (mk_axioms live) (chop (length goals) thms);
+
+        val bd_Card_order = #bd_card_order axioms RS @{thm conjunct2[OF card_order_on_Card_order]};
+        val bd_Cinfinite = @{thm conjI} OF [#bd_cinfinite axioms, bd_Card_order];
+        val bd_Cnotzero = bd_Cinfinite RS @{thm Cinfinite_Cnotzero};
+
+        fun mk_collect_set_map () =
+          let
+            val defT = mk_bnf_T Ts Calpha --> HOLogic.mk_setT T;
+            val collect_map = HOLogic.mk_comp
+              (mk_collect (map (mk_bnf_t Ts) bnf_sets) defT,
+              Term.list_comb (mk_bnf_map As' Ts, hs));
+            val image_collect = mk_collect
+              (map2 (fn h => fn set => HOLogic.mk_comp (mk_image h, set)) hs bnf_sets_As)
+              defT;
+            (*collect {set1 ... setm} o map f1 ... fm = collect {f1` o set1 ... fm` o setm}*)
+            val goal = fold_rev Logic.all hs (mk_Trueprop_eq (collect_map, image_collect));
+          in
+            Goal.prove_sorry lthy [] [] goal (K (mk_collect_set_map_tac (#set_map0 axioms)))
+            |> Thm.close_derivation
+          end;
+
+        val collect_set_map = Lazy.lazy mk_collect_set_map;
+
+        fun mk_in_mono () =
+          let
+            val prems_mono = map2 (HOLogic.mk_Trueprop oo mk_leq) As As_copy;
+            val in_mono_goal =
+              fold_rev Logic.all (As @ As_copy)
+                (Logic.list_implies (prems_mono, HOLogic.mk_Trueprop
+                  (mk_leq (mk_in As bnf_sets_As CA') (mk_in As_copy bnf_sets_As CA'))));
+          in
+            Goal.prove_sorry lthy [] [] in_mono_goal (K (mk_in_mono_tac live))
+            |> Thm.close_derivation
+          end;
+
+        val in_mono = Lazy.lazy mk_in_mono;
+
+        fun mk_in_cong () =
+          let
+            val prems_cong = map2 (curry mk_Trueprop_eq) As As_copy;
+            val in_cong_goal =
+              fold_rev Logic.all (As @ As_copy)
+                (Logic.list_implies (prems_cong,
+                  mk_Trueprop_eq (mk_in As bnf_sets_As CA', mk_in As_copy bnf_sets_As CA')));
+          in
+            Goal.prove_sorry lthy [] [] in_cong_goal
+              (K ((TRY o hyp_subst_tac lthy THEN' rtac refl) 1))
+            |> Thm.close_derivation
+          end;
+
+        val in_cong = Lazy.lazy mk_in_cong;
+
+        val map_id = Lazy.lazy (fn () => mk_map_id (#map_id0 axioms));
+        val map_comp = Lazy.lazy (fn () => mk_map_comp (#map_comp0 axioms));
+
+        fun mk_map_cong () =
+          let
+            val prem0 = mk_Trueprop_eq (x, x_copy);
+            val prems = map4 (mk_map_cong_prem x_copy) zs bnf_sets_As fs fs_copy;
+            val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x,
+              Term.list_comb (bnf_map_AsBs, fs_copy) $ x_copy);
+            val goal = fold_rev Logic.all (x :: x_copy :: fs @ fs_copy)
+              (Logic.list_implies (prem0 :: prems, eq));
+          in
+            Goal.prove_sorry lthy [] [] goal (fn _ => mk_map_cong_tac lthy (#map_cong0 axioms))
+            |> Thm.close_derivation
+          end;
+
+        val map_cong = Lazy.lazy mk_map_cong;
+
+        val set_map = map (fn thm => Lazy.lazy (fn () => mk_set_map thm)) (#set_map0 axioms);
+
+        val wit_thms =
+          if null nontriv_wit_thms then mk_wit_thms (map Lazy.force set_map) else nontriv_wit_thms;
+
+        fun mk_in_bd () =
+          let
+            val bdT = fst (dest_relT (fastype_of bnf_bd_As));
+            val bdTs = replicate live bdT;
+            val bd_bnfT = mk_bnf_T bdTs Calpha;
+            val surj_imp_ordLeq_inst = (if live = 0 then TrueI else
+              let
+                val ranTs = map (fn AT => mk_sumT (AT, HOLogic.unitT)) As';
+                val funTs = map (fn T => bdT --> T) ranTs;
+                val ran_bnfT = mk_bnf_T ranTs Calpha;
+                val (revTs, Ts) = `rev (bd_bnfT :: funTs);
+                val cTs = map (SOME o certifyT lthy) [ran_bnfT, Library.foldr1 HOLogic.mk_prodT Ts];
+                val tinst = fold (fn T => fn t => HOLogic.mk_split (Term.absdummy T t)) (tl revTs)
+                  (Term.absdummy (hd revTs) (Term.list_comb (mk_bnf_map bdTs ranTs,
+                    map Bound (live - 1 downto 0)) $ Bound live));
+                val cts = [NONE, SOME (certify lthy tinst)];
+              in
+                Drule.instantiate' cTs cts @{thm surj_imp_ordLeq}
+              end);
+            val bd = mk_cexp
+              (if live = 0 then ctwo
+                else mk_csum (Library.foldr1 (uncurry mk_csum) (map mk_card_of As)) ctwo)
+              (mk_csum bnf_bd_As (mk_card_of (HOLogic.mk_UNIV bd_bnfT)));
+            val in_bd_goal =
+              fold_rev Logic.all As
+                (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (mk_in As bnf_sets_As CA')) bd));
+          in
+            Goal.prove_sorry lthy [] [] in_bd_goal
+              (mk_in_bd_tac live surj_imp_ordLeq_inst
+                (Lazy.force map_comp) (Lazy.force map_id) (#map_cong0 axioms)
+                (map Lazy.force set_map) (#set_bd axioms) (#bd_card_order axioms)
+                bd_Card_order bd_Cinfinite bd_Cnotzero)
+            |> Thm.close_derivation
+          end;
+
+        val in_bd = Lazy.lazy mk_in_bd;
+
+        val rel_OO_Grp = #rel_OO_Grp axioms;
+        val rel_OO_Grps = no_refl [rel_OO_Grp];
+
+        fun mk_rel_Grp () =
+          let
+            val lhs = Term.list_comb (rel, map2 mk_Grp As fs);
+            val rhs = mk_Grp (mk_in As bnf_sets_As CA') (Term.list_comb (bnf_map_AsBs, fs));
+            val goal = fold_rev Logic.all (As @ fs) (mk_Trueprop_eq (lhs, rhs));
+          in
+            Goal.prove_sorry lthy [] [] goal
+              (mk_rel_Grp_tac rel_OO_Grps (#map_id0 axioms) (#map_cong0 axioms) (Lazy.force map_id)
+                (Lazy.force map_comp) (map Lazy.force set_map))
+            |> Thm.close_derivation
+          end;
+
+        val rel_Grp = Lazy.lazy mk_rel_Grp;
+
+        fun mk_rel_prems f = map2 (HOLogic.mk_Trueprop oo f) Rs Rs_copy
+        fun mk_rel_concl f = HOLogic.mk_Trueprop
+          (f (Term.list_comb (rel, Rs), Term.list_comb (rel, Rs_copy)));
+
+        fun mk_rel_mono () =
+          let
+            val mono_prems = mk_rel_prems mk_leq;
+            val mono_concl = mk_rel_concl (uncurry mk_leq);
+          in
+            Goal.prove_sorry lthy [] []
+              (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (mono_prems, mono_concl)))
+              (K (mk_rel_mono_tac rel_OO_Grps (Lazy.force in_mono)))
+            |> Thm.close_derivation
+          end;
+
+        fun mk_rel_cong () =
+          let
+            val cong_prems = mk_rel_prems (curry HOLogic.mk_eq);
+            val cong_concl = mk_rel_concl HOLogic.mk_eq;
+          in
+            Goal.prove_sorry lthy [] []
+              (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (cong_prems, cong_concl)))
+              (fn _ => (TRY o hyp_subst_tac lthy THEN' rtac refl) 1)
+            |> Thm.close_derivation
+          end;
+
+        val rel_mono = Lazy.lazy mk_rel_mono;
+        val rel_cong = Lazy.lazy mk_rel_cong;
+
+        fun mk_rel_eq () =
+          Goal.prove_sorry lthy [] []
+            (mk_Trueprop_eq (Term.list_comb (relAsAs, map HOLogic.eq_const As'),
+              HOLogic.eq_const CA'))
+            (K (mk_rel_eq_tac live (Lazy.force rel_Grp) (Lazy.force rel_cong) (#map_id0 axioms)))
+          |> Thm.close_derivation;
+
+        val rel_eq = Lazy.lazy mk_rel_eq;
+
+        fun mk_rel_conversep () =
+          let
+            val relBsAs = mk_bnf_rel pred2RT's CB' CA';
+            val lhs = Term.list_comb (relBsAs, map mk_conversep Rs);
+            val rhs = mk_conversep (Term.list_comb (rel, Rs));
+            val le_goal = fold_rev Logic.all Rs (HOLogic.mk_Trueprop (mk_leq lhs rhs));
+            val le_thm = Goal.prove_sorry lthy [] [] le_goal
+              (mk_rel_conversep_le_tac rel_OO_Grps (Lazy.force rel_eq) (#map_cong0 axioms)
+                (Lazy.force map_comp) (map Lazy.force set_map))
+              |> Thm.close_derivation
+            val goal = fold_rev Logic.all Rs (mk_Trueprop_eq (lhs, rhs));
+          in
+            Goal.prove_sorry lthy [] [] goal
+              (K (mk_rel_conversep_tac le_thm (Lazy.force rel_mono)))
+            |> Thm.close_derivation
+          end;
+
+        val rel_conversep = Lazy.lazy mk_rel_conversep;
+
+        fun mk_rel_OO () =
+          Goal.prove_sorry lthy [] []
+            (fold_rev Logic.all (Rs @ Ss) (HOLogic.mk_Trueprop (mk_leq rel_OO_lhs rel_OO_rhs)))
+            (mk_rel_OO_le_tac rel_OO_Grps (Lazy.force rel_eq) (#map_cong0 axioms)
+              (Lazy.force map_comp) (map Lazy.force set_map))
+          |> Thm.close_derivation
+          |> (fn thm => @{thm antisym} OF [thm, #le_rel_OO axioms]);
+
+        val rel_OO = Lazy.lazy mk_rel_OO;
+
+        fun mk_in_rel () = trans OF [rel_OO_Grp, @{thm OO_Grp_alt}] RS @{thm predicate2_eqD};
+
+        val in_rel = Lazy.lazy mk_in_rel;
+
+        fun mk_rel_flip () =
+          let
+            val rel_conversep_thm = Lazy.force rel_conversep;
+            val cts = map (SOME o certify lthy) Rs;
+            val rel_conversep_thm' = cterm_instantiate_pos cts rel_conversep_thm;
+          in
+            unfold_thms lthy @{thms conversep_iff} (rel_conversep_thm' RS @{thm predicate2_eqD})
+            |> singleton (Proof_Context.export names_lthy pre_names_lthy)
+          end;
+
+        val rel_flip = Lazy.lazy mk_rel_flip;
+
+        fun mk_rel_mono_strong () =
+          let
+            fun mk_prem setA setB R S a b =
+              HOLogic.mk_Trueprop
+                (mk_Ball (setA $ x) (Term.absfree (dest_Free a)
+                  (mk_Ball (setB $ y) (Term.absfree (dest_Free b)
+                    (HOLogic.mk_imp (R $ a $ b, S $ a $ b))))));
+            val prems = HOLogic.mk_Trueprop (Term.list_comb (rel, Rs) $ x $ y) :: 
+              map6 mk_prem bnf_sets_As bnf_sets_Bs Rs Rs_copy zs ys;
+            val concl = HOLogic.mk_Trueprop (Term.list_comb (rel, Rs_copy) $ x $ y);
+          in
+            Goal.prove_sorry lthy [] []
+              (fold_rev Logic.all (x :: y :: Rs @ Rs_copy) (Logic.list_implies (prems, concl)))
+              (mk_rel_mono_strong_tac (Lazy.force in_rel) (map Lazy.force set_map))
+            |> Thm.close_derivation
+          end;
+
+        val rel_mono_strong = Lazy.lazy mk_rel_mono_strong;
+
+        fun mk_map_transfer () =
+          let
+            val rels = map2 mk_fun_rel transfer_domRs transfer_ranRs;
+            val rel = mk_fun_rel
+              (Term.list_comb (mk_bnf_rel transfer_domRTs CA' CB1, transfer_domRs))
+              (Term.list_comb (mk_bnf_rel transfer_ranRTs CB' CB2, transfer_ranRs));
+            val concl = HOLogic.mk_Trueprop
+              (fold_rev mk_fun_rel rels rel $ bnf_map_AsBs $ mk_bnf_map B1Ts B2Ts);
+          in
+            Goal.prove_sorry lthy [] []
+              (fold_rev Logic.all (transfer_domRs @ transfer_ranRs) concl)
+              (mk_map_transfer_tac (Lazy.force rel_mono) (Lazy.force in_rel)
+                (map Lazy.force set_map) (#map_cong0 axioms) (Lazy.force map_comp))
+            |> Thm.close_derivation
+          end;
+
+        val map_transfer = Lazy.lazy mk_map_transfer;
+
+        val defs = mk_defs bnf_map_def bnf_set_defs bnf_rel_def;
+
+        val facts = mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_map in_bd in_cong
+          in_mono in_rel map_comp map_cong map_id map_transfer rel_eq rel_flip set_map
+          rel_cong rel_mono rel_mono_strong rel_Grp rel_conversep rel_OO;
+
+        val wits = map2 mk_witness bnf_wits wit_thms;
+
+        val bnf_rel =
+          Term.subst_atomic_types ((Ds ~~ deads) @ (As' ~~ alphas) @ (Bs' ~~ betas)) rel;
+
+        val bnf = mk_bnf bnf_b Calpha live alphas betas dead deads bnf_map bnf_sets bnf_bd axioms
+          defs facts wits bnf_rel;
+      in
+        (bnf, lthy |> note_bnf_thms fact_policy qualify bnf_b bnf)
+      end;
+
+    val one_step_defs =
+      no_reflexive (bnf_map_def :: bnf_bd_def :: bnf_set_defs @ bnf_wit_defs @ [bnf_rel_def]);
+  in
+    (key, goals, wit_goalss, after_qed, lthy, one_step_defs)
+  end;
+
+fun register_bnf key (bnf, lthy) =
+  (bnf, Local_Theory.declaration {syntax = false, pervasive = true}
+    (fn phi => Data.map (Symtab.default (key, morph_bnf phi bnf))) lthy);
+
+fun bnf_def const_policy fact_policy qualify tacs wit_tac Ds map_b rel_b set_bs =
+  (fn (_, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, one_step_defs) =>
+  let
+    fun mk_wits_tac set_maps =
+      K (TRYALL Goal.conjunction_tac) THEN'
+      (case triv_tac_opt of
+        SOME tac => tac set_maps
+      | NONE => fn {context = ctxt, prems} =>
+          unfold_thms_tac ctxt one_step_defs THEN wit_tac {context = ctxt, prems = prems});
+    val wit_goals = map Logic.mk_conjunction_balanced wit_goalss;
+    fun mk_wit_thms set_maps =
+      Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals) (mk_wits_tac set_maps)
+        |> Conjunction.elim_balanced (length wit_goals)
+        |> map2 (Conjunction.elim_balanced o length) wit_goalss
+        |> map (map (Thm.close_derivation o Thm.forall_elim_vars 0));
+  in
+    map2 (Thm.close_derivation oo Goal.prove_sorry lthy [] [])
+      goals (map (fn tac => fn {context = ctxt, prems} =>
+        unfold_thms_tac ctxt one_step_defs THEN tac {context = ctxt, prems = prems}) tacs)
+    |> (fn thms => after_qed mk_wit_thms (map single thms) lthy)
+  end) oo prepare_def const_policy fact_policy qualify (K I) (K I) Ds map_b rel_b set_bs;
+
+val bnf_cmd = (fn (key, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, defs) =>
+  let
+    val wit_goals = map Logic.mk_conjunction_balanced wit_goalss;
+    fun mk_triv_wit_thms tac set_maps =
+      Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals)
+        (K (TRYALL Goal.conjunction_tac) THEN' tac set_maps)
+        |> Conjunction.elim_balanced (length wit_goals)
+        |> map2 (Conjunction.elim_balanced o length) wit_goalss
+        |> map (map (Thm.close_derivation o Thm.forall_elim_vars 0));
+    val (mk_wit_thms, nontriv_wit_goals) = 
+      (case triv_tac_opt of
+        NONE => (fn _ => [], map (map (rpair [])) wit_goalss)
+      | SOME tac => (mk_triv_wit_thms tac, []));
+  in
+    Proof.unfolding ([[(defs, [])]])
+      (Proof.theorem NONE (snd o register_bnf key oo after_qed mk_wit_thms)
+        (map (single o rpair []) goals @ nontriv_wit_goals) lthy)
+  end) oo prepare_def Do_Inline (user_policy Note_Some) I Syntax.read_typ Syntax.read_term NONE
+    Binding.empty Binding.empty [];
+
+fun print_bnfs ctxt =
+  let
+    fun pretty_set sets i = Pretty.block
+      [Pretty.str (mk_setN (i + 1) ^ ":"), Pretty.brk 1,
+          Pretty.quote (Syntax.pretty_term ctxt (nth sets i))];
+
+    fun pretty_bnf (key, BNF {T = T, map = map, sets = sets, bd = bd,
+      live = live, lives = lives, dead = dead, deads = deads, ...}) =
+      Pretty.big_list
+        (Pretty.string_of (Pretty.block [Pretty.str key, Pretty.str ":", Pretty.brk 1,
+          Pretty.quote (Syntax.pretty_typ ctxt T)]))
+        ([Pretty.block [Pretty.str "live:", Pretty.brk 1, Pretty.str (string_of_int live),
+            Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) lives)],
+          Pretty.block [Pretty.str "dead:", Pretty.brk 1, Pretty.str (string_of_int dead),
+            Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) deads)],
+          Pretty.block [Pretty.str (mapN ^ ":"), Pretty.brk 1,
+            Pretty.quote (Syntax.pretty_term ctxt map)]] @
+          List.map (pretty_set sets) (0 upto length sets - 1) @
+          [Pretty.block [Pretty.str (bdN ^ ":"), Pretty.brk 1,
+            Pretty.quote (Syntax.pretty_term ctxt bd)]]);
+  in
+    Pretty.big_list "BNFs:" (map pretty_bnf (Symtab.dest (Data.get (Context.Proof ctxt))))
+    |> Pretty.writeln
+  end;
+
+val _ =
+  Outer_Syntax.improper_command @{command_spec "print_bnfs"}
+    "print all bounded natural functors"
+    (Scan.succeed (Toplevel.keep (print_bnfs o Toplevel.context_of)));
+
+val _ =
+  Outer_Syntax.local_theory_to_proof @{command_spec "bnf"}
+    "register a type as a bounded natural functor"
+    (parse_opt_binding_colon -- Parse.typ --|
+       (Parse.reserved "map" -- @{keyword ":"}) -- Parse.term --
+       (Scan.option ((Parse.reserved "sets" -- @{keyword ":"}) |--
+         Scan.repeat1 (Scan.unless (Parse.reserved "bd") Parse.term)) >> the_default []) --|
+       (Parse.reserved "bd" -- @{keyword ":"}) -- Parse.term --
+       (Scan.option ((Parse.reserved "wits" -- @{keyword ":"}) |--
+         Scan.repeat1 (Scan.unless (Parse.reserved "rel") Parse.term)) >> the_default []) --
+       Scan.option ((Parse.reserved "rel" -- @{keyword ":"}) |-- Parse.term)
+       >> bnf_cmd);
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_def_tactics.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,284 @@
+(*  Title:      HOL/BNF/Tools/bnf_def_tactics.ML
+    Author:     Dmitriy Traytel, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2012
+
+Tactics for definition of bounded natural functors.
+*)
+
+signature BNF_DEF_TACTICS =
+sig
+  val mk_collect_set_map_tac: thm list -> tactic
+  val mk_map_id: thm -> thm
+  val mk_map_comp: thm -> thm
+  val mk_map_cong_tac: Proof.context -> thm -> tactic
+  val mk_in_mono_tac: int -> tactic
+  val mk_set_map: thm -> thm
+
+  val mk_rel_Grp_tac: thm list -> thm -> thm -> thm -> thm -> thm list ->
+    {prems: thm list, context: Proof.context} -> tactic
+  val mk_rel_eq_tac: int -> thm -> thm -> thm -> tactic
+  val mk_rel_OO_le_tac: thm list -> thm -> thm -> thm -> thm list ->
+    {prems: thm list, context: Proof.context} -> tactic
+  val mk_rel_conversep_tac: thm -> thm -> tactic
+  val mk_rel_conversep_le_tac: thm list -> thm -> thm -> thm -> thm list ->
+    {prems: thm list, context: Proof.context} -> tactic
+  val mk_rel_mono_tac: thm list -> thm -> tactic
+  val mk_rel_mono_strong_tac: thm -> thm list -> {prems: 'a, context: Proof.context} -> tactic
+
+  val mk_map_transfer_tac: thm -> thm -> thm list -> thm -> thm ->
+    {prems: thm list, context: Proof.context} -> tactic
+
+  val mk_in_bd_tac: int -> thm -> thm -> thm -> thm -> thm list -> thm list -> thm -> thm -> thm ->
+    thm -> {prems: thm list, context: Proof.context} -> tactic
+
+  val mk_trivial_wit_tac: thm list -> thm list -> {prems: thm list, context: Proof.context} ->
+    tactic
+end;
+
+structure BNF_Def_Tactics : BNF_DEF_TACTICS =
+struct
+
+open BNF_Util
+open BNF_Tactics
+
+val ord_eq_le_trans = @{thm ord_eq_le_trans};
+val ord_le_eq_trans = @{thm ord_le_eq_trans};
+val conversep_shift = @{thm conversep_le_swap} RS iffD1;
+
+fun mk_map_id id = mk_trans (fun_cong OF [id]) @{thm id_apply};
+fun mk_map_comp comp = @{thm o_eq_dest_lhs} OF [mk_sym comp];
+fun mk_map_cong_tac ctxt cong0 =
+  (hyp_subst_tac ctxt THEN' rtac cong0 THEN'
+   REPEAT_DETERM o (dtac meta_spec THEN' etac meta_mp THEN' atac)) 1;
+fun mk_set_map set_map0 = set_map0 RS @{thm comp_eq_dest};
+fun mk_in_mono_tac n = if n = 0 then rtac subset_UNIV 1
+  else (rtac subsetI THEN'
+  rtac CollectI) 1 THEN
+  REPEAT_DETERM (eresolve_tac [CollectE, conjE] 1) THEN
+  REPEAT_DETERM_N (n - 1)
+    ((rtac conjI THEN' etac subset_trans THEN' atac) 1) THEN
+  (etac subset_trans THEN' atac) 1;
+
+fun mk_collect_set_map_tac set_map0s =
+  (rtac (@{thm collect_o} RS trans) THEN' rtac @{thm arg_cong[of _ _ collect]} THEN'
+  EVERY' (map (fn set_map0 =>
+    rtac (mk_trans @{thm image_insert} @{thm arg_cong2[of _ _ _ _ insert]}) THEN'
+    rtac set_map0) set_map0s) THEN'
+  rtac @{thm image_empty}) 1;
+
+fun mk_rel_Grp_tac rel_OO_Grps map_id0 map_cong0 map_id map_comp set_maps
+  {context = ctxt, prems = _} =
+  let
+    val n = length set_maps;
+    val rel_OO_Grps_tac = if null rel_OO_Grps then K all_tac else rtac (hd rel_OO_Grps RS trans);
+  in
+    if null set_maps then
+      unfold_thms_tac ctxt ((map_id0 RS @{thm Grp_UNIV_id}) :: rel_OO_Grps) THEN
+      rtac @{thm Grp_UNIV_idI[OF refl]} 1
+    else
+      EVERY' [rel_OO_Grps_tac, rtac @{thm antisym}, rtac @{thm predicate2I},
+        REPEAT_DETERM o
+          eresolve_tac [CollectE, exE, conjE, @{thm GrpE}, @{thm relcomppE}, @{thm conversepE}],
+        hyp_subst_tac ctxt, rtac @{thm GrpI}, rtac trans, rtac map_comp, rtac map_cong0,
+        REPEAT_DETERM_N n o EVERY' [rtac @{thm Collect_split_Grp_eqD}, etac @{thm set_mp}, atac],
+        rtac CollectI,
+        CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS ord_eq_le_trans),
+          rtac @{thm image_subsetI}, rtac @{thm Collect_split_Grp_inD}, etac @{thm set_mp}, atac])
+        set_maps,
+        rtac @{thm predicate2I}, REPEAT_DETERM o eresolve_tac [@{thm GrpE}, exE, conjE],
+        hyp_subst_tac ctxt,
+        rtac @{thm relcomppI}, rtac @{thm conversepI},
+        EVERY' (map2 (fn convol => fn map_id0 =>
+          EVERY' [rtac @{thm GrpI}, rtac (box_equals OF [map_cong0, map_comp RS sym, map_id0]),
+            REPEAT_DETERM_N n o rtac (convol RS fun_cong),
+            REPEAT_DETERM o eresolve_tac [CollectE, conjE],
+            rtac CollectI,
+            CONJ_WRAP' (fn thm =>
+              EVERY' [rtac ord_eq_le_trans, rtac thm, rtac @{thm image_subsetI},
+                rtac @{thm convol_mem_GrpI}, etac set_mp, atac])
+            set_maps])
+          @{thms fst_convol snd_convol} [map_id, refl])] 1
+  end;
+
+fun mk_rel_eq_tac n rel_Grp rel_cong map_id0 =
+  (EVERY' (rtac (rel_cong RS trans) :: replicate n (rtac @{thm eq_alt})) THEN'
+  rtac (rel_Grp RSN (2, @{thm box_equals[OF _ sym sym[OF eq_alt]]})) THEN'
+  (if n = 0 then rtac refl
+  else EVERY' [rtac @{thm arg_cong2[of _ _ _ _ "Grp"]},
+    rtac @{thm equalityI}, rtac subset_UNIV, rtac subsetI, rtac CollectI,
+    CONJ_WRAP' (K (rtac subset_UNIV)) (1 upto n), rtac map_id0])) 1;
+
+fun mk_rel_mono_tac rel_OO_Grps in_mono =
+  let
+    val rel_OO_Grps_tac = if null rel_OO_Grps then K all_tac
+      else rtac (hd rel_OO_Grps RS ord_eq_le_trans) THEN'
+        rtac (hd rel_OO_Grps RS sym RSN (2, ord_le_eq_trans));
+  in
+    EVERY' [rel_OO_Grps_tac, rtac @{thm relcompp_mono}, rtac @{thm iffD2[OF conversep_mono]},
+      rtac @{thm Grp_mono}, rtac in_mono, REPEAT_DETERM o etac @{thm Collect_split_mono},
+      rtac @{thm Grp_mono}, rtac in_mono, REPEAT_DETERM o etac @{thm Collect_split_mono}] 1
+  end;
+
+fun mk_rel_conversep_le_tac rel_OO_Grps rel_eq map_cong0 map_comp set_maps
+  {context = ctxt, prems = _} =
+  let
+    val n = length set_maps;
+    val rel_OO_Grps_tac = if null rel_OO_Grps then K all_tac
+      else rtac (hd rel_OO_Grps RS ord_eq_le_trans) THEN'
+        rtac (hd rel_OO_Grps RS sym RS @{thm arg_cong[of _ _ conversep]} RSN (2, ord_le_eq_trans));
+  in
+    if null set_maps then rtac (rel_eq RS @{thm leq_conversepI}) 1
+    else
+      EVERY' [rel_OO_Grps_tac, rtac @{thm predicate2I},
+        REPEAT_DETERM o
+          eresolve_tac [CollectE, exE, conjE, @{thm GrpE}, @{thm relcomppE}, @{thm conversepE}],
+        hyp_subst_tac ctxt, rtac @{thm conversepI}, rtac @{thm relcomppI}, rtac @{thm conversepI},
+        EVERY' (map (fn thm => EVERY' [rtac @{thm GrpI}, rtac sym, rtac trans,
+          rtac map_cong0, REPEAT_DETERM_N n o rtac thm,
+          rtac (map_comp RS sym), rtac CollectI,
+          CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS ord_eq_le_trans),
+            etac @{thm flip_pred}]) set_maps]) [@{thm snd_fst_flip}, @{thm fst_snd_flip}])] 1
+  end;
+
+fun mk_rel_conversep_tac le_conversep rel_mono =
+  EVERY' [rtac @{thm antisym}, rtac le_conversep, rtac @{thm xt1(6)}, rtac conversep_shift,
+    rtac le_conversep, rtac @{thm iffD2[OF conversep_mono]}, rtac rel_mono,
+    REPEAT_DETERM o rtac @{thm eq_refl[OF sym[OF conversep_conversep]]}] 1;
+
+fun mk_rel_OO_le_tac rel_OO_Grps rel_eq map_cong0 map_comp set_maps
+  {context = ctxt, prems = _} =
+  let
+    val n = length set_maps;
+    fun in_tac nthO_in = rtac CollectI THEN'
+        CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS ord_eq_le_trans),
+          rtac @{thm image_subsetI}, rtac nthO_in, etac set_mp, atac]) set_maps;
+    val rel_OO_Grps_tac = if null rel_OO_Grps then K all_tac
+      else rtac (hd rel_OO_Grps RS ord_eq_le_trans) THEN'
+        rtac (@{thm arg_cong2[of _ _ _ _ "op OO"]} OF (replicate 2 (hd rel_OO_Grps RS sym)) RSN
+          (2, ord_le_eq_trans));
+  in
+    if null set_maps then rtac (rel_eq RS @{thm leq_OOI}) 1
+    else
+      EVERY' [rel_OO_Grps_tac, rtac @{thm predicate2I},
+        REPEAT_DETERM o
+          eresolve_tac [CollectE, exE, conjE, @{thm GrpE}, @{thm relcomppE}, @{thm conversepE}],
+        hyp_subst_tac ctxt,
+        rtac @{thm relcomppI}, rtac @{thm relcomppI}, rtac @{thm conversepI}, rtac @{thm GrpI},
+        rtac trans, rtac map_comp, rtac sym, rtac map_cong0,
+        REPEAT_DETERM_N n o rtac @{thm fst_fstOp},
+        in_tac @{thm fstOp_in},
+        rtac @{thm GrpI}, rtac trans, rtac map_comp, rtac map_cong0,
+        REPEAT_DETERM_N n o EVERY' [rtac trans, rtac o_apply, 
+          rtac ballE, rtac subst,
+          rtac @{thm csquare_def}, rtac @{thm csquare_fstOp_sndOp}, atac, etac notE,
+          etac set_mp, atac],
+        in_tac @{thm fstOp_in},
+        rtac @{thm relcomppI}, rtac @{thm conversepI}, rtac @{thm GrpI},
+        rtac trans, rtac map_comp, rtac map_cong0,
+        REPEAT_DETERM_N n o rtac o_apply,
+        in_tac @{thm sndOp_in},
+        rtac @{thm GrpI}, rtac trans, rtac map_comp, rtac sym, rtac map_cong0,
+        REPEAT_DETERM_N n o rtac @{thm snd_sndOp},
+        in_tac @{thm sndOp_in}] 1
+  end;
+
+fun mk_rel_mono_strong_tac in_rel set_maps {context = ctxt, prems = _} =
+  if null set_maps then atac 1
+  else
+    unfold_tac ctxt [in_rel] THEN
+    REPEAT_DETERM (eresolve_tac [exE, CollectE, conjE] 1) THEN
+    hyp_subst_tac ctxt 1 THEN
+    unfold_tac ctxt set_maps THEN
+    EVERY' [rtac exI, rtac @{thm conjI[OF CollectI conjI[OF refl refl]]},
+      CONJ_WRAP' (K (etac @{thm Collect_split_mono_strong} THEN' atac)) set_maps] 1;
+
+fun mk_map_transfer_tac rel_mono in_rel set_maps map_cong0 map_comp
+  {context = ctxt, prems = _} =
+  let
+    val n = length set_maps;
+    val in_tac = if n = 0 then rtac UNIV_I else
+      rtac CollectI THEN' CONJ_WRAP' (fn thm =>
+        etac (thm RS
+          @{thm ord_eq_le_trans[OF _ subset_trans[OF image_mono convol_image_vimage2p]]}))
+      set_maps;
+  in
+    REPEAT_DETERM_N n (HEADGOAL (rtac @{thm fun_relI})) THEN
+    unfold_thms_tac ctxt @{thms fun_rel_iff_leq_vimage2p} THEN
+    HEADGOAL (EVERY' [rtac @{thm order_trans}, rtac rel_mono, REPEAT_DETERM_N n o atac,
+      rtac @{thm predicate2I}, dtac (in_rel RS iffD1),
+      REPEAT_DETERM o eresolve_tac [exE, CollectE, conjE], hyp_subst_tac ctxt,
+      rtac @{thm vimage2pI}, rtac (in_rel RS iffD2), rtac exI, rtac conjI, in_tac,
+      rtac conjI,
+      EVERY' (map (fn convol =>
+        rtac (box_equals OF [map_cong0, map_comp RS sym, map_comp RS sym]) THEN'
+        REPEAT_DETERM_N n o rtac (convol RS fun_cong)) @{thms fst_convol snd_convol})])
+  end;
+
+fun mk_in_bd_tac live surj_imp_ordLeq_inst map_comp map_id map_cong0 set_maps set_bds
+  bd_card_order bd_Card_order bd_Cinfinite bd_Cnotzero {context = ctxt, prems = _} =
+  if live = 0 then
+    rtac @{thm ordLeq_transitive[OF ordLeq_csum2[OF card_of_Card_order]
+      ordLeq_cexp2[OF ordLeq_refl[OF Card_order_ctwo] Card_order_csum]]} 1
+  else
+    let
+      val bd'_Cinfinite = bd_Cinfinite RS @{thm Cinfinite_csum1};
+      val inserts =
+        map (fn set_bd => 
+          iffD2 OF [@{thm card_of_ordLeq}, @{thm ordLeq_ordIso_trans} OF
+            [set_bd, bd_Card_order RS @{thm card_of_Field_ordIso} RS @{thm ordIso_symmetric}]])
+        set_bds;        
+    in
+      EVERY' [rtac (Drule.rotate_prems 1 ctrans), rtac @{thm cprod_cinfinite_bound},
+        rtac (ctrans OF @{thms ordLeq_csum2 ordLeq_cexp2}), rtac @{thm card_of_Card_order},
+        rtac @{thm ordLeq_csum2}, rtac @{thm Card_order_ctwo}, rtac @{thm Card_order_csum},
+        rtac @{thm ordIso_ordLeq_trans}, rtac @{thm cexp_cong1},
+        if live = 1 then rtac @{thm ordIso_refl[OF Card_order_csum]}
+        else
+          REPEAT_DETERM_N (live - 2) o rtac @{thm ordIso_transitive[OF csum_cong2]} THEN'
+          REPEAT_DETERM_N (live - 1) o rtac @{thm csum_csum},
+        rtac bd_Card_order, rtac (@{thm cexp_mono2_Cnotzero} RS ctrans), rtac @{thm ordLeq_csum1},
+        rtac bd_Card_order, rtac @{thm Card_order_csum}, rtac bd_Cnotzero,
+        rtac @{thm csum_Cfinite_cexp_Cinfinite},
+        rtac (if live = 1 then @{thm card_of_Card_order} else @{thm Card_order_csum}),
+        CONJ_WRAP_GEN' (rtac @{thm Cfinite_csum}) (K (rtac @{thm Cfinite_cone})) set_maps,
+        rtac bd'_Cinfinite, rtac @{thm card_of_Card_order},
+        rtac @{thm Card_order_cexp}, rtac @{thm Cinfinite_cexp}, rtac @{thm ordLeq_csum2},
+        rtac @{thm Card_order_ctwo}, rtac bd'_Cinfinite,
+        rtac (Drule.rotate_prems 1 (@{thm cprod_mono2} RSN (2, ctrans))),
+        REPEAT_DETERM_N (live - 1) o
+          (rtac (bd_Cinfinite RS @{thm cprod_cexp_csum_cexp_Cinfinite} RSN (2, ctrans)) THEN'
+           rtac @{thm ordLeq_ordIso_trans[OF cprod_mono2 ordIso_symmetric[OF cprod_cexp]]}),
+        rtac @{thm ordLeq_refl[OF Card_order_cexp]}] 1 THEN
+      unfold_thms_tac ctxt [bd_card_order RS @{thm card_order_csum_cone_cexp_def}] THEN
+      unfold_thms_tac ctxt @{thms cprod_def Field_card_of} THEN
+      EVERY' [rtac (Drule.rotate_prems 1 ctrans), rtac surj_imp_ordLeq_inst, rtac subsetI,
+        Method.insert_tac inserts, REPEAT_DETERM o dtac meta_spec,
+        REPEAT_DETERM o eresolve_tac [exE, Tactic.make_elim conjunct1], etac CollectE,
+        if live = 1 then K all_tac
+        else REPEAT_DETERM_N (live - 2) o (etac conjE THEN' rotate_tac ~1) THEN' etac conjE,
+        rtac (Drule.rotate_prems 1 @{thm image_eqI}), rtac @{thm SigmaI}, rtac @{thm UNIV_I},
+        CONJ_WRAP_GEN' (rtac @{thm SigmaI})
+          (K (etac @{thm If_the_inv_into_in_Func} THEN' atac)) set_maps,
+        rtac sym,
+        rtac (Drule.rotate_prems 1
+           ((box_equals OF [map_cong0 OF replicate live @{thm If_the_inv_into_f_f},
+             map_comp RS sym, map_id]) RSN (2, trans))),
+        REPEAT_DETERM_N (2 * live) o atac,
+        REPEAT_DETERM_N live o rtac (@{thm prod.cases} RS trans),
+        rtac refl,
+        rtac @{thm surj_imp_ordLeq}, rtac subsetI, rtac (Drule.rotate_prems 1 @{thm image_eqI}),
+        REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac CollectI,
+        CONJ_WRAP' (fn thm =>
+          rtac (thm RS ord_eq_le_trans) THEN' etac @{thm subset_trans[OF image_mono Un_upper1]})
+        set_maps,
+        rtac sym,
+        rtac (box_equals OF [map_cong0 OF replicate live @{thm fun_cong[OF sum_case_o_inj(1)]},
+           map_comp RS sym, map_id])] 1
+  end;
+
+fun mk_trivial_wit_tac wit_defs set_maps {context = ctxt, prems = _} =
+  unfold_thms_tac ctxt wit_defs THEN HEADGOAL (EVERY' (map (fn thm =>
+    dtac (thm RS equalityD1 RS set_mp) THEN' etac imageE THEN' atac) set_maps)) THEN ALLGOALS atac;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,1523 @@
+(*  Title:      HOL/BNF/Tools/bnf_fp_def_sugar.ML
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2012, 2013
+
+Sugared datatype and codatatype constructions.
+*)
+
+signature BNF_FP_DEF_SUGAR =
+sig
+  type fp_sugar =
+    {T: typ,
+     fp: BNF_FP_Util.fp_kind,
+     index: int,
+     pre_bnfs: BNF_Def.bnf list,
+     nested_bnfs: BNF_Def.bnf list,
+     nesting_bnfs: BNF_Def.bnf list,
+     fp_res: BNF_FP_Util.fp_result,
+     ctr_defss: thm list list,
+     ctr_sugars: Ctr_Sugar.ctr_sugar list,
+     co_iterss: term list list,
+     mapss: thm list list,
+     co_inducts: thm list,
+     co_iter_thmsss: thm list list list,
+     disc_co_itersss: thm list list list,
+     sel_co_iterssss: thm list list list list};
+
+  val of_fp_sugar: (fp_sugar -> 'a list) -> fp_sugar -> 'a
+  val eq_fp_sugar: fp_sugar * fp_sugar -> bool
+  val morph_fp_sugar: morphism -> fp_sugar -> fp_sugar
+  val transfer_fp_sugar: Proof.context -> fp_sugar -> fp_sugar
+  val fp_sugar_of: Proof.context -> string -> fp_sugar option
+  val fp_sugars_of: Proof.context -> fp_sugar list
+
+  val co_induct_of: 'a list -> 'a
+  val strong_co_induct_of: 'a list -> 'a
+
+  val tvar_subst: theory -> typ list -> typ list -> ((string * int) * typ) list
+  val exists_subtype_in: typ list -> typ -> bool
+  val flat_rec_arg_args: 'a list list -> 'a list
+  val flat_corec_preds_predsss_gettersss: 'a list -> 'a list list list -> 'a list list list ->
+    'a list
+  val mk_co_iter: theory -> BNF_FP_Util.fp_kind -> typ -> typ list -> term -> term
+  val nesty_bnfs: Proof.context -> typ list list list -> typ list -> BNF_Def.bnf list
+
+  type lfp_sugar_thms =
+    (thm list * thm * Args.src list)
+    * (thm list list * thm list list * Args.src list)
+
+  val morph_lfp_sugar_thms: morphism -> lfp_sugar_thms -> lfp_sugar_thms
+  val transfer_lfp_sugar_thms: Proof.context -> lfp_sugar_thms -> lfp_sugar_thms
+
+  type gfp_sugar_thms =
+    ((thm list * thm) list * Args.src list)
+    * (thm list list * thm list list * Args.src list)
+    * (thm list list * thm list list * Args.src list)
+    * (thm list list * thm list list * Args.src list)
+    * (thm list list list * thm list list list * Args.src list)
+
+  val morph_gfp_sugar_thms: morphism -> gfp_sugar_thms -> gfp_sugar_thms
+  val transfer_gfp_sugar_thms: Proof.context -> gfp_sugar_thms -> gfp_sugar_thms
+
+  val mk_co_iters_prelims: BNF_FP_Util.fp_kind -> typ list list list -> typ list -> typ list ->
+    int list -> int list list -> term list list -> Proof.context ->
+    (term list list
+     * (typ list list * typ list list list list * term list list
+        * term list list list list) list option
+     * (string * term list * term list list
+        * ((term list list * term list list list) * (typ list * typ list list)) list) option)
+    * Proof.context
+  val mk_iter_fun_arg_types: typ list list list -> int list -> int list list -> term ->
+    typ list list list list
+  val mk_coiter_fun_arg_types: typ list list list -> typ list -> int list -> term ->
+    typ list list
+    * (typ list list list list * typ list list list * typ list list list list * typ list)
+  val define_iters: string list ->
+    (typ list list * typ list list list list * term list list * term list list list list) list ->
+    (string -> binding) -> typ list -> typ list -> term list -> Proof.context ->
+    (term list * thm list) * Proof.context
+  val define_coiters: string list -> string * term list * term list list
+    * ((term list list * term list list list) * (typ list * typ list list)) list ->
+    (string -> binding) -> typ list -> typ list -> term list -> Proof.context ->
+    (term list * thm list) * Proof.context
+  val derive_induct_iters_thms_for_types: BNF_Def.bnf list ->
+    (typ list list * typ list list list list * term list list * term list list list list) list ->
+    thm -> thm list list -> BNF_Def.bnf list -> BNF_Def.bnf list -> typ list -> typ list ->
+    typ list -> typ list list list -> term list list -> thm list list -> term list list ->
+    thm list list -> local_theory -> lfp_sugar_thms
+  val derive_coinduct_coiters_thms_for_types: BNF_Def.bnf list ->
+    string * term list * term list list * ((term list list * term list list list)
+      * (typ list * typ list list)) list ->
+    thm -> thm list -> thm list -> thm list list -> BNF_Def.bnf list -> typ list -> typ list ->
+    typ list -> typ list list list -> int list list -> int list list -> int list -> thm list list ->
+    Ctr_Sugar.ctr_sugar list -> term list list -> thm list list -> (thm list -> thm list) ->
+    local_theory -> gfp_sugar_thms
+  val co_datatypes: BNF_FP_Util.fp_kind -> (mixfix list -> binding list -> binding list ->
+      binding list list -> binding list -> (string * sort) list -> typ list * typ list list ->
+      BNF_Def.bnf list -> local_theory -> BNF_FP_Util.fp_result * local_theory) ->
+    (bool * (bool * bool)) * (((((binding * (typ * sort)) list * binding) * (binding * binding))
+      * mixfix) * ((((binding * binding) * (binding * typ) list) * (binding * term) list) *
+        mixfix) list) list ->
+    local_theory -> local_theory
+  val parse_co_datatype_cmd: BNF_FP_Util.fp_kind -> (mixfix list -> binding list -> binding list ->
+      binding list list -> binding list -> (string * sort) list -> typ list * typ list list ->
+      BNF_Def.bnf list -> local_theory -> BNF_FP_Util.fp_result * local_theory) ->
+    (local_theory -> local_theory) parser
+end;
+
+structure BNF_FP_Def_Sugar : BNF_FP_DEF_SUGAR =
+struct
+
+open Ctr_Sugar
+open BNF_Util
+open BNF_Comp
+open BNF_Def
+open BNF_FP_Util
+open BNF_FP_Def_Sugar_Tactics
+
+val EqN = "Eq_";
+
+type fp_sugar =
+  {T: typ,
+   fp: fp_kind,
+   index: int,
+   pre_bnfs: bnf list,
+   nested_bnfs: bnf list,
+   nesting_bnfs: bnf list,
+   fp_res: fp_result,
+   ctr_defss: thm list list,
+   ctr_sugars: ctr_sugar list,
+   co_iterss: term list list,
+   mapss: thm list list,
+   co_inducts: thm list,
+   co_iter_thmsss: thm list list list,
+   disc_co_itersss: thm list list list,
+   sel_co_iterssss: thm list list list list};
+
+fun of_fp_sugar f (fp_sugar as ({index, ...}: fp_sugar)) = nth (f fp_sugar) index;
+
+fun eq_fp_sugar ({T = T1, fp = fp1, index = index1, fp_res = fp_res1, ...} : fp_sugar,
+    {T = T2, fp = fp2, index = index2, fp_res = fp_res2, ...} : fp_sugar) =
+  T1 = T2 andalso fp1 = fp2 andalso index1 = index2 andalso eq_fp_result (fp_res1, fp_res2);
+
+fun morph_fp_sugar phi ({T, fp, index, pre_bnfs, nested_bnfs, nesting_bnfs, fp_res, ctr_defss,
+    ctr_sugars, co_iterss, mapss, co_inducts, co_iter_thmsss, disc_co_itersss, sel_co_iterssss}
+    : fp_sugar) =
+  {T = Morphism.typ phi T, fp = fp, index = index, pre_bnfs = map (morph_bnf phi) pre_bnfs,
+    nested_bnfs = map (morph_bnf phi) nested_bnfs, nesting_bnfs = map (morph_bnf phi) nesting_bnfs,
+   fp_res = morph_fp_result phi fp_res,
+   ctr_defss = map (map (Morphism.thm phi)) ctr_defss,
+   ctr_sugars = map (morph_ctr_sugar phi) ctr_sugars,
+   co_iterss = map (map (Morphism.term phi)) co_iterss,
+   mapss = map (map (Morphism.thm phi)) mapss,
+   co_inducts = map (Morphism.thm phi) co_inducts,
+   co_iter_thmsss = map (map (map (Morphism.thm phi))) co_iter_thmsss,
+   disc_co_itersss = map (map (map (Morphism.thm phi))) disc_co_itersss,
+   sel_co_iterssss = map (map (map (map (Morphism.thm phi)))) sel_co_iterssss};
+
+val transfer_fp_sugar =
+  morph_fp_sugar o Morphism.transfer_morphism o Proof_Context.theory_of;
+
+structure Data = Generic_Data
+(
+  type T = fp_sugar Symtab.table;
+  val empty = Symtab.empty;
+  val extend = I;
+  val merge = Symtab.merge eq_fp_sugar;
+);
+
+fun fp_sugar_of ctxt =
+  Symtab.lookup (Data.get (Context.Proof ctxt))
+  #> Option.map (transfer_fp_sugar ctxt);
+
+fun fp_sugars_of ctxt =
+  Symtab.fold (cons o transfer_fp_sugar ctxt o snd) (Data.get (Context.Proof ctxt)) [];
+
+fun co_induct_of (i :: _) = i;
+fun strong_co_induct_of [_, s] = s;
+
+(* TODO: register "sum" and "prod" as datatypes to enable N2M reduction for them *)
+
+fun register_fp_sugar key fp_sugar =
+  Local_Theory.declaration {syntax = false, pervasive = true}
+    (fn phi => Data.map (Symtab.default (key, morph_fp_sugar phi fp_sugar)));
+
+fun register_fp_sugars fp pre_bnfs nested_bnfs nesting_bnfs (fp_res as {Ts, ...}) ctr_defss
+    ctr_sugars co_iterss mapss co_inducts co_iter_thmsss disc_co_itersss sel_co_iterssss lthy =
+  (0, lthy)
+  |> fold (fn T as Type (s, _) => fn (kk, lthy) => (kk + 1,
+    register_fp_sugar s {T = T, fp = fp, index = kk, pre_bnfs = pre_bnfs,
+        nested_bnfs = nested_bnfs, nesting_bnfs = nesting_bnfs, fp_res = fp_res,
+        ctr_defss = ctr_defss, ctr_sugars = ctr_sugars, co_iterss = co_iterss, mapss = mapss,
+        co_inducts = co_inducts, co_iter_thmsss = co_iter_thmsss, disc_co_itersss = disc_co_itersss,
+        sel_co_iterssss = sel_co_iterssss}
+      lthy)) Ts
+  |> snd;
+
+(* This function could produce clashes in contrived examples (e.g., "x.A", "x.x_A", "y.A"). *)
+fun quasi_unambiguous_case_names names =
+  let
+    val ps = map (`Long_Name.base_name) names;
+    val dups = Library.duplicates (op =) (map fst ps);
+    fun underscore s =
+      let val ss = space_explode Long_Name.separator s in
+        space_implode "_" (drop (length ss - 2) ss)
+      end;
+  in
+    map (fn (base, full) => if member (op =) dups base then underscore full else base) ps
+  end;
+
+val id_def = @{thm id_def};
+val mp_conj = @{thm mp_conj};
+
+val nitpicksimp_attrs = @{attributes [nitpick_simp]};
+val code_nitpicksimp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs;
+val simp_attrs = @{attributes [simp]};
+
+fun tvar_subst thy Ts Us =
+  Vartab.fold (cons o apsnd snd) (fold (Sign.typ_match thy) (Ts ~~ Us) Vartab.empty) [];
+
+val exists_subtype_in = Term.exists_subtype o member (op =);
+
+val lists_bmoc = fold (fn xs => fn t => Term.list_comb (t, xs));
+
+fun flat_rec_arg_args xss =
+  (* FIXME (once the old datatype package is phased out): The first line below gives the preferred
+     order. The second line is for compatibility with the old datatype package. *)
+(*
+  flat xss
+*)
+  map hd xss @ maps tl xss;
+
+fun flat_corec_predss_getterss qss fss = maps (op @) (qss ~~ fss);
+
+fun flat_corec_preds_predsss_gettersss [] [qss] [fss] = flat_corec_predss_getterss qss fss
+  | flat_corec_preds_predsss_gettersss (p :: ps) (qss :: qsss) (fss :: fsss) =
+    p :: flat_corec_predss_getterss qss fss @ flat_corec_preds_predsss_gettersss ps qsss fsss;
+
+fun mk_tupled_fun x f xs =
+  if xs = [x] then f else HOLogic.tupled_lambda x (Term.list_comb (f, xs));
+
+fun mk_uncurried2_fun f xss =
+  mk_tupled_fun (HOLogic.mk_tuple (map HOLogic.mk_tuple xss)) f (flat_rec_arg_args xss);
+
+fun mk_flip (x, Type (_, [T1, Type (_, [T2, T3])])) =
+  Abs ("x", T1, Abs ("y", T2, Var (x, T2 --> T1 --> T3) $ Bound 0 $ Bound 1));
+
+fun flip_rels lthy n thm =
+  let
+    val Rs = Term.add_vars (prop_of thm) [];
+    val Rs' = rev (drop (length Rs - n) Rs);
+    val cRs = map (fn f => (certify lthy (Var f), certify lthy (mk_flip f))) Rs';
+  in
+    Drule.cterm_instantiate cRs thm
+  end;
+
+fun mk_ctor_or_dtor get_T Ts t =
+  let val Type (_, Ts0) = get_T (fastype_of t) in
+    Term.subst_atomic_types (Ts0 ~~ Ts) t
+  end;
+
+val mk_ctor = mk_ctor_or_dtor range_type;
+val mk_dtor = mk_ctor_or_dtor domain_type;
+
+fun mk_co_iter thy fp fpT Cs t =
+  let
+    val (f_Cs, Type (_, [prebody, body])) = strip_fun_type (fastype_of t);
+    val fpT0 = fp_case fp prebody body;
+    val Cs0 = distinct (op =) (map (fp_case fp body_type domain_type) f_Cs);
+    val rho = tvar_subst thy (fpT0 :: Cs0) (fpT :: Cs);
+  in
+    Term.subst_TVars rho t
+  end;
+
+fun mk_co_iters thy fp fpTs Cs ts0 =
+  let
+    val nn = length fpTs;
+    val (fpTs0, Cs0) =
+      map ((fp = Greatest_FP ? swap) o dest_funT o snd o strip_typeN nn o fastype_of) ts0
+      |> split_list;
+    val rho = tvar_subst thy (fpTs0 @ Cs0) (fpTs @ Cs);
+  in
+    map (Term.subst_TVars rho) ts0
+  end;
+
+val mk_fp_iter_fun_types = binder_fun_types o fastype_of;
+
+fun unzip_recT (Type (@{type_name prod}, _)) T = [T]
+  | unzip_recT _ (Type (@{type_name prod}, Ts)) = Ts
+  | unzip_recT _ T = [T];
+
+fun unzip_corecT (Type (@{type_name sum}, _)) T = [T]
+  | unzip_corecT _ (Type (@{type_name sum}, Ts)) = Ts
+  | unzip_corecT _ T = [T];
+
+fun liveness_of_fp_bnf n bnf =
+  (case T_of_bnf bnf of
+    Type (_, Ts) => map (not o member (op =) (deads_of_bnf bnf)) Ts
+  | _ => replicate n false);
+
+fun cannot_merge_types () = error "Mutually recursive types must have the same type parameters";
+
+fun merge_type_arg T T' = if T = T' then T else cannot_merge_types ();
+
+fun merge_type_args (As, As') =
+  if length As = length As' then map2 merge_type_arg As As' else cannot_merge_types ();
+
+fun reassoc_conjs thm =
+  reassoc_conjs (thm RS @{thm conj_assoc[THEN iffD1]})
+  handle THM _ => thm;
+
+fun type_args_named_constrained_of ((((ncAs, _), _), _), _) = ncAs;
+fun type_binding_of ((((_, b), _), _), _) = b;
+fun map_binding_of (((_, (b, _)), _), _) = b;
+fun rel_binding_of (((_, (_, b)), _), _) = b;
+fun mixfix_of ((_, mx), _) = mx;
+fun ctr_specs_of (_, ctr_specs) = ctr_specs;
+
+fun disc_of ((((disc, _), _), _), _) = disc;
+fun ctr_of ((((_, ctr), _), _), _) = ctr;
+fun args_of (((_, args), _), _) = args;
+fun defaults_of ((_, ds), _) = ds;
+fun ctr_mixfix_of (_, mx) = mx;
+
+fun add_nesty_bnf_names Us =
+  let
+    fun add (Type (s, Ts)) ss =
+        let val (needs, ss') = fold_map add Ts ss in
+          if exists I needs then (true, insert (op =) s ss') else (false, ss')
+        end
+      | add T ss = (member (op =) Us T, ss);
+  in snd oo add end;
+
+fun nesty_bnfs ctxt ctr_Tsss Us =
+  map_filter (bnf_of ctxt) (fold (fold (fold (add_nesty_bnf_names Us))) ctr_Tsss []);
+
+fun indexify proj xs f p = f (find_index (curry (op =) (proj p)) xs) p;
+
+type lfp_sugar_thms =
+  (thm list * thm * Args.src list)
+  * (thm list list * thm list list * Args.src list)
+
+fun morph_lfp_sugar_thms phi ((inducts, induct, induct_attrs), (foldss, recss, iter_attrs)) =
+  ((map (Morphism.thm phi) inducts, Morphism.thm phi induct, induct_attrs),
+   (map (map (Morphism.thm phi)) foldss, map (map (Morphism.thm phi)) recss, iter_attrs));
+
+val transfer_lfp_sugar_thms =
+  morph_lfp_sugar_thms o Morphism.transfer_morphism o Proof_Context.theory_of;
+
+type gfp_sugar_thms =
+  ((thm list * thm) list * Args.src list)
+  * (thm list list * thm list list * Args.src list)
+  * (thm list list * thm list list * Args.src list)
+  * (thm list list * thm list list * Args.src list)
+  * (thm list list list * thm list list list * Args.src list);
+
+fun morph_gfp_sugar_thms phi ((coinducts_pairs, coinduct_attrs),
+    (unfoldss, corecss, coiter_attrs), (disc_unfoldss, disc_corecss, disc_iter_attrs),
+    (disc_unfold_iffss, disc_corec_iffss, disc_iter_iff_attrs),
+    (sel_unfoldsss, sel_corecsss, sel_iter_attrs)) =
+  ((map (apfst (map (Morphism.thm phi)) o apsnd (Morphism.thm phi)) coinducts_pairs,
+    coinduct_attrs),
+   (map (map (Morphism.thm phi)) unfoldss, map (map (Morphism.thm phi)) corecss, coiter_attrs),
+   (map (map (Morphism.thm phi)) disc_unfoldss, map (map (Morphism.thm phi)) disc_corecss,
+    disc_iter_attrs),
+   (map (map (Morphism.thm phi)) disc_unfold_iffss, map (map (Morphism.thm phi)) disc_corec_iffss,
+    disc_iter_iff_attrs),
+   (map (map (map (Morphism.thm phi))) sel_unfoldsss,
+    map (map (map (Morphism.thm phi))) sel_corecsss, sel_iter_attrs));
+
+val transfer_gfp_sugar_thms =
+  morph_gfp_sugar_thms o Morphism.transfer_morphism o Proof_Context.theory_of;
+
+fun mk_iter_fun_arg_types0 n ms = map2 dest_tupleT ms o dest_sumTN_balanced n o domain_type;
+
+fun mk_iter_fun_arg_types ctr_Tsss ns mss =
+  mk_fp_iter_fun_types
+  #> map3 mk_iter_fun_arg_types0 ns mss
+  #> map2 (map2 (map2 unzip_recT)) ctr_Tsss;
+
+fun mk_iters_args_types ctr_Tsss Cs ns mss ctor_iter_fun_Tss lthy =
+  let
+    val Css = map2 replicate ns Cs;
+    val y_Tsss = map3 mk_iter_fun_arg_types0 ns mss (map un_fold_of ctor_iter_fun_Tss);
+    val g_Tss = map2 (fn C => map (fn y_Ts => y_Ts ---> C)) Cs y_Tsss;
+
+    val ((gss, ysss), lthy) =
+      lthy
+      |> mk_Freess "f" g_Tss
+      ||>> mk_Freesss "x" y_Tsss;
+
+    val y_Tssss = map (map (map single)) y_Tsss;
+    val yssss = map (map (map single)) ysss;
+
+    val z_Tssss =
+      map4 (fn n => fn ms => fn ctr_Tss => fn ctor_iter_fun_Ts =>
+          map3 (fn m => fn ctr_Ts => fn ctor_iter_fun_T =>
+              map2 unzip_recT ctr_Ts (dest_tupleT m ctor_iter_fun_T))
+            ms ctr_Tss (dest_sumTN_balanced n (domain_type (co_rec_of ctor_iter_fun_Ts))))
+        ns mss ctr_Tsss ctor_iter_fun_Tss;
+
+    val z_Tsss' = map (map flat_rec_arg_args) z_Tssss;
+    val h_Tss = map2 (map2 (curry (op --->))) z_Tsss' Css;
+
+    val hss = map2 (map2 retype_free) h_Tss gss;
+    val zssss_hd = map2 (map2 (map2 (retype_free o hd))) z_Tssss ysss;
+    val (zssss_tl, lthy) =
+      lthy
+      |> mk_Freessss "y" (map (map (map tl)) z_Tssss);
+    val zssss = map2 (map2 (map2 cons)) zssss_hd zssss_tl;
+  in
+    ([(g_Tss, y_Tssss, gss, yssss), (h_Tss, z_Tssss, hss, zssss)], lthy)
+  end;
+
+fun mk_coiter_fun_arg_types0 ctr_Tsss Cs ns fun_Ts =
+  let
+    (*avoid "'a itself" arguments in coiterators*)
+    fun repair_arity [[]] = [[@{typ unit}]]
+      | repair_arity Tss = Tss;
+
+    val ctr_Tsss' = map repair_arity ctr_Tsss;
+    val f_sum_prod_Ts = map range_type fun_Ts;
+    val f_prod_Tss = map2 dest_sumTN_balanced ns f_sum_prod_Ts;
+    val f_Tsss = map2 (map2 (dest_tupleT o length)) ctr_Tsss' f_prod_Tss;
+    val f_Tssss = map3 (fn C => map2 (map2 (map (curry (op -->) C) oo unzip_corecT)))
+      Cs ctr_Tsss' f_Tsss;
+    val q_Tssss = map (map (map (fn [_] => [] | [_, T] => [mk_pred1T (domain_type T)]))) f_Tssss;
+  in
+    (q_Tssss, f_Tsss, f_Tssss, f_sum_prod_Ts)
+  end;
+
+fun mk_coiter_p_pred_types Cs ns = map2 (fn n => replicate (Int.max (0, n - 1)) o mk_pred1T) ns Cs;
+
+fun mk_coiter_fun_arg_types ctr_Tsss Cs ns dtor_coiter =
+  (mk_coiter_p_pred_types Cs ns,
+   mk_fp_iter_fun_types dtor_coiter |> mk_coiter_fun_arg_types0 ctr_Tsss Cs ns);
+
+fun mk_coiters_args_types ctr_Tsss Cs ns dtor_coiter_fun_Tss lthy =
+  let
+    val p_Tss = mk_coiter_p_pred_types Cs ns;
+
+    fun mk_types get_Ts =
+      let
+        val fun_Ts = map get_Ts dtor_coiter_fun_Tss;
+        val (q_Tssss, f_Tsss, f_Tssss, f_sum_prod_Ts) = mk_coiter_fun_arg_types0 ctr_Tsss Cs ns fun_Ts;
+        val pf_Tss = map3 flat_corec_preds_predsss_gettersss p_Tss q_Tssss f_Tssss;
+      in
+        (q_Tssss, f_Tsss, f_Tssss, (f_sum_prod_Ts, pf_Tss))
+      end;
+
+    val (r_Tssss, g_Tsss, g_Tssss, unfold_types) = mk_types un_fold_of;
+    val (s_Tssss, h_Tsss, h_Tssss, corec_types) = mk_types co_rec_of;
+
+    val ((((Free (z, _), cs), pss), gssss), lthy) =
+      lthy
+      |> yield_singleton (mk_Frees "z") dummyT
+      ||>> mk_Frees "a" Cs
+      ||>> mk_Freess "p" p_Tss
+      ||>> mk_Freessss "g" g_Tssss;
+    val rssss = map (map (map (fn [] => []))) r_Tssss;
+
+    val hssss_hd = map2 (map2 (map2 (fn T :: _ => fn [g] => retype_free T g))) h_Tssss gssss;
+    val ((sssss, hssss_tl), lthy) =
+      lthy
+      |> mk_Freessss "q" s_Tssss
+      ||>> mk_Freessss "h" (map (map (map tl)) h_Tssss);
+    val hssss = map2 (map2 (map2 cons)) hssss_hd hssss_tl;
+
+    val cpss = map2 (map o rapp) cs pss;
+
+    fun build_sum_inj mk_inj = build_map lthy (uncurry mk_inj o dest_sumT o snd);
+
+    fun build_dtor_coiter_arg _ [] [cf] = cf
+      | build_dtor_coiter_arg T [cq] [cf, cf'] =
+        mk_If cq (build_sum_inj Inl_const (fastype_of cf, T) $ cf)
+          (build_sum_inj Inr_const (fastype_of cf', T) $ cf');
+
+    fun mk_args qssss fssss f_Tsss =
+      let
+        val pfss = map3 flat_corec_preds_predsss_gettersss pss qssss fssss;
+        val cqssss = map2 (map o map o map o rapp) cs qssss;
+        val cfssss = map2 (map o map o map o rapp) cs fssss;
+        val cqfsss = map3 (map3 (map3 build_dtor_coiter_arg)) f_Tsss cqssss cfssss;
+      in (pfss, cqfsss) end;
+
+    val unfold_args = mk_args rssss gssss g_Tsss;
+    val corec_args = mk_args sssss hssss h_Tsss;
+  in
+    ((z, cs, cpss, [(unfold_args, unfold_types), (corec_args, corec_types)]), lthy)
+  end;
+
+fun mk_co_iters_prelims fp ctr_Tsss fpTs Cs ns mss xtor_co_iterss0 lthy =
+  let
+    val thy = Proof_Context.theory_of lthy;
+
+    val (xtor_co_iter_fun_Tss, xtor_co_iterss) =
+      map (mk_co_iters thy fp fpTs Cs #> `(mk_fp_iter_fun_types o hd)) (transpose xtor_co_iterss0)
+      |> apsnd transpose o apfst transpose o split_list;
+
+    val ((iters_args_types, coiters_args_types), lthy') =
+      if fp = Least_FP then
+        mk_iters_args_types ctr_Tsss Cs ns mss xtor_co_iter_fun_Tss lthy |>> (rpair NONE o SOME)
+      else
+        mk_coiters_args_types ctr_Tsss Cs ns xtor_co_iter_fun_Tss lthy |>> (pair NONE o SOME)
+  in
+    ((xtor_co_iterss, iters_args_types, coiters_args_types), lthy')
+  end;
+
+fun mk_preds_getterss_join c cps sum_prod_T cqfss =
+  let val n = length cqfss in
+    Term.lambda c (mk_IfN sum_prod_T cps
+      (map2 (mk_InN_balanced sum_prod_T n) (map HOLogic.mk_tuple cqfss) (1 upto n)))
+  end;
+
+fun define_co_iters fp fpT Cs binding_specs lthy0 =
+  let
+    val thy = Proof_Context.theory_of lthy0;
+
+    val maybe_conceal_def_binding = Thm.def_binding
+      #> Config.get lthy0 bnf_note_all = false ? Binding.conceal;
+
+    val ((csts, defs), (lthy', lthy)) = lthy0
+      |> apfst split_list o fold_map (fn (b, rhs) =>
+        Local_Theory.define ((b, NoSyn), ((maybe_conceal_def_binding b, []), rhs))
+        #>> apsnd snd) binding_specs
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy lthy';
+
+    val csts' = map (mk_co_iter thy fp fpT Cs o Morphism.term phi) csts;
+    val defs' = map (Morphism.thm phi) defs;
+  in
+    ((csts', defs'), lthy')
+  end;
+
+fun define_iters iterNs iter_args_typess' mk_binding fpTs Cs ctor_iters lthy =
+  let
+    val nn = length fpTs;
+
+    val fpT_to_C as Type (_, [fpT, _]) = snd (strip_typeN nn (fastype_of (hd ctor_iters)));
+
+    fun generate_iter pre (_, _, fss, xssss) ctor_iter =
+      (mk_binding pre,
+       fold_rev (fold_rev Term.lambda) fss (Term.list_comb (ctor_iter,
+         map2 (mk_sum_caseN_balanced oo map2 mk_uncurried2_fun) fss xssss)));
+  in
+    define_co_iters Least_FP fpT Cs (map3 generate_iter iterNs iter_args_typess' ctor_iters) lthy
+  end;
+
+fun define_coiters coiterNs (_, cs, cpss, coiter_args_typess') mk_binding fpTs Cs dtor_coiters
+    lthy =
+  let
+    val nn = length fpTs;
+
+    val C_to_fpT as Type (_, [_, fpT]) = snd (strip_typeN nn (fastype_of (hd dtor_coiters)));
+
+    fun generate_coiter pre ((pfss, cqfsss), (f_sum_prod_Ts, pf_Tss)) dtor_coiter =
+      (mk_binding pre,
+       fold_rev (fold_rev Term.lambda) pfss (Term.list_comb (dtor_coiter,
+         map4 mk_preds_getterss_join cs cpss f_sum_prod_Ts cqfsss)));
+  in
+    define_co_iters Greatest_FP fpT Cs
+      (map3 generate_coiter coiterNs coiter_args_typess' dtor_coiters) lthy
+  end;
+
+fun derive_induct_iters_thms_for_types pre_bnfs [fold_args_types, rec_args_types] ctor_induct
+    ctor_iter_thmss nesting_bnfs nested_bnfs fpTs Cs Xs ctrXs_Tsss ctrss ctr_defss iterss iter_defss
+    lthy =
+  let
+    val iterss' = transpose iterss;
+    val iter_defss' = transpose iter_defss;
+
+    val [folds, recs] = iterss';
+    val [fold_defs, rec_defs] = iter_defss';
+
+    val ctr_Tsss = map (map (binder_types o fastype_of)) ctrss;
+
+    val nn = length pre_bnfs;
+    val ns = map length ctr_Tsss;
+    val mss = map (map length) ctr_Tsss;
+
+    val pre_map_defs = map map_def_of_bnf pre_bnfs;
+    val pre_set_defss = map set_defs_of_bnf pre_bnfs;
+    val nesting_map_idents = map (unfold_thms lthy [id_def] o map_id0_of_bnf) nesting_bnfs;
+    val nested_map_idents = map (unfold_thms lthy [id_def] o map_id0_of_bnf) nested_bnfs;
+    val nested_set_maps = maps set_map_of_bnf nested_bnfs;
+
+    val fp_b_names = map base_name_of_typ fpTs;
+
+    val ((((ps, ps'), xsss), us'), names_lthy) =
+      lthy
+      |> mk_Frees' "P" (map mk_pred1T fpTs)
+      ||>> mk_Freesss "x" ctr_Tsss
+      ||>> Variable.variant_fixes fp_b_names;
+
+    val us = map2 (curry Free) us' fpTs;
+
+    fun mk_sets_nested bnf =
+      let
+        val Type (T_name, Us) = T_of_bnf bnf;
+        val lives = lives_of_bnf bnf;
+        val sets = sets_of_bnf bnf;
+        fun mk_set U =
+          (case find_index (curry (op =) U) lives of
+            ~1 => Term.dummy
+          | i => nth sets i);
+      in
+        (T_name, map mk_set Us)
+      end;
+
+    val setss_nested = map mk_sets_nested nested_bnfs;
+
+    val (induct_thms, induct_thm) =
+      let
+        fun mk_set Ts t =
+          let val Type (_, Ts0) = domain_type (fastype_of t) in
+            Term.subst_atomic_types (Ts0 ~~ Ts) t
+          end;
+
+        fun mk_raw_prem_prems _ (x as Free (_, Type _)) (X as TFree _) =
+            [([], (find_index (curry (op =) X) Xs + 1, x))]
+          | mk_raw_prem_prems names_lthy (x as Free (s, Type (T_name, Ts0))) (Type (_, Xs_Ts0)) =
+            (case AList.lookup (op =) setss_nested T_name of
+              NONE => []
+            | SOME raw_sets0 =>
+              let
+                val (Xs_Ts, (Ts, raw_sets)) =
+                  filter (exists_subtype_in Xs o fst) (Xs_Ts0 ~~ (Ts0 ~~ raw_sets0))
+                  |> split_list ||> split_list;
+                val sets = map (mk_set Ts0) raw_sets;
+                val (ys, names_lthy') = names_lthy |> mk_Frees s Ts;
+                val xysets = map (pair x) (ys ~~ sets);
+                val ppremss = map2 (mk_raw_prem_prems names_lthy') ys Xs_Ts;
+              in
+                flat (map2 (map o apfst o cons) xysets ppremss)
+              end)
+          | mk_raw_prem_prems _ _ _ = [];
+
+        fun close_prem_prem xs t =
+          fold_rev Logic.all (map Free (drop (nn + length xs)
+            (rev (Term.add_frees t (map dest_Free xs @ ps'))))) t;
+
+        fun mk_prem_prem xs (xysets, (j, x)) =
+          close_prem_prem xs (Logic.list_implies (map (fn (x', (y, set)) =>
+              HOLogic.mk_Trueprop (HOLogic.mk_mem (y, set $ x'))) xysets,
+            HOLogic.mk_Trueprop (nth ps (j - 1) $ x)));
+
+        fun mk_raw_prem phi ctr ctr_Ts ctrXs_Ts =
+          let
+            val (xs, names_lthy') = names_lthy |> mk_Frees "x" ctr_Ts;
+            val pprems = flat (map2 (mk_raw_prem_prems names_lthy') xs ctrXs_Ts);
+          in (xs, pprems, HOLogic.mk_Trueprop (phi $ Term.list_comb (ctr, xs))) end;
+
+        fun mk_prem (xs, raw_pprems, concl) =
+          fold_rev Logic.all xs (Logic.list_implies (map (mk_prem_prem xs) raw_pprems, concl));
+
+        val raw_premss = map4 (map3 o mk_raw_prem) ps ctrss ctr_Tsss ctrXs_Tsss;
+
+        val goal =
+          Library.foldr (Logic.list_implies o apfst (map mk_prem)) (raw_premss,
+            HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry (op $)) ps us)));
+
+        val kksss = map (map (map (fst o snd) o #2)) raw_premss;
+
+        val ctor_induct' = ctor_induct OF (map mk_sumEN_tupled_balanced mss);
+
+        val thm =
+          Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
+            mk_induct_tac ctxt nn ns mss kksss (flat ctr_defss) ctor_induct' nested_set_maps
+              pre_set_defss)
+          |> singleton (Proof_Context.export names_lthy lthy)
+          |> Thm.close_derivation;
+      in
+        `(conj_dests nn) thm
+      end;
+
+    val induct_cases = quasi_unambiguous_case_names (maps (map name_of_ctr) ctrss);
+    val induct_case_names_attr = Attrib.internal (K (Rule_Cases.case_names induct_cases));
+
+    val xctrss = map2 (map2 (curry Term.list_comb)) ctrss xsss;
+
+    fun mk_iter_thmss (_, x_Tssss, fss, _) iters iter_defs ctor_iter_thms =
+      let
+        val fiters = map (lists_bmoc fss) iters;
+
+        fun mk_goal fss fiter xctr f xs fxs =
+          fold_rev (fold_rev Logic.all) (xs :: fss)
+            (mk_Trueprop_eq (fiter $ xctr, Term.list_comb (f, fxs)));
+
+        fun maybe_tick (T, U) u f =
+          if try (fst o HOLogic.dest_prodT) U = SOME T then
+            Term.lambda u (HOLogic.mk_prod (u, f $ u))
+          else
+            f;
+
+        fun build_iter (x as Free (_, T)) U =
+          if T = U then
+            x
+          else
+            build_map lthy (indexify (perhaps (try (snd o HOLogic.dest_prodT)) o snd) Cs
+              (fn kk => fn TU => maybe_tick TU (nth us kk) (nth fiters kk))) (T, U) $ x;
+
+        val fxsss = map2 (map2 (flat_rec_arg_args oo map2 (map o build_iter))) xsss x_Tssss;
+
+        val goalss = map5 (map4 o mk_goal fss) fiters xctrss fss xsss fxsss;
+
+        val tacss =
+          map2 (map o mk_iter_tac pre_map_defs (nested_map_idents @ nesting_map_idents) iter_defs)
+            ctor_iter_thms ctr_defss;
+
+        fun prove goal tac =
+          Goal.prove_sorry lthy [] [] goal (tac o #context)
+          |> Thm.close_derivation;
+      in
+        map2 (map2 prove) goalss tacss
+      end;
+
+    val fold_thmss = mk_iter_thmss fold_args_types folds fold_defs (map un_fold_of ctor_iter_thmss);
+    val rec_thmss = mk_iter_thmss rec_args_types recs rec_defs (map co_rec_of ctor_iter_thmss);
+  in
+    ((induct_thms, induct_thm, [induct_case_names_attr]),
+     (fold_thmss, rec_thmss, code_nitpicksimp_attrs @ simp_attrs))
+  end;
+
+fun derive_coinduct_coiters_thms_for_types pre_bnfs (z, cs, cpss,
+      coiters_args_types as [((pgss, crgsss), _), ((phss, cshsss), _)])
+    dtor_coinduct dtor_injects dtor_ctors dtor_coiter_thmss nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss
+    mss ns ctr_defss (ctr_sugars : ctr_sugar list) coiterss coiter_defss export_args lthy =
+  let
+    fun mk_ctor_dtor_coiter_thm dtor_inject dtor_ctor coiter =
+      iffD1 OF [dtor_inject, trans OF [coiter, dtor_ctor RS sym]];
+
+    val ctor_dtor_coiter_thmss =
+      map3 (map oo mk_ctor_dtor_coiter_thm) dtor_injects dtor_ctors dtor_coiter_thmss;
+
+    val coiterss' = transpose coiterss;
+    val coiter_defss' = transpose coiter_defss;
+
+    val [unfold_defs, corec_defs] = coiter_defss';
+
+    val nn = length pre_bnfs;
+
+    val pre_map_defs = map map_def_of_bnf pre_bnfs;
+    val pre_rel_defs = map rel_def_of_bnf pre_bnfs;
+    val nesting_map_idents = map (unfold_thms lthy [id_def] o map_id0_of_bnf) nesting_bnfs;
+    val nesting_rel_eqs = map rel_eq_of_bnf nesting_bnfs;
+
+    val fp_b_names = map base_name_of_typ fpTs;
+
+    val ctrss = map #ctrs ctr_sugars;
+    val discss = map #discs ctr_sugars;
+    val selsss = map #selss ctr_sugars;
+    val exhausts = map #exhaust ctr_sugars;
+    val disc_thmsss = map #disc_thmss ctr_sugars;
+    val discIss = map #discIs ctr_sugars;
+    val sel_thmsss = map #sel_thmss ctr_sugars;
+
+    val (((rs, us'), vs'), names_lthy) =
+      lthy
+      |> mk_Frees "R" (map (fn T => mk_pred2T T T) fpTs)
+      ||>> Variable.variant_fixes fp_b_names
+      ||>> Variable.variant_fixes (map (suffix "'") fp_b_names);
+
+    val us = map2 (curry Free) us' fpTs;
+    val udiscss = map2 (map o rapp) us discss;
+    val uselsss = map2 (map o map o rapp) us selsss;
+
+    val vs = map2 (curry Free) vs' fpTs;
+    val vdiscss = map2 (map o rapp) vs discss;
+    val vselsss = map2 (map o map o rapp) vs selsss;
+
+    val coinduct_thms_pairs =
+      let
+        val uvrs = map3 (fn r => fn u => fn v => r $ u $ v) rs us vs;
+        val uv_eqs = map2 (curry HOLogic.mk_eq) us vs;
+        val strong_rs =
+          map4 (fn u => fn v => fn uvr => fn uv_eq =>
+            fold_rev Term.lambda [u, v] (HOLogic.mk_disj (uvr, uv_eq))) us vs uvrs uv_eqs;
+
+        fun build_the_rel rs' T Xs_T =
+          build_rel lthy (fn (_, X) => nth rs' (find_index (curry (op =) X) Xs)) (T, Xs_T)
+          |> Term.subst_atomic_types (Xs ~~ fpTs);
+
+        fun build_rel_app rs' usel vsel Xs_T =
+          fold rapp [usel, vsel] (build_the_rel rs' (fastype_of usel) Xs_T);
+
+        fun mk_prem_ctr_concls rs' n k udisc usels vdisc vsels ctrXs_Ts =
+          (if k = n then [] else [HOLogic.mk_eq (udisc, vdisc)]) @
+          (if null usels then
+             []
+           else
+             [Library.foldr HOLogic.mk_imp (if n = 1 then [] else [udisc, vdisc],
+                Library.foldr1 HOLogic.mk_conj (map3 (build_rel_app rs') usels vsels ctrXs_Ts))]);
+
+        fun mk_prem_concl rs' n udiscs uselss vdiscs vselss ctrXs_Tss =
+          Library.foldr1 HOLogic.mk_conj (flat (map6 (mk_prem_ctr_concls rs' n)
+            (1 upto n) udiscs uselss vdiscs vselss ctrXs_Tss))
+          handle List.Empty => @{term True};
+
+        fun mk_prem rs' uvr u v n udiscs uselss vdiscs vselss ctrXs_Tss =
+          fold_rev Logic.all [u, v] (Logic.mk_implies (HOLogic.mk_Trueprop uvr,
+            HOLogic.mk_Trueprop (mk_prem_concl rs' n udiscs uselss vdiscs vselss ctrXs_Tss)));
+
+        val concl =
+          HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+            (map3 (fn uvr => fn u => fn v => HOLogic.mk_imp (uvr, HOLogic.mk_eq (u, v)))
+               uvrs us vs));
+
+        fun mk_goal rs' =
+          Logic.list_implies (map9 (mk_prem rs') uvrs us vs ns udiscss uselsss vdiscss vselsss
+            ctrXs_Tsss, concl);
+
+        val goals = map mk_goal [rs, strong_rs];
+
+        fun prove dtor_coinduct' goal =
+          Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
+            mk_coinduct_tac ctxt nesting_rel_eqs nn ns dtor_coinduct' pre_rel_defs dtor_ctors
+              exhausts ctr_defss disc_thmsss sel_thmsss)
+          |> singleton (Proof_Context.export names_lthy lthy)
+          |> Thm.close_derivation;
+
+        fun postproc nn thm =
+          Thm.permute_prems 0 nn
+            (if nn = 1 then thm RS mp else funpow nn (fn thm => reassoc_conjs (thm RS mp_conj)) thm)
+          |> Drule.zero_var_indexes
+          |> `(conj_dests nn);
+
+        val rel_eqs = map rel_eq_of_bnf pre_bnfs;
+        val rel_monos = map rel_mono_of_bnf pre_bnfs;
+        val dtor_coinducts =
+          [dtor_coinduct, mk_strong_coinduct_thm dtor_coinduct rel_eqs rel_monos lthy];
+      in
+        map2 (postproc nn oo prove) dtor_coinducts goals
+      end;
+
+    fun mk_coinduct_concls ms discs ctrs =
+      let
+        fun mk_disc_concl disc = [name_of_disc disc];
+        fun mk_ctr_concl 0 _ = []
+          | mk_ctr_concl _ ctor = [name_of_ctr ctor];
+        val disc_concls = map mk_disc_concl (fst (split_last discs)) @ [[]];
+        val ctr_concls = map2 mk_ctr_concl ms ctrs;
+      in
+        flat (map2 append disc_concls ctr_concls)
+      end;
+
+    val coinduct_cases = quasi_unambiguous_case_names (map (prefix EqN) fp_b_names);
+    val coinduct_conclss =
+      map3 (quasi_unambiguous_case_names ooo mk_coinduct_concls) mss discss ctrss;
+
+    fun mk_maybe_not pos = not pos ? HOLogic.mk_not;
+
+    val fcoiterss' as [gunfolds, hcorecs] =
+      map2 (fn (pfss, _) => map (lists_bmoc pfss)) (map fst coiters_args_types) coiterss';
+
+    val (unfold_thmss, corec_thmss) =
+      let
+        fun mk_goal pfss c cps fcoiter n k ctr m cfs' =
+          fold_rev (fold_rev Logic.all) ([c] :: pfss)
+            (Logic.list_implies (seq_conds (HOLogic.mk_Trueprop oo mk_maybe_not) n k cps,
+               mk_Trueprop_eq (fcoiter $ c, Term.list_comb (ctr, take m cfs'))));
+
+        fun mk_U maybe_mk_sumT =
+          typ_subst_nonatomic (map2 (fn C => fn fpT => (maybe_mk_sumT fpT C, fpT)) Cs fpTs);
+
+        fun tack z_name (c, u) f =
+          let val z = Free (z_name, mk_sumT (fastype_of u, fastype_of c)) in
+            Term.lambda z (mk_sum_case (Term.lambda u u, Term.lambda c (f $ c)) $ z)
+          end;
+
+        fun build_coiter fcoiters maybe_mk_sumT maybe_tack cqf =
+          let val T = fastype_of cqf in
+            if exists_subtype_in Cs T then
+              let val U = mk_U maybe_mk_sumT T in
+                build_map lthy (indexify snd fpTs (fn kk => fn _ =>
+                  maybe_tack (nth cs kk, nth us kk) (nth fcoiters kk))) (T, U) $ cqf
+              end
+            else
+              cqf
+          end;
+
+        val crgsss' = map (map (map (build_coiter (un_fold_of fcoiterss') (K I) (K I)))) crgsss;
+        val cshsss' = map (map (map (build_coiter (co_rec_of fcoiterss') (curry mk_sumT) (tack z))))
+          cshsss;
+
+        val unfold_goalss = map8 (map4 oooo mk_goal pgss) cs cpss gunfolds ns kss ctrss mss crgsss';
+        val corec_goalss = map8 (map4 oooo mk_goal phss) cs cpss hcorecs ns kss ctrss mss cshsss';
+
+        val unfold_tacss =
+          map3 (map oo mk_coiter_tac unfold_defs nesting_map_idents)
+            (map un_fold_of ctor_dtor_coiter_thmss) pre_map_defs ctr_defss;
+        val corec_tacss =
+          map3 (map oo mk_coiter_tac corec_defs nesting_map_idents)
+            (map co_rec_of ctor_dtor_coiter_thmss) pre_map_defs ctr_defss;
+
+        fun prove goal tac =
+          Goal.prove_sorry lthy [] [] goal (tac o #context)
+          |> Thm.close_derivation;
+
+        val unfold_thmss = map2 (map2 prove) unfold_goalss unfold_tacss;
+        val corec_thmss =
+          map2 (map2 prove) corec_goalss corec_tacss
+          |> map (map (unfold_thms lthy @{thms sum_case_if}));
+      in
+        (unfold_thmss, corec_thmss)
+      end;
+
+    val (disc_unfold_iff_thmss, disc_corec_iff_thmss) =
+      let
+        fun mk_goal c cps fcoiter n k disc =
+          mk_Trueprop_eq (disc $ (fcoiter $ c),
+            if n = 1 then @{const True}
+            else Library.foldr1 HOLogic.mk_conj (seq_conds mk_maybe_not n k cps));
+
+        val unfold_goalss = map6 (map2 oooo mk_goal) cs cpss gunfolds ns kss discss;
+        val corec_goalss = map6 (map2 oooo mk_goal) cs cpss hcorecs ns kss discss;
+
+        fun mk_case_split' cp = Drule.instantiate' [] [SOME (certify lthy cp)] @{thm case_split};
+
+        val case_splitss' = map (map mk_case_split') cpss;
+
+        val unfold_tacss =
+          map3 (map oo mk_disc_coiter_iff_tac) case_splitss' unfold_thmss disc_thmsss;
+        val corec_tacss =
+          map3 (map oo mk_disc_coiter_iff_tac) case_splitss' corec_thmss disc_thmsss;
+
+        fun prove goal tac =
+          Goal.prove_sorry lthy [] [] goal (tac o #context)
+          |> singleton export_args
+          |> singleton (Proof_Context.export names_lthy lthy)
+          |> Thm.close_derivation;
+
+        fun proves [_] [_] = []
+          | proves goals tacs = map2 prove goals tacs;
+      in
+        (map2 proves unfold_goalss unfold_tacss, map2 proves corec_goalss corec_tacss)
+      end;
+
+    fun mk_disc_coiter_thms coiters discIs = map (op RS) (coiters ~~ discIs);
+
+    val disc_unfold_thmss = map2 mk_disc_coiter_thms unfold_thmss discIss;
+    val disc_corec_thmss = map2 mk_disc_coiter_thms corec_thmss discIss;
+
+    fun mk_sel_coiter_thm coiter_thm sel sel_thm =
+      let
+        val (domT, ranT) = dest_funT (fastype_of sel);
+        val arg_cong' =
+          Drule.instantiate' (map (SOME o certifyT lthy) [domT, ranT])
+            [NONE, NONE, SOME (certify lthy sel)] arg_cong
+          |> Thm.varifyT_global;
+        val sel_thm' = sel_thm RSN (2, trans);
+      in
+        coiter_thm RS arg_cong' RS sel_thm'
+      end;
+
+    fun mk_sel_coiter_thms coiter_thmss =
+      map3 (map3 (map2 o mk_sel_coiter_thm)) coiter_thmss selsss sel_thmsss;
+
+    val sel_unfold_thmsss = mk_sel_coiter_thms unfold_thmss;
+    val sel_corec_thmsss = mk_sel_coiter_thms corec_thmss;
+
+    val coinduct_consumes_attr = Attrib.internal (K (Rule_Cases.consumes nn));
+    val coinduct_case_names_attr = Attrib.internal (K (Rule_Cases.case_names coinduct_cases));
+    val coinduct_case_concl_attrs =
+      map2 (fn casex => fn concls =>
+          Attrib.internal (K (Rule_Cases.case_conclusion (casex, concls))))
+        coinduct_cases coinduct_conclss;
+    val coinduct_case_attrs =
+      coinduct_consumes_attr :: coinduct_case_names_attr :: coinduct_case_concl_attrs;
+  in
+    ((coinduct_thms_pairs, coinduct_case_attrs),
+     (unfold_thmss, corec_thmss, code_nitpicksimp_attrs),
+     (disc_unfold_thmss, disc_corec_thmss, []),
+     (disc_unfold_iff_thmss, disc_corec_iff_thmss, simp_attrs),
+     (sel_unfold_thmsss, sel_corec_thmsss, simp_attrs))
+  end;
+
+fun define_co_datatypes prepare_constraint prepare_typ prepare_term fp construct_fp
+    (wrap_opts as (no_discs_sels, (_, rep_compat)), specs) no_defs_lthy0 =
+  let
+    (* TODO: sanity checks on arguments *)
+
+    val _ = if fp = Greatest_FP andalso no_discs_sels then
+        error "Cannot define codatatypes without discriminators and selectors"
+      else
+        ();
+
+    fun qualify mandatory fp_b_name =
+      Binding.qualify mandatory fp_b_name o (rep_compat ? Binding.qualify false rep_compat_prefix);
+
+    val nn = length specs;
+    val fp_bs = map type_binding_of specs;
+    val fp_b_names = map Binding.name_of fp_bs;
+    val fp_common_name = mk_common_name fp_b_names;
+    val map_bs = map map_binding_of specs;
+    val rel_bs = map rel_binding_of specs;
+
+    fun prepare_type_arg (_, (ty, c)) =
+      let val TFree (s, _) = prepare_typ no_defs_lthy0 ty in
+        TFree (s, prepare_constraint no_defs_lthy0 c)
+      end;
+
+    val Ass0 = map (map prepare_type_arg o type_args_named_constrained_of) specs;
+    val unsorted_Ass0 = map (map (resort_tfree HOLogic.typeS)) Ass0;
+    val unsorted_As = Library.foldr1 merge_type_args unsorted_Ass0;
+    val num_As = length unsorted_As;
+    val set_bss = map (map fst o type_args_named_constrained_of) specs;
+
+    val (((Bs0, Cs), Xs), no_defs_lthy) =
+      no_defs_lthy0
+      |> fold (Variable.declare_typ o resort_tfree dummyS) unsorted_As
+      |> mk_TFrees num_As
+      ||>> mk_TFrees nn
+      ||>> variant_tfrees fp_b_names;
+
+    fun add_fake_type spec = Typedecl.basic_typedecl (type_binding_of spec, num_As, mixfix_of spec);
+
+    val (fake_T_names, fake_lthy) = fold_map add_fake_type specs no_defs_lthy0;
+
+    val qsoty = quote o Syntax.string_of_typ fake_lthy;
+
+    val _ = (case Library.duplicates (op =) unsorted_As of [] => ()
+      | A :: _ => error ("Duplicate type parameter " ^ qsoty A ^ " in " ^ co_prefix fp ^
+          "datatype specification"));
+
+    val bad_args =
+      map (Logic.type_map (singleton (Variable.polymorphic no_defs_lthy0))) unsorted_As
+      |> filter_out Term.is_TVar;
+    val _ = null bad_args orelse
+      error ("Locally fixed type argument " ^ qsoty (hd bad_args) ^ " in " ^ co_prefix fp ^
+        "datatype specification");
+
+    val mixfixes = map mixfix_of specs;
+
+    val _ = (case Library.duplicates Binding.eq_name fp_bs of [] => ()
+      | b :: _ => error ("Duplicate type name declaration " ^ quote (Binding.name_of b)));
+
+    val ctr_specss = map ctr_specs_of specs;
+
+    val disc_bindingss = map (map disc_of) ctr_specss;
+    val ctr_bindingss =
+      map2 (fn fp_b_name => map (qualify false fp_b_name o ctr_of)) fp_b_names ctr_specss;
+    val ctr_argsss = map (map args_of) ctr_specss;
+    val ctr_mixfixess = map (map ctr_mixfix_of) ctr_specss;
+
+    val sel_bindingsss = map (map (map fst)) ctr_argsss;
+    val fake_ctr_Tsss0 = map (map (map (prepare_typ fake_lthy o snd))) ctr_argsss;
+    val raw_sel_defaultsss = map (map defaults_of) ctr_specss;
+
+    val (As :: _) :: fake_ctr_Tsss =
+      burrow (burrow (Syntax.check_typs fake_lthy)) (Ass0 :: fake_ctr_Tsss0);
+    val As' = map dest_TFree As;
+
+    val rhs_As' = fold (fold (fold Term.add_tfreesT)) fake_ctr_Tsss [];
+    val _ = (case subtract (op =) As' rhs_As' of [] => ()
+      | extras => error ("Extra type variables on right-hand side: " ^
+          commas (map (qsoty o TFree) extras)));
+
+    val fake_Ts = map (fn s => Type (s, As)) fake_T_names;
+
+    fun eq_fpT_check (T as Type (s, Ts)) (T' as Type (s', Ts')) =
+        s = s' andalso (Ts = Ts' orelse
+          error ("Wrong type arguments in " ^ co_prefix fp ^ "recursive type " ^ qsoty T ^
+            " (expected " ^ qsoty T' ^ ")"))
+      | eq_fpT_check _ _ = false;
+
+    fun freeze_fp (T as Type (s, Ts)) =
+        (case find_index (eq_fpT_check T) fake_Ts of
+          ~1 => Type (s, map freeze_fp Ts)
+        | kk => nth Xs kk)
+      | freeze_fp T = T;
+
+    val unfreeze_fp = Term.typ_subst_atomic (Xs ~~ fake_Ts);
+
+    val ctrXs_Tsss = map (map (map freeze_fp)) fake_ctr_Tsss;
+    val ctrXs_sum_prod_Ts = map (mk_sumTN_balanced o map HOLogic.mk_tupleT) ctrXs_Tsss;
+
+    val fp_eqs =
+      map dest_TFree Xs ~~ map (Term.typ_subst_atomic (As ~~ unsorted_As)) ctrXs_sum_prod_Ts;
+
+    val rhsXs_As' = fold (fold (fold Term.add_tfreesT)) ctrXs_Tsss [];
+    val _ = (case subtract (op =) rhsXs_As' As' of [] => ()
+      | extras => List.app (fn extra => warning ("Unused type variable on right-hand side of " ^
+          co_prefix fp ^ "datatype definition: " ^ qsoty (TFree extra))) extras);
+
+    val (pre_bnfs, (fp_res as {bnfs = fp_bnfs as any_fp_bnf :: _, ctors = ctors0, dtors = dtors0,
+           xtor_co_iterss = xtor_co_iterss0, xtor_co_induct, dtor_ctors, ctor_dtors, ctor_injects,
+           dtor_injects, xtor_map_thms, xtor_set_thmss, xtor_rel_thms, xtor_co_iter_thmss, ...},
+           lthy)) =
+      fp_bnf (construct_fp mixfixes map_bs rel_bs set_bss) fp_bs (map dest_TFree unsorted_As) fp_eqs
+        no_defs_lthy0
+      handle BAD_DEAD (X, X_backdrop) =>
+        (case X_backdrop of
+          Type (bad_tc, _) =>
+          let
+            val fake_T = qsoty (unfreeze_fp X);
+            val fake_T_backdrop = qsoty (unfreeze_fp X_backdrop);
+            fun register_hint () =
+              "\nUse the " ^ quote (fst (fst @{command_spec "bnf"})) ^ " command to register " ^
+              quote bad_tc ^ " as a bounded natural functor to allow nested (co)recursion through \
+              \it";
+          in
+            if is_some (bnf_of no_defs_lthy bad_tc) orelse
+               is_some (fp_sugar_of no_defs_lthy bad_tc) then
+              error ("Inadmissible " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^
+                " in type expression " ^ fake_T_backdrop)
+            else if is_some (Datatype_Data.get_info (Proof_Context.theory_of no_defs_lthy)
+                bad_tc) then
+              error ("Unsupported " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^
+                " via the old-style datatype " ^ quote bad_tc ^ " in type expression " ^
+                fake_T_backdrop ^ register_hint ())
+            else
+              error ("Unsupported " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^
+                " via type constructor " ^ quote bad_tc ^ " in type expression " ^ fake_T_backdrop ^
+                register_hint ())
+          end);
+
+    val time = time lthy;
+    val timer = time (Timer.startRealTimer ());
+
+    val nesting_bnfs = nesty_bnfs lthy ctrXs_Tsss As;
+    val nested_bnfs = nesty_bnfs lthy ctrXs_Tsss Xs;
+
+    val pre_map_defs = map map_def_of_bnf pre_bnfs;
+    val pre_set_defss = map set_defs_of_bnf pre_bnfs;
+    val pre_rel_defs = map rel_def_of_bnf pre_bnfs;
+    val nesting_set_maps = maps set_map_of_bnf nesting_bnfs;
+    val nested_set_maps = maps set_map_of_bnf nested_bnfs;
+
+    val live = live_of_bnf any_fp_bnf;
+    val _ =
+      if live = 0 andalso exists (not o Binding.is_empty) (map_bs @ rel_bs) then
+        warning "Map function and relator names ignored"
+      else
+        ();
+
+    val Bs =
+      map3 (fn alive => fn A as TFree (_, S) => fn B => if alive then resort_tfree S B else A)
+        (liveness_of_fp_bnf num_As any_fp_bnf) As Bs0;
+
+    val B_ify = Term.typ_subst_atomic (As ~~ Bs);
+
+    val ctors = map (mk_ctor As) ctors0;
+    val dtors = map (mk_dtor As) dtors0;
+
+    val fpTs = map (domain_type o fastype_of) dtors;
+
+    fun massage_simple_notes base =
+      filter_out (null o #2)
+      #> map (fn (thmN, thms, attrs) =>
+        ((qualify true base (Binding.name thmN), attrs), [(thms, [])]));
+
+    val massage_multi_notes =
+      maps (fn (thmN, thmss, attrs) =>
+        map3 (fn fp_b_name => fn Type (T_name, _) => fn thms =>
+            ((qualify true fp_b_name (Binding.name thmN), attrs T_name), [(thms, [])]))
+          fp_b_names fpTs thmss)
+      #> filter_out (null o fst o hd o snd);
+
+    val ctr_Tsss = map (map (map (Term.typ_subst_atomic (Xs ~~ fpTs)))) ctrXs_Tsss;
+    val ns = map length ctr_Tsss;
+    val kss = map (fn n => 1 upto n) ns;
+    val mss = map (map length) ctr_Tsss;
+
+    val ((xtor_co_iterss, iters_args_types, coiters_args_types), lthy') =
+      mk_co_iters_prelims fp ctr_Tsss fpTs Cs ns mss xtor_co_iterss0 lthy;
+
+    fun define_ctrs_dtrs_for_type (((((((((((((((((((((((fp_bnf, fp_b), fpT), ctor), dtor),
+            xtor_co_iters), ctor_dtor), dtor_ctor), ctor_inject), pre_map_def), pre_set_defs),
+          pre_rel_def), fp_map_thm), fp_set_thms), fp_rel_thm), n), ks), ms), ctr_bindings),
+        ctr_mixfixes), ctr_Tss), disc_bindings), sel_bindingss), raw_sel_defaultss) no_defs_lthy =
+      let
+        val fp_b_name = Binding.name_of fp_b;
+
+        val dtorT = domain_type (fastype_of ctor);
+        val ctr_prod_Ts = map HOLogic.mk_tupleT ctr_Tss;
+        val ctr_sum_prod_T = mk_sumTN_balanced ctr_prod_Ts;
+
+        val ((((w, xss), yss), u'), names_lthy) =
+          no_defs_lthy
+          |> yield_singleton (mk_Frees "w") dtorT
+          ||>> mk_Freess "x" ctr_Tss
+          ||>> mk_Freess "y" (map (map B_ify) ctr_Tss)
+          ||>> yield_singleton Variable.variant_fixes fp_b_name;
+
+        val u = Free (u', fpT);
+
+        val tuple_xs = map HOLogic.mk_tuple xss;
+        val tuple_ys = map HOLogic.mk_tuple yss;
+
+        val ctr_rhss =
+          map3 (fn k => fn xs => fn tuple_x => fold_rev Term.lambda xs (ctor $
+            mk_InN_balanced ctr_sum_prod_T n tuple_x k)) ks xss tuple_xs;
+
+        val maybe_conceal_def_binding = Thm.def_binding
+          #> Config.get no_defs_lthy bnf_note_all = false ? Binding.conceal;
+
+        val ((raw_ctrs, raw_ctr_defs), (lthy', lthy)) = no_defs_lthy
+          |> apfst split_list o fold_map3 (fn b => fn mx => fn rhs =>
+              Local_Theory.define ((b, mx), ((maybe_conceal_def_binding b, []), rhs)) #>> apsnd snd)
+            ctr_bindings ctr_mixfixes ctr_rhss
+          ||> `Local_Theory.restore;
+
+        val phi = Proof_Context.export_morphism lthy lthy';
+
+        val ctr_defs = map (Morphism.thm phi) raw_ctr_defs;
+        val ctr_defs' =
+          map2 (fn m => fn def => mk_unabs_def m (def RS meta_eq_to_obj_eq)) ms ctr_defs;
+
+        val ctrs0 = map (Morphism.term phi) raw_ctrs;
+        val ctrs = map (mk_ctr As) ctrs0;
+
+        fun wrap_ctrs lthy =
+          let
+            fun exhaust_tac {context = ctxt, prems = _} =
+              let
+                val ctor_iff_dtor_thm =
+                  let
+                    val goal =
+                      fold_rev Logic.all [w, u]
+                        (mk_Trueprop_eq (HOLogic.mk_eq (u, ctor $ w), HOLogic.mk_eq (dtor $ u, w)));
+                  in
+                    Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
+                      mk_ctor_iff_dtor_tac ctxt (map (SOME o certifyT lthy) [dtorT, fpT])
+                        (certify lthy ctor) (certify lthy dtor) ctor_dtor dtor_ctor)
+                    |> Thm.close_derivation
+                    |> Morphism.thm phi
+                  end;
+
+                val sumEN_thm' =
+                  unfold_thms lthy @{thms unit_all_eq1}
+                    (Drule.instantiate' (map (SOME o certifyT lthy) ctr_prod_Ts) []
+                       (mk_sumEN_balanced n))
+                  |> Morphism.thm phi;
+              in
+                mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor_thm sumEN_thm'
+              end;
+
+            val inject_tacss =
+              map2 (fn 0 => K [] | _ => fn ctr_def => [fn {context = ctxt, ...} =>
+                mk_inject_tac ctxt ctr_def ctor_inject]) ms ctr_defs;
+
+            val half_distinct_tacss =
+              map (map (fn (def, def') => fn {context = ctxt, ...} =>
+                mk_half_distinct_tac ctxt ctor_inject [def, def'])) (mk_half_pairss (`I ctr_defs));
+
+            val tacss = [exhaust_tac] :: inject_tacss @ half_distinct_tacss;
+
+            val sel_defaultss = map (map (apsnd (prepare_term lthy))) raw_sel_defaultss
+          in
+            wrap_free_constructors tacss (((wrap_opts, ctrs0), standard_binding), (disc_bindings,
+              (sel_bindingss, sel_defaultss))) lthy
+          end;
+
+        fun derive_maps_sets_rels (ctr_sugar, lthy) =
+          if live = 0 then
+            ((([], [], [], []), ctr_sugar), lthy)
+          else
+            let
+              val rel_flip = rel_flip_of_bnf fp_bnf;
+              val nones = replicate live NONE;
+
+              val ctor_cong =
+                if fp = Least_FP then
+                  Drule.dummy_thm
+                else
+                  let val ctor' = mk_ctor Bs ctor in
+                    cterm_instantiate_pos [NONE, NONE, SOME (certify lthy ctor')] arg_cong
+                  end;
+
+              fun mk_cIn ify =
+                certify lthy o (fp = Greatest_FP ? curry (op $) (map_types ify ctor)) oo
+                mk_InN_balanced (ify ctr_sum_prod_T) n;
+
+              val cxIns = map2 (mk_cIn I) tuple_xs ks;
+              val cyIns = map2 (mk_cIn B_ify) tuple_ys ks;
+
+              fun mk_map_thm ctr_def' cxIn =
+                fold_thms lthy [ctr_def']
+                  (unfold_thms lthy (pre_map_def ::
+                       (if fp = Least_FP then [] else [ctor_dtor, dtor_ctor]) @ sum_prod_thms_map)
+                     (cterm_instantiate_pos (nones @ [SOME cxIn])
+                        (if fp = Least_FP then fp_map_thm else fp_map_thm RS ctor_cong)))
+                |> singleton (Proof_Context.export names_lthy no_defs_lthy);
+
+              fun mk_set_thm fp_set_thm ctr_def' cxIn =
+                fold_thms lthy [ctr_def']
+                  (unfold_thms lthy (pre_set_defs @ nested_set_maps @ nesting_set_maps @
+                       (if fp = Least_FP then [] else [dtor_ctor]) @ sum_prod_thms_set)
+                     (cterm_instantiate_pos [SOME cxIn] fp_set_thm))
+                |> singleton (Proof_Context.export names_lthy no_defs_lthy);
+
+              fun mk_set_thms fp_set_thm = map2 (mk_set_thm fp_set_thm) ctr_defs' cxIns;
+
+              val map_thms = map2 mk_map_thm ctr_defs' cxIns;
+              val set_thmss = map mk_set_thms fp_set_thms;
+
+              val rel_infos = (ctr_defs' ~~ cxIns, ctr_defs' ~~ cyIns);
+
+              fun mk_rel_thm postproc ctr_defs' cxIn cyIn =
+                fold_thms lthy ctr_defs'
+                  (unfold_thms lthy (@{thm Inl_Inr_False} :: pre_rel_def ::
+                       (if fp = Least_FP then [] else [dtor_ctor]) @ sum_prod_thms_rel)
+                     (cterm_instantiate_pos (nones @ [SOME cxIn, SOME cyIn]) fp_rel_thm))
+                |> postproc
+                |> singleton (Proof_Context.export names_lthy no_defs_lthy);
+
+              fun mk_rel_inject_thm ((ctr_def', cxIn), (_, cyIn)) =
+                mk_rel_thm (unfold_thms lthy @{thms eq_sym_Unity_conv}) [ctr_def'] cxIn cyIn;
+
+              val rel_inject_thms = map mk_rel_inject_thm (op ~~ rel_infos);
+
+              fun mk_half_rel_distinct_thm ((xctr_def', cxIn), (yctr_def', cyIn)) =
+                mk_rel_thm (fn thm => thm RS @{thm eq_False[THEN iffD1]}) [xctr_def', yctr_def']
+                  cxIn cyIn;
+
+              fun mk_other_half_rel_distinct_thm thm =
+                flip_rels lthy live thm
+                RS (rel_flip RS sym RS @{thm arg_cong[of _ _ Not]} RS iffD2);
+
+              val half_rel_distinct_thmss =
+                map (map mk_half_rel_distinct_thm) (mk_half_pairss rel_infos);
+              val other_half_rel_distinct_thmss =
+                map (map mk_other_half_rel_distinct_thm) half_rel_distinct_thmss;
+              val (rel_distinct_thms, _) =
+                join_halves n half_rel_distinct_thmss other_half_rel_distinct_thmss;
+
+              val anonymous_notes =
+                [(map (fn th => th RS @{thm eq_False[THEN iffD2]}) rel_distinct_thms,
+                  code_nitpicksimp_attrs),
+                 (map2 (fn th => fn 0 => th RS @{thm eq_True[THEN iffD2]} | _ => th)
+                    rel_inject_thms ms, code_nitpicksimp_attrs)]
+                |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
+
+              val notes =
+                [(mapN, map_thms, code_nitpicksimp_attrs @ simp_attrs),
+                 (rel_distinctN, rel_distinct_thms, simp_attrs),
+                 (rel_injectN, rel_inject_thms, simp_attrs),
+                 (setN, flat set_thmss, code_nitpicksimp_attrs @ simp_attrs)]
+                |> massage_simple_notes fp_b_name;
+            in
+              (((map_thms, rel_inject_thms, rel_distinct_thms, set_thmss), ctr_sugar),
+               lthy |> Local_Theory.notes (anonymous_notes @ notes) |> snd)
+            end;
+
+        fun mk_binding pre = qualify false fp_b_name (Binding.prefix_name (pre ^ "_") fp_b);
+
+        fun massage_res (((maps_sets_rels, ctr_sugar), co_iter_res), lthy) =
+          (((maps_sets_rels, (ctrs, xss, ctr_defs, ctr_sugar)), co_iter_res), lthy);
+      in
+        (wrap_ctrs
+         #> derive_maps_sets_rels
+         ##>>
+           (if fp = Least_FP then define_iters [foldN, recN] (the iters_args_types)
+            else define_coiters [unfoldN, corecN] (the coiters_args_types))
+             mk_binding fpTs Cs xtor_co_iters
+         #> massage_res, lthy')
+      end;
+
+    fun wrap_types_etc (wrap_types_etcs, lthy) =
+      fold_map I wrap_types_etcs lthy
+      |>> apsnd split_list o apfst (apsnd split_list4 o apfst split_list4 o split_list)
+        o split_list;
+
+    fun mk_simp_thms ({injects, distincts, case_thms, ...} : ctr_sugar) un_folds co_recs
+        mapsx rel_injects rel_distincts setss =
+      injects @ distincts @ case_thms @ co_recs @ un_folds @ mapsx @ rel_injects @ rel_distincts
+      @ flat setss;
+
+    fun derive_note_induct_iters_thms_for_types
+        ((((mapss, rel_injects, rel_distincts, setss), (ctrss, _, ctr_defss, ctr_sugars)),
+          (iterss, iter_defss)), lthy) =
+      let
+        val ((induct_thms, induct_thm, induct_attrs), (fold_thmss, rec_thmss, iter_attrs)) =
+          derive_induct_iters_thms_for_types pre_bnfs (the iters_args_types) xtor_co_induct
+            xtor_co_iter_thmss nesting_bnfs nested_bnfs fpTs Cs Xs ctrXs_Tsss ctrss ctr_defss iterss
+            iter_defss lthy;
+
+        val induct_type_attr = Attrib.internal o K o Induct.induct_type;
+
+        val simp_thmss =
+          map7 mk_simp_thms ctr_sugars fold_thmss rec_thmss mapss rel_injects rel_distincts setss;
+
+        val common_notes =
+          (if nn > 1 then [(inductN, [induct_thm], induct_attrs)] else [])
+          |> massage_simple_notes fp_common_name;
+
+        val notes =
+          [(foldN, fold_thmss, K iter_attrs),
+           (inductN, map single induct_thms, fn T_name => induct_attrs @ [induct_type_attr T_name]),
+           (recN, rec_thmss, K iter_attrs),
+           (simpsN, simp_thmss, K [])]
+          |> massage_multi_notes;
+      in
+        lthy
+        |> Local_Theory.notes (common_notes @ notes) |> snd
+        |> register_fp_sugars Least_FP pre_bnfs nested_bnfs nesting_bnfs fp_res ctr_defss ctr_sugars
+          iterss mapss [induct_thm] (transpose [fold_thmss, rec_thmss]) [] []
+      end;
+
+    fun derive_note_coinduct_coiters_thms_for_types
+        ((((mapss, rel_injects, rel_distincts, setss), (_, _, ctr_defss, ctr_sugars)),
+          (coiterss, coiter_defss)), lthy) =
+      let
+        val (([(coinduct_thms, coinduct_thm), (strong_coinduct_thms, strong_coinduct_thm)],
+              coinduct_attrs),
+             (unfold_thmss, corec_thmss, coiter_attrs),
+             (disc_unfold_thmss, disc_corec_thmss, disc_coiter_attrs),
+             (disc_unfold_iff_thmss, disc_corec_iff_thmss, disc_coiter_iff_attrs),
+             (sel_unfold_thmsss, sel_corec_thmsss, sel_coiter_attrs)) =
+          derive_coinduct_coiters_thms_for_types pre_bnfs (the coiters_args_types) xtor_co_induct
+            dtor_injects dtor_ctors xtor_co_iter_thmss nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss ns
+            ctr_defss ctr_sugars coiterss coiter_defss (Proof_Context.export lthy' no_defs_lthy)
+            lthy;
+
+        val sel_unfold_thmss = map flat sel_unfold_thmsss;
+        val sel_corec_thmss = map flat sel_corec_thmsss;
+
+        val coinduct_type_attr = Attrib.internal o K o Induct.coinduct_type;
+
+        val flat_coiter_thms = append oo append;
+
+        val simp_thmss =
+          map7 mk_simp_thms ctr_sugars
+            (map3 flat_coiter_thms disc_unfold_thmss disc_unfold_iff_thmss sel_unfold_thmss)
+            (map3 flat_coiter_thms disc_corec_thmss disc_corec_iff_thmss sel_corec_thmss)
+            mapss rel_injects rel_distincts setss;
+
+        val common_notes =
+          (if nn > 1 then
+             [(coinductN, [coinduct_thm], coinduct_attrs),
+              (strong_coinductN, [strong_coinduct_thm], coinduct_attrs)]
+           else
+             [])
+          |> massage_simple_notes fp_common_name;
+
+        val notes =
+          [(coinductN, map single coinduct_thms,
+            fn T_name => coinduct_attrs @ [coinduct_type_attr T_name]),
+           (corecN, corec_thmss, K coiter_attrs),
+           (disc_corecN, disc_corec_thmss, K disc_coiter_attrs),
+           (disc_corec_iffN, disc_corec_iff_thmss, K disc_coiter_iff_attrs),
+           (disc_unfoldN, disc_unfold_thmss, K disc_coiter_attrs),
+           (disc_unfold_iffN, disc_unfold_iff_thmss, K disc_coiter_iff_attrs),
+           (sel_corecN, sel_corec_thmss, K sel_coiter_attrs),
+           (sel_unfoldN, sel_unfold_thmss, K sel_coiter_attrs),
+           (simpsN, simp_thmss, K []),
+           (strong_coinductN, map single strong_coinduct_thms, K coinduct_attrs),
+           (unfoldN, unfold_thmss, K coiter_attrs)]
+          |> massage_multi_notes;
+
+        fun is_codatatype (Type (s, _)) =
+            (case fp_sugar_of lthy s of SOME {fp = Greatest_FP, ...} => true | _ => false)
+          | is_codatatype _ = false;
+
+        val nitpick_supported = forall (is_codatatype o T_of_bnf) nested_bnfs;
+
+        fun register_nitpick fpT ({ctrs, casex, ...} : ctr_sugar) =
+          Nitpick_HOL.register_codatatype fpT (fst (dest_Const casex))
+            (map (dest_Const o mk_ctr As) ctrs)
+          |> Generic_Target.theory_declaration;
+      in
+        lthy
+        |> Local_Theory.notes (common_notes @ notes) |> snd
+        |> register_fp_sugars Greatest_FP pre_bnfs nested_bnfs nesting_bnfs fp_res ctr_defss
+          ctr_sugars coiterss mapss [coinduct_thm, strong_coinduct_thm]
+          (transpose [unfold_thmss, corec_thmss]) (transpose [disc_unfold_thmss, disc_corec_thmss])
+          (transpose [sel_unfold_thmsss, sel_corec_thmsss])
+        |> nitpick_supported ? fold2 register_nitpick fpTs ctr_sugars
+      end;
+
+    val lthy'' = lthy'
+      |> fold_map define_ctrs_dtrs_for_type (fp_bnfs ~~ fp_bs ~~ fpTs ~~ ctors ~~ dtors ~~
+        xtor_co_iterss ~~ ctor_dtors ~~ dtor_ctors ~~ ctor_injects ~~ pre_map_defs ~~
+        pre_set_defss ~~ pre_rel_defs ~~ xtor_map_thms ~~ xtor_set_thmss ~~ xtor_rel_thms ~~ ns ~~
+        kss ~~ mss ~~ ctr_bindingss ~~ ctr_mixfixess ~~ ctr_Tsss ~~ disc_bindingss ~~
+        sel_bindingsss ~~ raw_sel_defaultsss)
+      |> wrap_types_etc
+      |> fp_case fp derive_note_induct_iters_thms_for_types
+           derive_note_coinduct_coiters_thms_for_types;
+
+    val timer = time (timer ("Constructors, discriminators, selectors, etc., for the new " ^
+      co_prefix fp ^ "datatype"));
+  in
+    timer; lthy''
+  end;
+
+fun co_datatypes x = define_co_datatypes (K I) (K I) (K I) x;
+
+fun co_datatype_cmd x =
+  define_co_datatypes Typedecl.read_constraint Syntax.parse_typ Syntax.parse_term x;
+
+val parse_ctr_arg =
+  @{keyword "("} |-- parse_binding_colon -- Parse.typ --| @{keyword ")"} ||
+  (Parse.typ >> pair Binding.empty);
+
+val parse_defaults =
+  @{keyword "("} |-- Parse.reserved "defaults" |-- Scan.repeat parse_bound_term --| @{keyword ")"};
+
+val parse_type_arg_constrained =
+  Parse.type_ident -- Scan.option (@{keyword "::"} |-- Parse.!!! Parse.sort);
+
+val parse_type_arg_named_constrained = parse_opt_binding_colon -- parse_type_arg_constrained;
+
+(*FIXME: use parse_type_args_named_constrained from BNF_Util and thus 
+  allow users to kill certain arguments of a (co)datatype*)
+val parse_type_args_named_constrained =
+  parse_type_arg_constrained >> (single o pair Binding.empty) ||
+  @{keyword "("} |-- Parse.!!! (Parse.list1 parse_type_arg_named_constrained --| @{keyword ")"}) ||
+  Scan.succeed [];
+
+val parse_ctr_spec =
+  parse_opt_binding_colon -- parse_binding -- Scan.repeat parse_ctr_arg --
+  Scan.optional parse_defaults [] -- Parse.opt_mixfix;
+
+val parse_spec =
+  parse_type_args_named_constrained -- parse_binding -- parse_map_rel_bindings --
+  Parse.opt_mixfix -- (@{keyword "="} |-- Parse.enum1 "|" parse_ctr_spec);
+
+val parse_co_datatype = parse_wrap_free_constructors_options -- Parse.and_list1 parse_spec;
+
+fun parse_co_datatype_cmd fp construct_fp = parse_co_datatype >> co_datatype_cmd fp construct_fp;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,181 @@
+(*  Title:      HOL/BNF/Tools/bnf_fp_def_sugar_tactics.ML
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2012
+
+Tactics for datatype and codatatype sugar.
+*)
+
+signature BNF_FP_DEF_SUGAR_TACTICS =
+sig
+  val sum_prod_thms_map: thm list
+  val sum_prod_thms_set: thm list
+  val sum_prod_thms_rel: thm list
+
+  val mk_coinduct_tac: Proof.context -> thm list -> int -> int list -> thm -> thm list ->
+    thm list -> thm list -> thm list list -> thm list list list -> thm list list list -> tactic
+  val mk_coiter_tac: thm list -> thm list -> thm -> thm -> thm -> Proof.context -> tactic
+  val mk_ctor_iff_dtor_tac: Proof.context -> ctyp option list -> cterm -> cterm -> thm -> thm ->
+    tactic
+  val mk_disc_coiter_iff_tac: thm list -> thm list -> thm list -> Proof.context -> tactic
+  val mk_exhaust_tac: Proof.context -> int -> thm list -> thm -> thm -> tactic
+  val mk_half_distinct_tac: Proof.context -> thm -> thm list -> tactic
+  val mk_induct_tac: Proof.context -> int -> int list -> int list list -> int list list list ->
+    thm list -> thm -> thm list -> thm list list -> tactic
+  val mk_inject_tac: Proof.context -> thm -> thm -> tactic
+  val mk_iter_tac: thm list -> thm list -> thm list -> thm -> thm -> Proof.context -> tactic
+end;
+
+structure BNF_FP_Def_Sugar_Tactics : BNF_FP_DEF_SUGAR_TACTICS =
+struct
+
+open BNF_Tactics
+open BNF_Util
+open BNF_FP_Util
+
+val basic_simp_thms = @{thms simp_thms(7,8,12,14,22,24)};
+val more_simp_thms = basic_simp_thms @ @{thms simp_thms(11,15,16,21)};
+
+val sum_prod_thms_map = @{thms id_apply map_pair_simp prod.cases sum.cases sum_map.simps};
+val sum_prod_thms_set0 =
+  @{thms SUP_empty Sup_empty Sup_insert UN_insert Un_empty_left Un_empty_right Un_iff
+      Union_Un_distrib collect_def[abs_def] image_def o_apply map_pair_simp
+      mem_Collect_eq mem_UN_compreh_eq prod_set_simps sum_map.simps sum_set_simps};
+val sum_prod_thms_set = @{thms UN_compreh_eq_eq} @ sum_prod_thms_set0;
+val sum_prod_thms_rel = @{thms prod_rel_simp sum_rel_simps id_apply};
+
+fun hhf_concl_conv cv ctxt ct =
+  (case Thm.term_of ct of
+    Const (@{const_name all}, _) $ Abs _ =>
+    Conv.arg_conv (Conv.abs_conv (hhf_concl_conv cv o snd) ctxt) ct
+  | _ => Conv.concl_conv ~1 cv ct);
+
+fun co_induct_inst_as_projs ctxt k thm =
+  let
+    val fs = Term.add_vars (prop_of thm) []
+      |> filter (fn (_, Type (@{type_name fun}, [_, T'])) => T' <> HOLogic.boolT | _ => false);
+    fun mk_cfp (f as (_, T)) =
+      (certify ctxt (Var f), certify ctxt (mk_proj T (num_binder_types T) k));
+    val cfps = map mk_cfp fs;
+  in
+    Drule.cterm_instantiate cfps thm
+  end;
+
+val co_induct_inst_as_projs_tac = PRIMITIVE oo co_induct_inst_as_projs;
+
+fun mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor sumEN' =
+  unfold_thms_tac ctxt (ctor_iff_dtor :: ctr_defs) THEN HEADGOAL (rtac sumEN') THEN
+  unfold_thms_tac ctxt @{thms split_paired_all} THEN
+  HEADGOAL (EVERY' (maps (fn k => [select_prem_tac n (rotate_tac 1) k,
+    REPEAT_DETERM o dtac meta_spec, etac meta_mp, atac]) (1 upto n)));
+
+fun mk_ctor_iff_dtor_tac ctxt cTs cctor cdtor ctor_dtor dtor_ctor =
+  HEADGOAL (rtac iffI THEN'
+    EVERY' (map3 (fn cTs => fn cx => fn th =>
+      dtac (Drule.instantiate' cTs [NONE, NONE, SOME cx] arg_cong) THEN'
+      SELECT_GOAL (unfold_thms_tac ctxt [th]) THEN'
+      atac) [rev cTs, cTs] [cdtor, cctor] [dtor_ctor, ctor_dtor]));
+
+fun mk_half_distinct_tac ctxt ctor_inject ctr_defs =
+  unfold_thms_tac ctxt (ctor_inject :: @{thms sum.inject} @ ctr_defs) THEN
+  HEADGOAL (rtac @{thm sum.distinct(1)});
+
+fun mk_inject_tac ctxt ctr_def ctor_inject =
+  unfold_thms_tac ctxt [ctr_def] THEN HEADGOAL (rtac (ctor_inject RS ssubst)) THEN
+  unfold_thms_tac ctxt @{thms sum.inject Pair_eq conj_assoc} THEN HEADGOAL (rtac refl);
+
+val iter_unfold_thms =
+  @{thms comp_def convol_def fst_conv id_def prod_case_Pair_iden snd_conv
+      split_conv unit_case_Unity} @ sum_prod_thms_map;
+
+fun mk_iter_tac pre_map_defs map_idents iter_defs ctor_iter ctr_def ctxt =
+  unfold_thms_tac ctxt (ctr_def :: ctor_iter :: iter_defs @ pre_map_defs @ map_idents @
+    iter_unfold_thms) THEN HEADGOAL (rtac refl);
+
+val coiter_unfold_thms = @{thms id_def} @ sum_prod_thms_map;
+val ss_if_True_False = simpset_of (ss_only @{thms if_True if_False} @{context});
+
+fun mk_coiter_tac coiter_defs map_idents ctor_dtor_coiter pre_map_def ctr_def ctxt =
+  unfold_thms_tac ctxt (ctr_def :: coiter_defs) THEN
+  HEADGOAL (rtac (ctor_dtor_coiter RS trans) THEN'
+    asm_simp_tac (put_simpset ss_if_True_False ctxt)) THEN_MAYBE
+  (unfold_thms_tac ctxt (pre_map_def :: map_idents @ coiter_unfold_thms) THEN
+   HEADGOAL (rtac refl ORELSE' rtac (@{thm unit_eq} RS arg_cong)));
+
+fun mk_disc_coiter_iff_tac case_splits' coiters discs ctxt =
+  EVERY (map3 (fn case_split_tac => fn coiter_thm => fn disc =>
+      HEADGOAL case_split_tac THEN unfold_thms_tac ctxt [coiter_thm] THEN
+      HEADGOAL (asm_simp_tac (ss_only basic_simp_thms ctxt)) THEN
+      (if is_refl disc then all_tac else HEADGOAL (rtac disc)))
+    (map rtac case_splits' @ [K all_tac]) coiters discs);
+
+fun solve_prem_prem_tac ctxt =
+  REPEAT o (eresolve_tac @{thms bexE rev_bexI} ORELSE' rtac @{thm rev_bexI[OF UNIV_I]} ORELSE'
+    hyp_subst_tac ctxt ORELSE' resolve_tac @{thms disjI1 disjI2}) THEN'
+  (rtac refl ORELSE' atac ORELSE' rtac @{thm singletonI});
+
+fun mk_induct_leverage_prem_prems_tac ctxt nn kks set_maps pre_set_defs =
+  HEADGOAL (EVERY' (maps (fn kk => [select_prem_tac nn (dtac meta_spec) kk, etac meta_mp,
+    SELECT_GOAL (unfold_thms_tac ctxt (pre_set_defs @ set_maps @ sum_prod_thms_set0)),
+    solve_prem_prem_tac ctxt]) (rev kks)));
+
+fun mk_induct_discharge_prem_tac ctxt nn n set_maps pre_set_defs m k kks =
+  let val r = length kks in
+    HEADGOAL (EVERY' [select_prem_tac n (rotate_tac 1) k, rotate_tac ~1, hyp_subst_tac ctxt,
+      REPEAT_DETERM_N m o (dtac meta_spec THEN' rotate_tac ~1)]) THEN
+    EVERY [REPEAT_DETERM_N r
+        (HEADGOAL (rotate_tac ~1 THEN' dtac meta_mp THEN' rotate_tac 1) THEN prefer_tac 2),
+      if r > 0 then ALLGOALS (Goal.norm_hhf_tac ctxt) else all_tac, HEADGOAL atac,
+      mk_induct_leverage_prem_prems_tac ctxt nn kks set_maps pre_set_defs]
+  end;
+
+fun mk_induct_tac ctxt nn ns mss kkss ctr_defs ctor_induct' set_maps pre_set_defss =
+  let val n = Integer.sum ns in
+    unfold_thms_tac ctxt ctr_defs THEN HEADGOAL (rtac ctor_induct') THEN
+    co_induct_inst_as_projs_tac ctxt 0 THEN
+    EVERY (map4 (EVERY oooo map3 o mk_induct_discharge_prem_tac ctxt nn n set_maps) pre_set_defss
+      mss (unflat mss (1 upto n)) kkss)
+  end;
+
+fun mk_coinduct_same_ctr_tac ctxt rel_eqs pre_rel_def dtor_ctor ctr_def discs sels =
+  hyp_subst_tac ctxt THEN'
+  CONVERSION (hhf_concl_conv
+    (Conv.top_conv (K (Conv.try_conv (Conv.rewr_conv ctr_def))) ctxt) ctxt) THEN'
+  SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: dtor_ctor :: sels)) THEN'
+  SELECT_GOAL (unfold_thms_tac ctxt (pre_rel_def :: dtor_ctor :: sels @ sum_prod_thms_rel)) THEN'
+  (atac ORELSE' REPEAT o etac conjE THEN'
+     full_simp_tac
+       (ss_only (@{thm prod.inject} :: no_refl discs @ rel_eqs @ more_simp_thms) ctxt) THEN'
+     REPEAT o etac conjE THEN_MAYBE' REPEAT o hyp_subst_tac ctxt THEN'
+     REPEAT o (resolve_tac [refl, conjI] ORELSE' atac));
+
+fun mk_coinduct_distinct_ctrs_tac ctxt discs discs' =
+  let
+    val discs'' = map (perhaps (try (fn th => th RS @{thm notnotD}))) (discs @ discs')
+      |> distinct Thm.eq_thm_prop;
+  in
+    hyp_subst_tac ctxt THEN' REPEAT o etac conjE THEN'
+    full_simp_tac (ss_only (refl :: no_refl discs'' @ basic_simp_thms) ctxt)
+  end;
+
+fun mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn kk n pre_rel_def dtor_ctor exhaust ctr_defs
+    discss selss =
+  let val ks = 1 upto n in
+    EVERY' ([rtac allI, rtac allI, rtac impI, select_prem_tac nn (dtac meta_spec) kk,
+        dtac meta_spec, dtac meta_mp, atac, rtac exhaust, K (co_induct_inst_as_projs_tac ctxt 0),
+        hyp_subst_tac ctxt] @
+      map4 (fn k => fn ctr_def => fn discs => fn sels =>
+        EVERY' ([rtac exhaust, K (co_induct_inst_as_projs_tac ctxt 1)] @
+          map2 (fn k' => fn discs' =>
+            if k' = k then
+              mk_coinduct_same_ctr_tac ctxt rel_eqs' pre_rel_def dtor_ctor ctr_def discs sels
+            else
+              mk_coinduct_distinct_ctrs_tac ctxt discs discs') ks discss)) ks ctr_defs discss selss)
+  end;
+
+fun mk_coinduct_tac ctxt rel_eqs' nn ns dtor_coinduct' pre_rel_defs dtor_ctors exhausts ctr_defss
+    discsss selsss =
+  HEADGOAL (rtac dtor_coinduct' THEN'
+    EVERY' (map8 (mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn)
+      (1 upto nn) ns pre_rel_defs dtor_ctors exhausts ctr_defss discsss selsss));
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_fp_n2m.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,378 @@
+(*  Title:      HOL/BNF/Tools/bnf_fp_n2m.ML
+    Author:     Dmitriy Traytel, TU Muenchen
+    Copyright   2013
+
+Flattening of nested to mutual (co)recursion.
+*)
+
+signature BNF_FP_N2M =
+sig
+  val construct_mutualized_fp: BNF_FP_Util.fp_kind  -> typ list -> BNF_FP_Def_Sugar.fp_sugar list ->
+    binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list ->
+    local_theory -> BNF_FP_Util.fp_result * local_theory
+end;
+
+structure BNF_FP_N2M : BNF_FP_N2M =
+struct
+
+open BNF_Def
+open BNF_Util
+open BNF_FP_Util
+open BNF_FP_Def_Sugar
+open BNF_Tactics
+open BNF_FP_N2M_Tactics
+
+fun force_typ ctxt T =
+  map_types Type_Infer.paramify_vars
+  #> Type.constraint T
+  #> Syntax.check_term ctxt
+  #> singleton (Variable.polymorphic ctxt);
+
+fun mk_prod_map f g =
+  let
+    val ((fAT, fBT), fT) = `dest_funT (fastype_of f);
+    val ((gAT, gBT), gT) = `dest_funT (fastype_of g);
+  in
+    Const (@{const_name map_pair},
+      fT --> gT --> HOLogic.mk_prodT (fAT, gAT) --> HOLogic.mk_prodT (fBT, gBT)) $ f $ g
+  end;
+
+fun mk_sum_map f g =
+  let
+    val ((fAT, fBT), fT) = `dest_funT (fastype_of f);
+    val ((gAT, gBT), gT) = `dest_funT (fastype_of g);
+  in
+    Const (@{const_name sum_map}, fT --> gT --> mk_sumT (fAT, gAT) --> mk_sumT (fBT, gBT)) $ f $ g
+  end;
+
+fun construct_mutualized_fp fp fpTs fp_sugars bs resBs (resDs, Dss) bnfs lthy =
+  let
+    fun steal get = map (of_fp_sugar (get o #fp_res)) fp_sugars;
+
+    val n = length bnfs;
+    val deads = fold (union (op =)) Dss resDs;
+    val As = subtract (op =) deads (map TFree resBs);
+    val names_lthy = fold Variable.declare_typ (As @ deads) lthy;
+    val m = length As;
+    val live = m + n;
+    val ((Xs, Bs), names_lthy) = names_lthy
+      |> mk_TFrees n
+      ||>> mk_TFrees m;
+    val allAs = As @ Xs;
+    val phiTs = map2 mk_pred2T As Bs;
+    val theta = As ~~ Bs;
+    val fpTs' = map (Term.typ_subst_atomic theta) fpTs;
+    val pre_phiTs = map2 mk_pred2T fpTs fpTs';
+
+    fun mk_co_algT T U = fp_case fp (T --> U) (U --> T);
+    fun co_swap pair = fp_case fp I swap pair;
+    val dest_co_algT = co_swap o dest_funT;
+    val co_alg_argT = fp_case fp range_type domain_type;
+    val co_alg_funT = fp_case fp domain_type range_type;
+    val mk_co_product = curry (fp_case fp mk_convol mk_sum_case);
+    val mk_map_co_product = fp_case fp mk_prod_map mk_sum_map;
+    val co_proj1_const = fp_case fp (fst_const o fst) (uncurry Inl_const o dest_sumT o snd);
+    val mk_co_productT = curry (fp_case fp HOLogic.mk_prodT mk_sumT);
+    val dest_co_productT = fp_case fp HOLogic.dest_prodT dest_sumT;
+
+    val ((ctors, dtors), (xtor's, xtors)) =
+      let
+        val ctors = map2 (force_typ names_lthy o (fn T => dummyT --> T)) fpTs (steal #ctors);
+        val dtors = map2 (force_typ names_lthy o (fn T => T --> dummyT)) fpTs (steal #dtors);
+      in
+        ((ctors, dtors), `(map (Term.subst_atomic_types theta)) (fp_case fp ctors dtors))
+      end;
+
+    val xTs = map (domain_type o fastype_of) xtors;
+    val yTs = map (domain_type o fastype_of) xtor's;
+
+    val (((((phis, phis'), pre_phis), xs), ys), names_lthy) = names_lthy
+      |> mk_Frees' "R" phiTs
+      ||>> mk_Frees "S" pre_phiTs
+      ||>> mk_Frees "x" xTs
+      ||>> mk_Frees "y" yTs;
+
+    val fp_bnfs = steal #bnfs;
+    val pre_bnfs = map (of_fp_sugar #pre_bnfs) fp_sugars;
+    val pre_bnfss = map #pre_bnfs fp_sugars;
+    val nesty_bnfss = map (fn sugar => #nested_bnfs sugar @ #nesting_bnfs sugar) fp_sugars;
+    val fp_nesty_bnfss = fp_bnfs :: nesty_bnfss;
+    val fp_nesty_bnfs = distinct eq_bnf (flat fp_nesty_bnfss);
+
+    val rels =
+      let
+        fun find_rel T As Bs = fp_nesty_bnfss
+          |> map (filter_out (curry eq_bnf BNF_Comp.DEADID_bnf))
+          |> get_first (find_first (fn bnf => Type.could_unify (T_of_bnf bnf, T)))
+          |> Option.map (fn bnf =>
+            let val live = live_of_bnf bnf;
+            in (mk_rel live As Bs (rel_of_bnf bnf), live) end)
+          |> the_default (HOLogic.eq_const T, 0);
+
+        fun mk_rel (T as Type (_, Ts)) (Type (_, Us)) =
+              let
+                val (rel, live) = find_rel T Ts Us;
+                val (Ts', Us') = fastype_of rel |> strip_typeN live |> fst |> map_split dest_pred2T;
+                val rels = map2 mk_rel Ts' Us';
+              in
+                Term.list_comb (rel, rels)
+              end
+          | mk_rel (T as TFree _) _ = (nth phis (find_index (curry op = T) As)
+              handle General.Subscript => HOLogic.eq_const T)
+          | mk_rel _ _ = raise Fail "fpTs contains schematic type variables";
+      in
+        map2 (fold_rev Term.absfree phis' oo mk_rel) fpTs fpTs'
+      end;
+
+    val pre_rels = map2 (fn Ds => mk_rel_of_bnf Ds (As @ fpTs) (Bs @ fpTs')) Dss bnfs;
+
+    val rel_unfoldss = map (maps (fn bnf => no_refl [rel_def_of_bnf bnf])) pre_bnfss;
+    val rel_xtor_co_inducts = steal (split_conj_thm o #rel_xtor_co_induct_thm)
+      |> map2 (fn unfs => unfold_thms lthy (id_apply :: unfs)) rel_unfoldss;
+
+    val rel_defs = map rel_def_of_bnf bnfs;
+    val rel_monos = map rel_mono_of_bnf bnfs;
+
+    val rel_xtor_co_induct_thm =
+      mk_rel_xtor_co_induct_thm fp pre_rels pre_phis rels phis xs ys xtors xtor's
+        (mk_rel_xtor_co_induct_tactic fp rel_xtor_co_inducts rel_defs rel_monos) lthy;
+
+    val rel_eqs = no_refl (map rel_eq_of_bnf fp_nesty_bnfs);
+    val map_id0s = no_refl (map map_id0_of_bnf bnfs);
+
+    val xtor_co_induct_thm =
+      (case fp of
+        Least_FP =>
+          let
+            val (Ps, names_lthy) = names_lthy
+              |> mk_Frees "P" (map (fn T => T --> HOLogic.boolT) fpTs);
+            fun mk_Grp_id P =
+              let val T = domain_type (fastype_of P);
+              in mk_Grp (HOLogic.Collect_const T $ P) (HOLogic.id_const T) end;
+            val cts = map (SOME o certify lthy) (map HOLogic.eq_const As @ map mk_Grp_id Ps);
+          in
+            cterm_instantiate_pos cts rel_xtor_co_induct_thm
+            |> singleton (Proof_Context.export names_lthy lthy)
+            |> unfold_thms lthy (@{thms eq_le_Grp_id_iff all_simps(1,2)[symmetric]} @ rel_eqs)
+            |> funpow n (fn thm => thm RS spec)
+            |> unfold_thms lthy (@{thm eq_alt} :: map rel_Grp_of_bnf bnfs @ map_id0s)
+            |> unfold_thms lthy @{thms Grp_id_mono_subst eqTrueI[OF subset_UNIV] simp_thms(22)}
+            |> unfold_thms lthy @{thms subset_iff mem_Collect_eq
+               atomize_conjL[symmetric] atomize_all[symmetric] atomize_imp[symmetric]}
+            |> unfold_thms lthy (maps set_defs_of_bnf bnfs)
+          end
+      | Greatest_FP =>
+          let
+            val cts = NONE :: map (SOME o certify lthy) (map HOLogic.eq_const As);
+          in
+            cterm_instantiate_pos cts rel_xtor_co_induct_thm
+            |> unfold_thms lthy (@{thms le_fun_def le_bool_def all_simps(1,2)[symmetric]} @ rel_eqs)
+            |> funpow (2 * n) (fn thm => thm RS spec)
+            |> Conv.fconv_rule (Object_Logic.atomize lthy)
+            |> funpow n (fn thm => thm RS mp)
+          end);
+
+    val fold_preTs = map2 (fn Ds => mk_T_of_bnf Ds allAs) Dss bnfs;
+    val fold_pre_deads_only_Ts = map2 (fn Ds => mk_T_of_bnf Ds (replicate live dummyT)) Dss bnfs;
+    val rec_theta = Xs ~~ map2 mk_co_productT fpTs Xs;
+    val rec_preTs = map (Term.typ_subst_atomic rec_theta) fold_preTs;
+
+    val fold_strTs = map2 mk_co_algT fold_preTs Xs;
+    val rec_strTs = map2 mk_co_algT rec_preTs Xs;
+    val resTs = map2 mk_co_algT fpTs Xs;
+
+    val (((fold_strs, fold_strs'), (rec_strs, rec_strs')), names_lthy) = names_lthy
+      |> mk_Frees' "s" fold_strTs
+      ||>> mk_Frees' "s" rec_strTs;
+
+    val co_iters = steal #xtor_co_iterss;
+    val ns = map (length o #pre_bnfs) fp_sugars;
+    fun substT rho (Type (@{type_name "fun"}, [T, U])) = substT rho T --> substT rho U
+      | substT rho (Type (s, Ts)) = Type (s, map (typ_subst_nonatomic rho) Ts)
+      | substT _ T = T;
+    fun force_iter is_rec i TU TU_rec raw_iters =
+      let
+        val approx_fold = un_fold_of raw_iters
+          |> force_typ names_lthy
+            (replicate (nth ns i) dummyT ---> (if is_rec then TU_rec else TU));
+        val TUs = binder_fun_types (Term.typ_subst_atomic (Xs ~~ fpTs) (fastype_of approx_fold));
+        val js = find_indices Type.could_unify
+          TUs (map (Term.typ_subst_atomic (Xs ~~ fpTs)) fold_strTs);
+        val Tpats = map (fn j => mk_co_algT (nth fold_pre_deads_only_Ts j) (nth Xs j)) js;
+        val iter = raw_iters |> (if is_rec then co_rec_of else un_fold_of);
+      in
+        force_typ names_lthy (Tpats ---> TU) iter
+      end;
+
+    fun mk_iter b_opt is_rec iters lthy TU =
+      let
+        val x = co_alg_argT TU;
+        val i = find_index (fn T => x = T) Xs;
+        val TUiter =
+          (case find_first (fn f => body_fun_type (fastype_of f) = TU) iters of
+            NONE => nth co_iters i
+              |> force_iter is_rec i
+                (TU |> (is_none b_opt andalso not is_rec) ? substT (fpTs ~~ Xs))
+                (TU |> (is_none b_opt) ? substT (map2 mk_co_productT fpTs Xs ~~ Xs))
+          | SOME f => f);
+        val TUs = binder_fun_types (fastype_of TUiter);
+        val iter_preTs = if is_rec then rec_preTs else fold_preTs;
+        val iter_strs = if is_rec then rec_strs else fold_strs;
+        fun mk_s TU' =
+          let
+            val i = find_index (fn T => co_alg_argT TU' = T) Xs;
+            val sF = co_alg_funT TU';
+            val F = nth iter_preTs i;
+            val s = nth iter_strs i;
+          in
+            (if sF = F then s
+            else
+              let
+                val smapT = replicate live dummyT ---> mk_co_algT sF F;
+                fun hidden_to_unit t =
+                  Term.subst_TVars (map (rpair HOLogic.unitT) (Term.add_tvar_names t [])) t;
+                val smap = map_of_bnf (nth bnfs i)
+                  |> force_typ names_lthy smapT
+                  |> hidden_to_unit;
+                val smap_argTs = strip_typeN live (fastype_of smap) |> fst;
+                fun mk_smap_arg TU =
+                  (if domain_type TU = range_type TU then
+                    HOLogic.id_const (domain_type TU)
+                  else if is_rec then
+                    let
+                      val (TY, (U, X)) = TU |> dest_co_algT ||> dest_co_productT;
+                      val T = mk_co_algT TY U;
+                    in
+                      (case try (force_typ lthy T o build_map lthy co_proj1_const o dest_funT) T of
+                        SOME f => mk_co_product f
+                          (fst (fst (mk_iter NONE is_rec iters lthy (mk_co_algT TY X))))
+                      | NONE => mk_map_co_product
+                          (build_map lthy co_proj1_const
+                            (dest_funT (mk_co_algT (dest_co_productT TY |> fst) U)))
+                          (HOLogic.id_const X))
+                    end
+                  else
+                    fst (fst (mk_iter NONE is_rec iters lthy TU)))
+                val smap_args = map mk_smap_arg smap_argTs;
+              in
+                HOLogic.mk_comp (co_swap (s, Term.list_comb (smap, smap_args)))
+              end)
+          end;
+        val t = Term.list_comb (TUiter, map mk_s TUs);
+      in
+        (case b_opt of
+          NONE => ((t, Drule.dummy_thm), lthy)
+        | SOME b => Local_Theory.define ((b, NoSyn), ((Binding.conceal (Thm.def_binding b), []),
+            fold_rev Term.absfree (if is_rec then rec_strs' else fold_strs') t)) lthy |>> apsnd snd)
+      end;
+
+    fun mk_iters is_rec name lthy =
+      fold2 (fn TU => fn b => fn ((iters, defs), lthy) =>
+        mk_iter (SOME b) is_rec iters lthy TU |>> (fn (f, d) => (f :: iters, d :: defs)))
+      resTs (map (Binding.suffix_name ("_" ^ name)) bs) (([], []), lthy)
+      |>> apfst rev o apsnd rev;
+    val foldN = fp_case fp ctor_foldN dtor_unfoldN;
+    val recN = fp_case fp ctor_recN dtor_corecN;
+    val (((raw_un_folds, raw_un_fold_defs), (raw_co_recs, raw_co_rec_defs)), (lthy, raw_lthy)) =
+      lthy
+      |> mk_iters false foldN
+      ||>> mk_iters true recN
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism raw_lthy lthy;
+
+    val un_folds = map (Morphism.term phi) raw_un_folds;
+    val co_recs = map (Morphism.term phi) raw_co_recs;
+
+    val (xtor_un_fold_thms, xtor_co_rec_thms) =
+      let
+        val folds = map (fn f => Term.list_comb (f, fold_strs)) raw_un_folds;
+        val recs = map (fn r => Term.list_comb (r, rec_strs)) raw_co_recs;
+        val fold_mapTs = co_swap (As @ fpTs, As @ Xs);
+        val rec_mapTs = co_swap (As @ fpTs, As @ map2 mk_co_productT fpTs Xs);
+        val pre_fold_maps =
+          map2 (fn Ds => fn bnf =>
+            Term.list_comb (uncurry (mk_map_of_bnf Ds) fold_mapTs bnf,
+              map HOLogic.id_const As @ folds))
+          Dss bnfs;
+        val pre_rec_maps =
+          map2 (fn Ds => fn bnf =>
+            Term.list_comb (uncurry (mk_map_of_bnf Ds) rec_mapTs bnf,
+              map HOLogic.id_const As @ map2 (mk_co_product o HOLogic.id_const) fpTs recs))
+          Dss bnfs;
+
+        fun mk_goals f xtor s smap =
+          ((f, xtor), (s, smap))
+          |> pairself (HOLogic.mk_comp o co_swap)
+          |> HOLogic.mk_eq;
+
+        val fold_goals = map4 mk_goals folds xtors fold_strs pre_fold_maps
+        val rec_goals = map4 mk_goals recs xtors rec_strs pre_rec_maps;
+
+        fun mk_thms ss goals tac =
+          Library.foldr1 HOLogic.mk_conj goals
+          |> HOLogic.mk_Trueprop
+          |> fold_rev Logic.all ss
+          |> (fn goal => Goal.prove_sorry raw_lthy [] [] goal tac)
+          |> Thm.close_derivation
+          |> Morphism.thm phi
+          |> split_conj_thm
+          |> map (fn thm => thm RS @{thm comp_eq_dest});
+
+        val pre_map_defs = no_refl (map map_def_of_bnf bnfs);
+        val fp_pre_map_defs = no_refl (map map_def_of_bnf pre_bnfs);
+
+        val map_unfoldss = map (maps (fn bnf => no_refl [map_def_of_bnf bnf])) pre_bnfss;
+        val unfold_map = map2 (fn unfs => unfold_thms lthy (id_apply :: unfs)) map_unfoldss;
+
+        val fp_xtor_co_iterss = steal #xtor_co_iter_thmss;
+        val fp_xtor_un_folds = map (mk_pointfree lthy o un_fold_of) fp_xtor_co_iterss |> unfold_map;
+        val fp_xtor_co_recs = map (mk_pointfree lthy o co_rec_of) fp_xtor_co_iterss |> unfold_map;
+
+        val fp_co_iter_o_mapss = steal #xtor_co_iter_o_map_thmss;
+        val fp_fold_o_maps = map un_fold_of fp_co_iter_o_mapss |> unfold_map;
+        val fp_rec_o_maps = map co_rec_of fp_co_iter_o_mapss |> unfold_map;
+        val fold_thms = fp_case fp @{thm o_assoc[symmetric]} @{thm o_assoc} ::
+          @{thms id_apply o_apply o_id id_o map_pair.comp map_pair.id sum_map.comp sum_map.id};
+        val rec_thms = fold_thms @ fp_case fp
+          @{thms fst_convol map_pair_o_convol convol_o}
+          @{thms sum_case_o_inj(1) sum_case_o_sum_map o_sum_case};
+        val map_thms = no_refl (maps (fn bnf =>
+          [map_comp0_of_bnf bnf RS sym, map_id0_of_bnf bnf]) fp_nesty_bnfs);
+
+        fun mk_tac defs o_map_thms xtor_thms thms {context = ctxt, prems = _} =
+          unfold_thms_tac ctxt
+            (flat [thms, defs, pre_map_defs, fp_pre_map_defs, xtor_thms, o_map_thms, map_thms]) THEN
+          CONJ_WRAP (K (HEADGOAL (rtac refl))) bnfs;
+
+        val fold_tac = mk_tac raw_un_fold_defs fp_fold_o_maps fp_xtor_un_folds fold_thms;
+        val rec_tac = mk_tac raw_co_rec_defs fp_rec_o_maps fp_xtor_co_recs rec_thms;
+      in
+        (mk_thms fold_strs fold_goals fold_tac, mk_thms rec_strs rec_goals rec_tac)
+      end;
+
+    (* These results are half broken. This is deliberate. We care only about those fields that are
+       used by "primrec_new", "primcorecursive", and "datatype_new_compat". *)
+    val fp_res =
+      ({Ts = fpTs,
+        bnfs = steal #bnfs,
+        dtors = dtors,
+        ctors = ctors,
+        xtor_co_iterss = transpose [un_folds, co_recs],
+        xtor_co_induct = xtor_co_induct_thm,
+        dtor_ctors = steal #dtor_ctors (*too general types*),
+        ctor_dtors = steal #ctor_dtors (*too general types*),
+        ctor_injects = steal #ctor_injects (*too general types*),
+        dtor_injects = steal #dtor_injects (*too general types*),
+        xtor_map_thms = steal #xtor_map_thms (*too general types and terms*),
+        xtor_set_thmss = steal #xtor_set_thmss (*too general types and terms*),
+        xtor_rel_thms = steal #xtor_rel_thms (*too general types and terms*),
+        xtor_co_iter_thmss = transpose [xtor_un_fold_thms, xtor_co_rec_thms],
+        xtor_co_iter_o_map_thmss = steal #xtor_co_iter_o_map_thmss (*theorem about old constant*),
+        rel_xtor_co_induct_thm = rel_xtor_co_induct_thm}
+       |> morph_fp_result (Morphism.term_morphism "BNF" (singleton (Variable.polymorphic lthy))));
+  in
+    (fp_res, lthy)
+  end;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,394 @@
+(*  Title:      HOL/BNF/Tools/bnf_fp_n2m_sugar.ML
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2013
+
+Suggared flattening of nested to mutual (co)recursion.
+*)
+
+signature BNF_FP_N2M_SUGAR =
+sig
+  val unfold_let: term -> term
+  val dest_map: Proof.context -> string -> term -> term * term list
+
+  val mutualize_fp_sugars: BNF_FP_Util.fp_kind -> binding list -> typ list -> (term -> int list) ->
+    term list list list list -> BNF_FP_Def_Sugar.fp_sugar list -> local_theory ->
+    (BNF_FP_Def_Sugar.fp_sugar list
+     * (BNF_FP_Def_Sugar.lfp_sugar_thms option * BNF_FP_Def_Sugar.gfp_sugar_thms option))
+    * local_theory
+  val indexify_callsss: BNF_FP_Def_Sugar.fp_sugar -> (term * term list list) list ->
+    term list list list
+  val nested_to_mutual_fps: BNF_FP_Util.fp_kind -> binding list -> typ list -> (term -> int list) ->
+    (term * term list list) list list -> local_theory ->
+    (typ list * int list * BNF_FP_Def_Sugar.fp_sugar list
+     * (BNF_FP_Def_Sugar.lfp_sugar_thms option * BNF_FP_Def_Sugar.gfp_sugar_thms option))
+    * local_theory
+end;
+
+structure BNF_FP_N2M_Sugar : BNF_FP_N2M_SUGAR =
+struct
+
+open Ctr_Sugar
+open BNF_Util
+open BNF_Def
+open BNF_FP_Util
+open BNF_FP_Def_Sugar
+open BNF_FP_N2M
+
+val n2mN = "n2m_"
+
+type n2m_sugar = fp_sugar list * (lfp_sugar_thms option * gfp_sugar_thms option);
+
+structure Data = Generic_Data
+(
+  type T = n2m_sugar Typtab.table;
+  val empty = Typtab.empty;
+  val extend = I;
+  val merge = Typtab.merge (eq_fst (eq_list eq_fp_sugar));
+);
+
+fun morph_n2m_sugar phi (fp_sugars, (lfp_sugar_thms_opt, gfp_sugar_thms_opt)) =
+  (map (morph_fp_sugar phi) fp_sugars,
+   (Option.map (morph_lfp_sugar_thms phi) lfp_sugar_thms_opt,
+    Option.map (morph_gfp_sugar_thms phi) gfp_sugar_thms_opt));
+
+val transfer_n2m_sugar =
+  morph_n2m_sugar o Morphism.transfer_morphism o Proof_Context.theory_of;
+
+fun n2m_sugar_of ctxt =
+  Typtab.lookup (Data.get (Context.Proof ctxt))
+  #> Option.map (transfer_n2m_sugar ctxt);
+
+fun register_n2m_sugar key n2m_sugar =
+  Local_Theory.declaration {syntax = false, pervasive = false}
+    (fn phi => Data.map (Typtab.default (key, morph_n2m_sugar phi n2m_sugar)));
+
+fun unfold_let (Const (@{const_name Let}, _) $ arg1 $ arg2) = unfold_let (betapply (arg2, arg1))
+  | unfold_let (Const (@{const_name prod_case}, _) $ t) =
+    (case unfold_let t of
+      t' as Abs (s1, T1, Abs (s2, T2, _)) =>
+      let val v = Var ((s1 ^ s2, Term.maxidx_of_term t' + 1), HOLogic.mk_prodT (T1, T2)) in
+        lambda v (incr_boundvars 1 (betapplys (t', [HOLogic.mk_fst v, HOLogic.mk_snd v])))
+      end
+    | _ => t)
+  | unfold_let (t $ u) = betapply (unfold_let t, unfold_let u)
+  | unfold_let (Abs (s, T, t)) = Abs (s, T, unfold_let t)
+  | unfold_let t = t;
+
+fun mk_map_pattern ctxt s =
+  let
+    val bnf = the (bnf_of ctxt s);
+    val mapx = map_of_bnf bnf;
+    val live = live_of_bnf bnf;
+    val (f_Ts, _) = strip_typeN live (fastype_of mapx);
+    val fs = map_index (fn (i, T) => Var (("?f", i), T)) f_Ts;
+  in
+    (mapx, betapplys (mapx, fs))
+  end;
+
+fun dest_map ctxt s call =
+  let
+    val (map0, pat) = mk_map_pattern ctxt s;
+    val (_, tenv) = fo_match ctxt call pat;
+  in
+    (map0, Vartab.fold_rev (fn (_, (_, f)) => cons f) tenv [])
+  end;
+
+fun dest_abs_or_applied_map _ _ (Abs (_, _, t)) = (Term.dummy, [t])
+  | dest_abs_or_applied_map ctxt s (t1 $ _) = dest_map ctxt s t1;
+
+fun map_partition f xs =
+  fold_rev (fn x => fn (ys, (good, bad)) =>
+      case f x of SOME y => (y :: ys, (x :: good, bad)) | NONE => (ys, (good, x :: bad)))
+    xs ([], ([], []));
+
+fun key_of_fp_eqs fp fpTs fp_eqs =
+  Type (fp_case fp "l" "g", fpTs @ maps (fn (x, T) => [TFree x, T]) fp_eqs);
+
+(* TODO: test with sort constraints on As *)
+fun mutualize_fp_sugars fp bs fpTs get_indices callssss fp_sugars0 no_defs_lthy0 =
+  let
+    val thy = Proof_Context.theory_of no_defs_lthy0;
+
+    val qsotm = quote o Syntax.string_of_term no_defs_lthy0;
+
+    fun incompatible_calls t1 t2 =
+      error ("Incompatible " ^ co_prefix fp ^ "recursive calls: " ^ qsotm t1 ^ " vs. " ^ qsotm t2);
+    fun nested_self_call t =
+      error ("Unsupported nested self-call " ^ qsotm t);
+
+    val b_names = map Binding.name_of bs;
+    val fp_b_names = map base_name_of_typ fpTs;
+
+    val nn = length fpTs;
+
+    fun target_ctr_sugar_of_fp_sugar fpT ({T, index, ctr_sugars, ...} : fp_sugar) =
+      let
+        val rho = Vartab.fold (cons o apsnd snd) (Sign.typ_match thy (T, fpT) Vartab.empty) [];
+        val phi = Morphism.term_morphism "BNF" (Term.subst_TVars rho);
+      in
+        morph_ctr_sugar phi (nth ctr_sugars index)
+      end;
+
+    val ctr_defss = map (of_fp_sugar #ctr_defss) fp_sugars0;
+    val mapss = map (of_fp_sugar #mapss) fp_sugars0;
+    val ctr_sugars = map2 target_ctr_sugar_of_fp_sugar fpTs fp_sugars0;
+
+    val ctrss = map #ctrs ctr_sugars;
+    val ctr_Tss = map (map fastype_of) ctrss;
+
+    val As' = fold (fold Term.add_tfreesT) ctr_Tss [];
+    val As = map TFree As';
+
+    val ((Cs, Xs), no_defs_lthy) =
+      no_defs_lthy0
+      |> fold Variable.declare_typ As
+      |> mk_TFrees nn
+      ||>> variant_tfrees fp_b_names;
+
+    fun check_call_dead live_call call =
+      if null (get_indices call) then () else incompatible_calls live_call call;
+
+    fun freeze_fpTs_simple (T as Type (s, Ts)) =
+        (case find_index (curry (op =) T) fpTs of
+          ~1 => Type (s, map freeze_fpTs_simple Ts)
+        | kk => nth Xs kk)
+      | freeze_fpTs_simple T = T;
+
+    fun freeze_fpTs_map (fpT as Type (_, Ts')) (callss, (live_call :: _, dead_calls))
+        (T as Type (s, Ts)) =
+      if Ts' = Ts then
+        nested_self_call live_call
+      else
+        (List.app (check_call_dead live_call) dead_calls;
+         Type (s, map2 (freeze_fpTs fpT) (flatten_type_args_of_bnf (the (bnf_of no_defs_lthy s)) []
+           (transpose callss)) Ts))
+    and freeze_fpTs fpT calls (T as Type (s, _)) =
+        (case map_partition (try (snd o dest_map no_defs_lthy s)) calls of
+          ([], _) =>
+          (case map_partition (try (snd o dest_abs_or_applied_map no_defs_lthy s)) calls of
+            ([], _) => freeze_fpTs_simple T
+          | callsp => freeze_fpTs_map fpT callsp T)
+        | callsp => freeze_fpTs_map fpT callsp T)
+      | freeze_fpTs _ _ T = T;
+
+    val ctr_Tsss = map (map binder_types) ctr_Tss;
+    val ctrXs_Tsss = map3 (map2 o map2 o freeze_fpTs) fpTs callssss ctr_Tsss;
+    val ctrXs_sum_prod_Ts = map (mk_sumTN_balanced o map HOLogic.mk_tupleT) ctrXs_Tsss;
+    val ctr_Ts = map (body_type o hd) ctr_Tss;
+
+    val ns = map length ctr_Tsss;
+    val kss = map (fn n => 1 upto n) ns;
+    val mss = map (map length) ctr_Tsss;
+
+    val fp_eqs = map dest_TFree Xs ~~ ctrXs_sum_prod_Ts;
+    val key = key_of_fp_eqs fp fpTs fp_eqs;
+  in
+    (case n2m_sugar_of no_defs_lthy key of
+      SOME n2m_sugar => (n2m_sugar, no_defs_lthy)
+    | NONE =>
+      let
+        val base_fp_names = Name.variant_list [] fp_b_names;
+        val fp_bs = map2 (fn b_name => fn base_fp_name =>
+            Binding.qualify true b_name (Binding.name (n2mN ^ base_fp_name)))
+          b_names base_fp_names;
+
+        val (pre_bnfs, (fp_res as {xtor_co_iterss = xtor_co_iterss0, xtor_co_induct, dtor_injects,
+               dtor_ctors, xtor_co_iter_thmss, ...}, lthy)) =
+          fp_bnf (construct_mutualized_fp fp fpTs fp_sugars0) fp_bs As' fp_eqs no_defs_lthy;
+
+        val nesting_bnfs = nesty_bnfs lthy ctrXs_Tsss As;
+        val nested_bnfs = nesty_bnfs lthy ctrXs_Tsss Xs;
+
+        val ((xtor_co_iterss, iters_args_types, coiters_args_types), _) =
+          mk_co_iters_prelims fp ctr_Tsss fpTs Cs ns mss xtor_co_iterss0 lthy;
+
+        fun mk_binding b suf = Binding.suffix_name ("_" ^ suf) b;
+
+        val ((co_iterss, co_iter_defss), lthy) =
+          fold_map2 (fn b =>
+            (if fp = Least_FP then define_iters [foldN, recN] (the iters_args_types)
+             else define_coiters [unfoldN, corecN] (the coiters_args_types))
+              (mk_binding b) fpTs Cs) fp_bs xtor_co_iterss lthy
+          |>> split_list;
+
+        val ((co_inducts, un_fold_thmss, co_rec_thmss, disc_unfold_thmss, disc_corec_thmss,
+              sel_unfold_thmsss, sel_corec_thmsss), fp_sugar_thms) =
+          if fp = Least_FP then
+            derive_induct_iters_thms_for_types pre_bnfs (the iters_args_types) xtor_co_induct
+              xtor_co_iter_thmss nesting_bnfs nested_bnfs fpTs Cs Xs ctrXs_Tsss ctrss ctr_defss
+              co_iterss co_iter_defss lthy
+            |> `(fn ((_, induct, _), (fold_thmss, rec_thmss, _)) =>
+              ([induct], fold_thmss, rec_thmss, [], [], [], []))
+            ||> (fn info => (SOME info, NONE))
+          else
+            derive_coinduct_coiters_thms_for_types pre_bnfs (the coiters_args_types) xtor_co_induct
+              dtor_injects dtor_ctors xtor_co_iter_thmss nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss
+              ns ctr_defss ctr_sugars co_iterss co_iter_defss
+              (Proof_Context.export lthy no_defs_lthy) lthy
+            |> `(fn ((coinduct_thms_pairs, _), (unfold_thmss, corec_thmss, _),
+                    (disc_unfold_thmss, disc_corec_thmss, _), _,
+                    (sel_unfold_thmsss, sel_corec_thmsss, _)) =>
+              (map snd coinduct_thms_pairs, unfold_thmss, corec_thmss, disc_unfold_thmss,
+               disc_corec_thmss, sel_unfold_thmsss, sel_corec_thmsss))
+            ||> (fn info => (NONE, SOME info));
+
+        val phi = Proof_Context.export_morphism no_defs_lthy no_defs_lthy0;
+
+        fun mk_target_fp_sugar (kk, T) =
+          {T = T, fp = fp, index = kk, pre_bnfs = pre_bnfs, nested_bnfs = nested_bnfs,
+           nesting_bnfs = nesting_bnfs, fp_res = fp_res, ctr_defss = ctr_defss,
+           ctr_sugars = ctr_sugars, co_iterss = co_iterss, mapss = mapss, co_inducts = co_inducts,
+           co_iter_thmsss = transpose [un_fold_thmss, co_rec_thmss],
+           disc_co_itersss = transpose [disc_unfold_thmss, disc_corec_thmss],
+           sel_co_iterssss = transpose [sel_unfold_thmsss, sel_corec_thmsss]}
+          |> morph_fp_sugar phi;
+
+        val n2m_sugar = (map_index mk_target_fp_sugar fpTs, fp_sugar_thms);
+      in
+        (n2m_sugar, lthy |> register_n2m_sugar key n2m_sugar)
+      end)
+  end;
+
+fun indexify_callsss fp_sugar callsss =
+  let
+    val {ctrs, ...} = of_fp_sugar #ctr_sugars fp_sugar;
+    fun indexify_ctr ctr =
+      (case AList.lookup Term.aconv_untyped callsss ctr of
+        NONE => replicate (num_binder_types (fastype_of ctr)) []
+      | SOME callss => map (map (Envir.beta_eta_contract o unfold_let)) callss);
+  in
+    map indexify_ctr ctrs
+  end;
+
+fun retypargs tyargs (Type (s, _)) = Type (s, tyargs);
+
+fun fold_subtype_pairs f (T as Type (s, Ts), U as Type (s', Us)) =
+    f (T, U) #> (if s = s' then fold (fold_subtype_pairs f) (Ts ~~ Us) else I)
+  | fold_subtype_pairs f TU = f TU;
+
+fun nested_to_mutual_fps fp actual_bs actual_Ts get_indices actual_callssss0 lthy =
+  let
+    val qsoty = quote o Syntax.string_of_typ lthy;
+    val qsotys = space_implode " or " o map qsoty;
+
+    fun duplicate_datatype T = error (qsoty T ^ " is not mutually recursive with itself");
+    fun not_co_datatype0 T = error (qsoty T ^ " is not a " ^ co_prefix fp ^ "datatype");
+    fun not_co_datatype (T as Type (s, _)) =
+        if fp = Least_FP andalso
+           is_some (Datatype_Data.get_info (Proof_Context.theory_of lthy) s) then
+          error (qsoty T ^ " is not a new-style datatype (cf. \"datatype_new\")")
+        else
+          not_co_datatype0 T
+      | not_co_datatype T = not_co_datatype0 T;
+    fun not_mutually_nested_rec Ts1 Ts2 =
+      error (qsotys Ts1 ^ " is neither mutually recursive with " ^ qsotys Ts2 ^
+        " nor nested recursive via " ^ qsotys Ts2);
+
+    val _ = (case Library.duplicates (op =) actual_Ts of [] => () | T :: _ => duplicate_datatype T);
+
+    val perm_actual_Ts =
+      sort (prod_ord int_ord Term_Ord.typ_ord o pairself (`Term.size_of_typ)) actual_Ts;
+
+    fun the_ctrs_of (Type (s, Ts)) = map (mk_ctr Ts) (#ctrs (the (ctr_sugar_of lthy s)));
+
+    fun the_fp_sugar_of (T as Type (T_name, _)) =
+      (case fp_sugar_of lthy T_name of
+        SOME (fp_sugar as {fp = fp', ...}) => if fp = fp' then fp_sugar else not_co_datatype T
+      | NONE => not_co_datatype T);
+
+    fun gen_rhss_in gen_Ts rho subTs =
+      let
+        fun maybe_insert (T, Type (_, gen_tyargs)) =
+            if member (op =) subTs T then insert (op =) gen_tyargs else I
+          | maybe_insert _ = I;
+
+        val ctrs = maps the_ctrs_of gen_Ts;
+        val gen_ctr_Ts = maps (binder_types o fastype_of) ctrs;
+        val ctr_Ts = map (Term.typ_subst_atomic rho) gen_ctr_Ts;
+      in
+        fold (fold_subtype_pairs maybe_insert) (ctr_Ts ~~ gen_ctr_Ts) []
+      end;
+
+    fun gather_types _ _ num_groups seen gen_seen [] = (num_groups, seen, gen_seen)
+      | gather_types lthy rho num_groups seen gen_seen ((T as Type (_, tyargs)) :: Ts) =
+        let
+          val {fp_res = {Ts = mutual_Ts0, ...}, ...} = the_fp_sugar_of T;
+          val mutual_Ts = map (retypargs tyargs) mutual_Ts0;
+
+          val _ = seen = [] orelse exists (exists_subtype_in seen) mutual_Ts orelse
+            not_mutually_nested_rec mutual_Ts seen;
+
+          fun fresh_tyargs () =
+            let
+              (* The name "'z" is unlikely to clash with the context, yielding more cache hits. *)
+              val (gen_tyargs, lthy') =
+                variant_tfrees (replicate (length tyargs) "z") lthy
+                |>> map Logic.varifyT_global;
+              val rho' = (gen_tyargs ~~ tyargs) @ rho;
+            in
+              (rho', gen_tyargs, gen_seen, lthy')
+            end;
+
+          val (rho', gen_tyargs, gen_seen', lthy') =
+            if exists (exists_subtype_in seen) mutual_Ts then
+              (case gen_rhss_in gen_seen rho mutual_Ts of
+                [] => fresh_tyargs ()
+              | gen_tyargs :: gen_tyargss_tl =>
+                let
+                  val unify_pairs = split_list (maps (curry (op ~~) gen_tyargs) gen_tyargss_tl);
+                  val mgu = Type.raw_unifys unify_pairs Vartab.empty;
+                  val gen_tyargs' = map (Envir.subst_type mgu) gen_tyargs;
+                  val gen_seen' = map (Envir.subst_type mgu) gen_seen;
+                in
+                  (rho, gen_tyargs', gen_seen', lthy)
+                end)
+            else
+              fresh_tyargs ();
+
+          val gen_mutual_Ts = map (retypargs gen_tyargs) mutual_Ts0;
+          val Ts' = filter_out (member (op =) mutual_Ts) Ts;
+        in
+          gather_types lthy' rho' (num_groups + 1) (seen @ mutual_Ts) (gen_seen' @ gen_mutual_Ts)
+            Ts'
+        end
+      | gather_types _ _ _ _ _ (T :: _) = not_co_datatype T;
+
+    val (num_groups, perm_Ts, perm_gen_Ts) = gather_types lthy [] 0 [] [] perm_actual_Ts;
+    val perm_frozen_gen_Ts = map Logic.unvarifyT_global perm_gen_Ts;
+
+    val missing_Ts = perm_Ts |> subtract (op =) actual_Ts;
+    val Ts = actual_Ts @ missing_Ts;
+
+    val nn = length Ts;
+    val kks = 0 upto nn - 1;
+
+    val callssss0 = pad_list [] nn actual_callssss0;
+
+    val common_name = mk_common_name (map Binding.name_of actual_bs);
+    val bs = pad_list (Binding.name common_name) nn actual_bs;
+
+    fun permute xs = permute_like (op =) Ts perm_Ts xs;
+    fun unpermute perm_xs = permute_like (op =) perm_Ts Ts perm_xs;
+
+    val perm_bs = permute bs;
+    val perm_kks = permute kks;
+    val perm_callssss0 = permute callssss0;
+    val perm_fp_sugars0 = map (the o fp_sugar_of lthy o fst o dest_Type) perm_Ts;
+
+    val perm_callssss = map2 indexify_callsss perm_fp_sugars0 perm_callssss0;
+
+    val get_perm_indices = map (fn kk => find_index (curry (op =) kk) perm_kks) o get_indices;
+
+    val ((perm_fp_sugars, fp_sugar_thms), lthy) =
+      if num_groups > 1 then
+        mutualize_fp_sugars fp perm_bs perm_frozen_gen_Ts get_perm_indices perm_callssss
+          perm_fp_sugars0 lthy
+      else
+        ((perm_fp_sugars0, (NONE, NONE)), lthy);
+
+    val fp_sugars = unpermute perm_fp_sugars;
+  in
+    ((missing_Ts, perm_kks, fp_sugars, fp_sugar_thms), lthy)
+  end;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_fp_n2m_tactics.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,41 @@
+(*  Title:      HOL/BNF/Tools/bnf_fp_n2m_tactics.ML
+    Author:     Dmitriy Traytel, TU Muenchen
+    Copyright   2013
+
+Tactics for mutualization of nested (co)datatypes.
+*)
+
+signature BNF_FP_N2M_TACTICS =
+sig
+  val mk_rel_xtor_co_induct_tactic: BNF_FP_Util.fp_kind -> thm list -> thm list -> thm list ->
+    {prems: thm list, context: Proof.context} -> tactic
+end;
+
+structure BNF_FP_N2M_Tactics : BNF_FP_N2M_TACTICS =
+struct
+
+open BNF_Util
+open BNF_FP_Util
+
+fun mk_rel_xtor_co_induct_tactic fp co_inducts rel_defs rel_monos
+  {context = ctxt, prems = raw_C_IHs} =
+  let
+    val unfolds = map (fn def => unfold_thms ctxt (id_apply :: no_reflexive [def])) rel_defs;
+    val folded_C_IHs = map (fn thm => thm RS @{thm spec2} RS mp) raw_C_IHs;
+    val C_IHs = map2 (curry op |>) folded_C_IHs unfolds;
+    val C_IH_monos =
+      map3 (fn C_IH => fn mono => fn unfold =>
+        (mono RSN (2, @{thm rev_predicate2D}), C_IH)
+        |> fp = Greatest_FP ? swap
+        |> op RS
+        |> unfold)
+      folded_C_IHs rel_monos unfolds;
+  in
+    HEADGOAL (CONJ_WRAP_GEN' (rtac @{thm context_conjI})
+      (fn thm => rtac thm THEN_ALL_NEW (rotate_tac ~1 THEN'
+         REPEAT_ALL_NEW (FIRST' [eresolve_tac C_IHs, eresolve_tac C_IH_monos,
+           rtac @{thm order_refl}, atac, resolve_tac co_inducts])))
+    co_inducts)
+  end;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_fp_rec_sugar_util.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,67 @@
+(*  Title:      HOL/BNF/Tools/bnf_fp_rec_sugar_util.ML
+    Author:     Lorenz Panny, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2013
+
+Library for recursor and corecursor sugar.
+*)
+
+signature BNF_FP_REC_SUGAR_UTIL =
+sig
+  val indexed: 'a list -> int -> int list * int
+  val indexedd: 'a list list -> int -> int list list * int
+  val indexeddd: 'a list list list -> int -> int list list list * int
+  val indexedddd: 'a list list list list -> int -> int list list list list * int
+  val find_index_eq: ''a list -> ''a -> int
+  val finds: ('a * 'b -> bool) -> 'a list -> 'b list -> ('a * 'b list) list * 'b list
+
+  val drop_all: term -> term
+
+  val mk_partial_compN: int -> typ -> term -> term
+  val mk_partial_comp: typ -> typ -> term -> term
+  val mk_compN: int -> typ list -> term * term -> term
+  val mk_comp: typ list -> term * term -> term
+
+  val get_indices: ((binding * typ) * 'a) list -> term -> int list
+end;
+
+structure BNF_FP_Rec_Sugar_Util : BNF_FP_REC_SUGAR_UTIL =
+struct
+
+fun indexe _ h = (h, h + 1);
+fun indexed xs = fold_map indexe xs;
+fun indexedd xss = fold_map indexed xss;
+fun indexeddd xsss = fold_map indexedd xsss;
+fun indexedddd xssss = fold_map indexeddd xssss;
+
+fun find_index_eq hs h = find_index (curry (op =) h) hs;
+
+fun finds eq = fold_map (fn x => List.partition (curry eq x) #>> pair x);
+
+fun drop_all t =
+  subst_bounds (strip_qnt_vars @{const_name all} t |> map Free |> rev,
+    strip_qnt_body @{const_name all} t);
+
+fun mk_partial_comp gT fT g =
+  let val T = domain_type fT --> range_type gT in
+    Const (@{const_name Fun.comp}, gT --> fT --> T) $ g
+  end;
+
+fun mk_partial_compN 0 _ g = g
+  | mk_partial_compN n fT g =
+    let val g' = mk_partial_compN (n - 1) (range_type fT) g in
+      mk_partial_comp (fastype_of g') fT g'
+    end;
+
+fun mk_compN n bound_Ts (g, f) =
+  let val typof = curry fastype_of1 bound_Ts in
+    mk_partial_compN n (typof f) g $ f
+  end;
+
+val mk_comp = mk_compN 1;
+
+fun get_indices fixes t = map (fst #>> Binding.name_of #> Free) fixes
+  |> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE)
+  |> map_filter I;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_fp_util.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,635 @@
+(*  Title:      HOL/BNF/Tools/bnf_fp_util.ML
+    Author:     Dmitriy Traytel, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2012, 2013
+
+Shared library for the datatype and codatatype constructions.
+*)
+
+signature BNF_FP_UTIL =
+sig
+  datatype fp_kind = Least_FP | Greatest_FP
+  val fp_case: fp_kind -> 'a -> 'a -> 'a
+
+  type fp_result =
+    {Ts: typ list,
+     bnfs: BNF_Def.bnf list,
+     ctors: term list,
+     dtors: term list,
+     xtor_co_iterss: term list list,
+     xtor_co_induct: thm,
+     dtor_ctors: thm list,
+     ctor_dtors: thm list,
+     ctor_injects: thm list,
+     dtor_injects: thm list,
+     xtor_map_thms: thm list,
+     xtor_set_thmss: thm list list,
+     xtor_rel_thms: thm list,
+     xtor_co_iter_thmss: thm list list,
+     xtor_co_iter_o_map_thmss: thm list list,
+     rel_xtor_co_induct_thm: thm}
+
+  val morph_fp_result: morphism -> fp_result -> fp_result
+  val eq_fp_result: fp_result * fp_result -> bool
+  val un_fold_of: 'a list -> 'a
+  val co_rec_of: 'a list -> 'a
+
+  val time: Proof.context -> Timer.real_timer -> string -> Timer.real_timer
+
+  val IITN: string
+  val LevN: string
+  val algN: string
+  val behN: string
+  val bisN: string
+  val carTN: string
+  val caseN: string
+  val coN: string
+  val coinductN: string
+  val corecN: string
+  val ctorN: string
+  val ctor_dtorN: string
+  val ctor_exhaustN: string
+  val ctor_induct2N: string
+  val ctor_inductN: string
+  val ctor_injectN: string
+  val ctor_foldN: string
+  val ctor_fold_o_mapN: string
+  val ctor_fold_transferN: string
+  val ctor_fold_uniqueN: string
+  val ctor_mapN: string
+  val ctor_map_uniqueN: string
+  val ctor_recN: string
+  val ctor_rec_o_mapN: string
+  val ctor_rec_uniqueN: string
+  val ctor_relN: string
+  val ctor_set_inclN: string
+  val ctor_set_set_inclN: string
+  val disc_unfoldN: string
+  val disc_unfold_iffN: string
+  val disc_corecN: string
+  val disc_corec_iffN: string
+  val dtorN: string
+  val dtor_coinductN: string
+  val dtor_corecN: string
+  val dtor_corec_o_mapN: string
+  val dtor_corec_uniqueN: string
+  val dtor_ctorN: string
+  val dtor_exhaustN: string
+  val dtor_injectN: string
+  val dtor_mapN: string
+  val dtor_map_coinductN: string
+  val dtor_map_strong_coinductN: string
+  val dtor_map_uniqueN: string
+  val dtor_relN: string
+  val dtor_set_inclN: string
+  val dtor_set_set_inclN: string
+  val dtor_strong_coinductN: string
+  val dtor_unfoldN: string
+  val dtor_unfold_o_mapN: string
+  val dtor_unfold_transferN: string
+  val dtor_unfold_uniqueN: string
+  val exhaustN: string
+  val foldN: string
+  val hsetN: string
+  val hset_recN: string
+  val inductN: string
+  val injectN: string
+  val isNodeN: string
+  val lsbisN: string
+  val mapN: string
+  val map_uniqueN: string
+  val min_algN: string
+  val morN: string
+  val nchotomyN: string
+  val recN: string
+  val rel_coinductN: string
+  val rel_inductN: string
+  val rel_injectN: string
+  val rel_distinctN: string
+  val rvN: string
+  val sel_corecN: string
+  val set_inclN: string
+  val set_set_inclN: string
+  val sel_unfoldN: string
+  val setN: string
+  val simpsN: string
+  val strTN: string
+  val str_initN: string
+  val strong_coinductN: string
+  val sum_bdN: string
+  val sum_bdTN: string
+  val unfoldN: string
+  val uniqueN: string
+
+  (* TODO: Don't index set facts. Isabelle packages traditionally generate uniform names. *)
+  val mk_ctor_setN: int -> string
+  val mk_dtor_setN: int -> string
+  val mk_dtor_set_inductN: int -> string
+  val mk_set_inductN: int -> string
+
+  val co_prefix: fp_kind -> string
+
+  val base_name_of_typ: typ -> string
+  val mk_common_name: string list -> string
+
+  val split_conj_thm: thm -> thm list
+  val split_conj_prems: int -> thm -> thm
+
+  val mk_sumTN: typ list -> typ
+  val mk_sumTN_balanced: typ list -> typ
+
+  val mk_proj: typ -> int -> int -> term
+
+  val mk_convol: term * term -> term
+
+  val Inl_const: typ -> typ -> term
+  val Inr_const: typ -> typ -> term
+
+  val mk_Inl: typ -> term -> term
+  val mk_Inr: typ -> term -> term
+  val mk_InN: typ list -> term -> int -> term
+  val mk_InN_balanced: typ -> int -> term -> int -> term
+  val mk_sum_case: term * term -> term
+  val mk_sum_caseN: term list -> term
+  val mk_sum_caseN_balanced: term list -> term
+
+  val dest_sumT: typ -> typ * typ
+  val dest_sumTN: int -> typ -> typ list
+  val dest_sumTN_balanced: int -> typ -> typ list
+  val dest_tupleT: int -> typ -> typ list
+
+  val If_const: typ -> term
+
+  val mk_Field: term -> term
+  val mk_If: term -> term -> term -> term
+  val mk_union: term * term -> term
+
+  val mk_sumEN: int -> thm
+  val mk_sumEN_balanced: int -> thm
+  val mk_sumEN_tupled_balanced: int list -> thm
+  val mk_sum_casesN: int -> int -> thm
+  val mk_sum_casesN_balanced: int -> int -> thm
+
+  val fixpoint: ('a * 'a -> bool) -> ('a list -> 'a list) -> 'a list -> 'a list
+
+  val mk_rel_xtor_co_induct_thm: fp_kind -> term list -> term list -> term list -> term list ->
+    term list -> term list -> term list -> term list ->
+    ({prems: thm list, context: Proof.context} -> tactic) -> Proof.context -> thm
+  val mk_un_fold_transfer_thms: fp_kind -> term list -> term list -> term list -> term list ->
+    term list -> term list -> ({prems: thm list, context: Proof.context} -> tactic) ->
+    Proof.context -> thm list
+  val mk_xtor_un_fold_o_map_thms: fp_kind -> bool -> int -> thm -> thm list -> thm list ->
+    thm list -> thm list -> thm list
+
+  val mk_strong_coinduct_thm: thm -> thm list -> thm list -> Proof.context -> thm
+
+  val fp_bnf: (binding list -> (string * sort) list -> typ list * typ list list ->
+      BNF_Def.bnf list -> local_theory -> 'a) ->
+    binding list -> (string * sort) list -> ((string * sort) * typ) list -> local_theory ->
+    BNF_Def.bnf list * 'a
+end;
+
+structure BNF_FP_Util : BNF_FP_UTIL =
+struct
+
+open BNF_Comp
+open BNF_Def
+open BNF_Util
+
+datatype fp_kind = Least_FP | Greatest_FP;
+
+fun fp_case Least_FP l _ = l
+  | fp_case Greatest_FP _ g = g;
+
+type fp_result =
+  {Ts: typ list,
+   bnfs: BNF_Def.bnf list,
+   ctors: term list,
+   dtors: term list,
+   xtor_co_iterss: term list list,
+   xtor_co_induct: thm,
+   dtor_ctors: thm list,
+   ctor_dtors: thm list,
+   ctor_injects: thm list,
+   dtor_injects: thm list,
+   xtor_map_thms: thm list,
+   xtor_set_thmss: thm list list,
+   xtor_rel_thms: thm list,
+   xtor_co_iter_thmss: thm list list,
+   xtor_co_iter_o_map_thmss: thm list list,
+   rel_xtor_co_induct_thm: thm};
+
+fun morph_fp_result phi {Ts, bnfs, ctors, dtors, xtor_co_iterss, xtor_co_induct, dtor_ctors,
+    ctor_dtors, ctor_injects, dtor_injects, xtor_map_thms, xtor_set_thmss, xtor_rel_thms,
+    xtor_co_iter_thmss, xtor_co_iter_o_map_thmss, rel_xtor_co_induct_thm} =
+  {Ts = map (Morphism.typ phi) Ts,
+   bnfs = map (morph_bnf phi) bnfs,
+   ctors = map (Morphism.term phi) ctors,
+   dtors = map (Morphism.term phi) dtors,
+   xtor_co_iterss = map (map (Morphism.term phi)) xtor_co_iterss,
+   xtor_co_induct = Morphism.thm phi xtor_co_induct,
+   dtor_ctors = map (Morphism.thm phi) dtor_ctors,
+   ctor_dtors = map (Morphism.thm phi) ctor_dtors,
+   ctor_injects = map (Morphism.thm phi) ctor_injects,
+   dtor_injects = map (Morphism.thm phi) dtor_injects,
+   xtor_map_thms = map (Morphism.thm phi) xtor_map_thms,
+   xtor_set_thmss = map (map (Morphism.thm phi)) xtor_set_thmss,
+   xtor_rel_thms = map (Morphism.thm phi) xtor_rel_thms,
+   xtor_co_iter_thmss = map (map (Morphism.thm phi)) xtor_co_iter_thmss,
+   xtor_co_iter_o_map_thmss = map (map (Morphism.thm phi)) xtor_co_iter_o_map_thmss,
+   rel_xtor_co_induct_thm = Morphism.thm phi rel_xtor_co_induct_thm};
+
+fun eq_fp_result ({bnfs = bnfs1, ...} : fp_result, {bnfs = bnfs2, ...} : fp_result) =
+  eq_list eq_bnf (bnfs1, bnfs2);
+
+fun un_fold_of [f, _] = f;
+fun co_rec_of [_, r] = r;
+
+
+fun time ctxt timer msg = (if Config.get ctxt bnf_timing
+  then warning (msg ^ ": " ^ ATP_Util.string_of_time (Timer.checkRealTimer timer))
+  else (); Timer.startRealTimer ());
+
+val preN = "pre_"
+val rawN = "raw_"
+
+val coN = "co"
+val unN = "un"
+val algN = "alg"
+val IITN = "IITN"
+val foldN = "fold"
+val unfoldN = unN ^ foldN
+val uniqueN = "_unique"
+val transferN = "_transfer"
+val simpsN = "simps"
+val ctorN = "ctor"
+val dtorN = "dtor"
+val ctor_foldN = ctorN ^ "_" ^ foldN
+val dtor_unfoldN = dtorN ^ "_" ^ unfoldN
+val ctor_fold_uniqueN = ctor_foldN ^ uniqueN
+val ctor_fold_o_mapN = ctor_foldN ^ "_o_" ^ mapN
+val dtor_unfold_uniqueN = dtor_unfoldN ^ uniqueN
+val dtor_unfold_o_mapN = dtor_unfoldN ^ "_o_" ^ mapN
+val ctor_fold_transferN = ctor_foldN ^ transferN
+val dtor_unfold_transferN = dtor_unfoldN ^ transferN
+val ctor_mapN = ctorN ^ "_" ^ mapN
+val dtor_mapN = dtorN ^ "_" ^ mapN
+val map_uniqueN = mapN ^ uniqueN
+val ctor_map_uniqueN = ctorN ^ "_" ^ map_uniqueN
+val dtor_map_uniqueN = dtorN ^ "_" ^ map_uniqueN
+val min_algN = "min_alg"
+val morN = "mor"
+val bisN = "bis"
+val lsbisN = "lsbis"
+val sum_bdTN = "sbdT"
+val sum_bdN = "sbd"
+val carTN = "carT"
+val strTN = "strT"
+val isNodeN = "isNode"
+val LevN = "Lev"
+val rvN = "recover"
+val behN = "beh"
+val setN = "set"
+val mk_ctor_setN = prefix (ctorN ^ "_") o mk_setN
+val mk_dtor_setN = prefix (dtorN ^ "_") o mk_setN
+fun mk_set_inductN i = mk_setN i ^ "_induct"
+val mk_dtor_set_inductN = prefix (dtorN ^ "_") o mk_set_inductN
+
+val str_initN = "str_init"
+val recN = "rec"
+val corecN = coN ^ recN
+val ctor_recN = ctorN ^ "_" ^ recN
+val ctor_rec_o_mapN = ctor_recN ^ "_o_" ^ mapN
+val ctor_rec_uniqueN = ctor_recN ^ uniqueN
+val dtor_corecN = dtorN ^ "_" ^ corecN
+val dtor_corec_o_mapN = dtor_corecN ^ "_o_" ^ mapN
+val dtor_corec_uniqueN = dtor_corecN ^ uniqueN
+
+val ctor_dtorN = ctorN ^ "_" ^ dtorN
+val dtor_ctorN = dtorN ^ "_" ^ ctorN
+val nchotomyN = "nchotomy"
+val injectN = "inject"
+val exhaustN = "exhaust"
+val ctor_injectN = ctorN ^ "_" ^ injectN
+val ctor_exhaustN = ctorN ^ "_" ^ exhaustN
+val dtor_injectN = dtorN ^ "_" ^ injectN
+val dtor_exhaustN = dtorN ^ "_" ^ exhaustN
+val ctor_relN = ctorN ^ "_" ^ relN
+val dtor_relN = dtorN ^ "_" ^ relN
+val inductN = "induct"
+val coinductN = coN ^ inductN
+val ctor_inductN = ctorN ^ "_" ^ inductN
+val ctor_induct2N = ctor_inductN ^ "2"
+val dtor_map_coinductN = dtor_mapN ^ "_" ^ coinductN
+val dtor_coinductN = dtorN ^ "_" ^ coinductN
+val strong_coinductN = "strong_" ^ coinductN
+val dtor_map_strong_coinductN = dtor_mapN ^ "_" ^ strong_coinductN
+val dtor_strong_coinductN = dtorN ^ "_" ^ strong_coinductN
+val hsetN = "Hset"
+val hset_recN = hsetN ^ "_rec"
+val set_inclN = "set_incl"
+val ctor_set_inclN = ctorN ^ "_" ^ set_inclN
+val dtor_set_inclN = dtorN ^ "_" ^ set_inclN
+val set_set_inclN = "set_set_incl"
+val ctor_set_set_inclN = ctorN ^ "_" ^ set_set_inclN
+val dtor_set_set_inclN = dtorN ^ "_" ^ set_set_inclN
+
+val caseN = "case"
+val discN = "disc"
+val disc_unfoldN = discN ^ "_" ^ unfoldN
+val disc_corecN = discN ^ "_" ^ corecN
+val iffN = "_iff"
+val disc_unfold_iffN = discN ^ "_" ^ unfoldN ^ iffN
+val disc_corec_iffN = discN ^ "_" ^ corecN ^ iffN
+val distinctN = "distinct"
+val rel_distinctN = relN ^ "_" ^ distinctN
+val injectN = "inject"
+val rel_injectN = relN ^ "_" ^ injectN
+val rel_coinductN = relN ^ "_" ^ coinductN
+val rel_inductN = relN ^ "_" ^ inductN
+val selN = "sel"
+val sel_unfoldN = selN ^ "_" ^ unfoldN
+val sel_corecN = selN ^ "_" ^ corecN
+
+fun co_prefix fp = (if fp = Greatest_FP then "co" else "");
+
+fun add_components_of_typ (Type (s, Ts)) =
+    cons (Long_Name.base_name s) #> fold_rev add_components_of_typ Ts
+  | add_components_of_typ _ = I;
+
+fun base_name_of_typ T = space_implode "_" (add_components_of_typ T []);
+
+val mk_common_name = space_implode "_";
+
+fun dest_sumT (Type (@{type_name sum}, [T, T'])) = (T, T');
+
+fun dest_sumTN 1 T = [T]
+  | dest_sumTN n (Type (@{type_name sum}, [T, T'])) = T :: dest_sumTN (n - 1) T';
+
+val dest_sumTN_balanced = Balanced_Tree.dest dest_sumT;
+
+(* TODO: move something like this to "HOLogic"? *)
+fun dest_tupleT 0 @{typ unit} = []
+  | dest_tupleT 1 T = [T]
+  | dest_tupleT n (Type (@{type_name prod}, [T, T'])) = T :: dest_tupleT (n - 1) T';
+
+val mk_sumTN = Library.foldr1 mk_sumT;
+val mk_sumTN_balanced = Balanced_Tree.make mk_sumT;
+
+fun mk_proj T n k =
+  let val (binders, _) = strip_typeN n T in
+    fold_rev (fn T => fn t => Abs (Name.uu, T, t)) binders (Bound (n - k - 1))
+  end;
+
+fun mk_convol (f, g) =
+  let
+    val (fU, fTU) = `range_type (fastype_of f);
+    val ((gT, gU), gTU) = `dest_funT (fastype_of g);
+    val convolT = fTU --> gTU --> gT --> HOLogic.mk_prodT (fU, gU);
+  in Const (@{const_name convol}, convolT) $ f $ g end;
+
+fun Inl_const LT RT = Const (@{const_name Inl}, LT --> mk_sumT (LT, RT));
+fun mk_Inl RT t = Inl_const (fastype_of t) RT $ t;
+
+fun Inr_const LT RT = Const (@{const_name Inr}, RT --> mk_sumT (LT, RT));
+fun mk_Inr LT t = Inr_const LT (fastype_of t) $ t;
+
+fun mk_InN [_] t 1 = t
+  | mk_InN (_ :: Ts) t 1 = mk_Inl (mk_sumTN Ts) t
+  | mk_InN (LT :: Ts) t m = mk_Inr LT (mk_InN Ts t (m - 1))
+  | mk_InN Ts t _ = raise (TYPE ("mk_InN", Ts, [t]));
+
+fun mk_InN_balanced sum_T n t k =
+  let
+    fun repair_types T (Const (s as @{const_name Inl}, _) $ t) = repair_inj_types T s fst t
+      | repair_types T (Const (s as @{const_name Inr}, _) $ t) = repair_inj_types T s snd t
+      | repair_types _ t = t
+    and repair_inj_types T s get t =
+      let val T' = get (dest_sumT T) in
+        Const (s, T' --> T) $ repair_types T' t
+      end;
+  in
+    Balanced_Tree.access {left = mk_Inl dummyT, right = mk_Inr dummyT, init = t} n k
+    |> repair_types sum_T
+  end;
+
+fun mk_sum_case (f, g) =
+  let
+    val fT = fastype_of f;
+    val gT = fastype_of g;
+  in
+    Const (@{const_name sum_case},
+      fT --> gT --> mk_sumT (domain_type fT, domain_type gT) --> range_type fT) $ f $ g
+  end;
+
+val mk_sum_caseN = Library.foldr1 mk_sum_case;
+val mk_sum_caseN_balanced = Balanced_Tree.make mk_sum_case;
+
+fun If_const T = Const (@{const_name If}, HOLogic.boolT --> T --> T --> T);
+fun mk_If p t f = let val T = fastype_of t in If_const T $ p $ t $ f end;
+
+fun mk_Field r =
+  let val T = fst (dest_relT (fastype_of r));
+  in Const (@{const_name Field}, mk_relT (T, T) --> HOLogic.mk_setT T) $ r end;
+
+val mk_union = HOLogic.mk_binop @{const_name sup};
+
+(*dangerous; use with monotonic, converging functions only!*)
+fun fixpoint eq f X = if subset eq (f X, X) then X else fixpoint eq f (f X);
+
+(* stolen from "~~/src/HOL/Tools/Datatype/datatype_aux.ML" *)
+fun split_conj_thm th =
+  ((th RS conjunct1) :: split_conj_thm (th RS conjunct2)) handle THM _ => [th];
+
+fun split_conj_prems limit th =
+  let
+    fun split n i th =
+      if i = n then th else split n (i + 1) (conjI RSN (i, th)) handle THM _ => th;
+  in split limit 1 th end;
+
+fun mk_sumEN 1 = @{thm one_pointE}
+  | mk_sumEN 2 = @{thm sumE}
+  | mk_sumEN n =
+    (fold (fn i => fn thm => @{thm obj_sumE_f} RSN (i, thm)) (2 upto n - 1) @{thm obj_sumE}) OF
+      replicate n (impI RS allI);
+
+fun mk_obj_sumEN_balanced n =
+  Balanced_Tree.make (fn (thm1, thm2) => thm1 RSN (1, thm2 RSN (2, @{thm obj_sumE_f})))
+    (replicate n asm_rl);
+
+fun mk_sumEN_balanced' n all_impIs = mk_obj_sumEN_balanced n OF all_impIs RS @{thm obj_one_pointE};
+
+fun mk_sumEN_balanced 1 = @{thm one_pointE} (*optimization*)
+  | mk_sumEN_balanced 2 = @{thm sumE} (*optimization*)
+  | mk_sumEN_balanced n = mk_sumEN_balanced' n (replicate n (impI RS allI));
+
+fun mk_tupled_allIN 0 = @{thm unit_all_impI}
+  | mk_tupled_allIN 1 = @{thm impI[THEN allI]}
+  | mk_tupled_allIN 2 = @{thm prod_all_impI} (*optimization*)
+  | mk_tupled_allIN n = mk_tupled_allIN (n - 1) RS @{thm prod_all_impI_step};
+
+fun mk_sumEN_tupled_balanced ms =
+  let val n = length ms in
+    if forall (curry op = 1) ms then mk_sumEN_balanced n
+    else mk_sumEN_balanced' n (map mk_tupled_allIN ms)
+  end;
+
+fun mk_sum_casesN 1 1 = refl
+  | mk_sum_casesN _ 1 = @{thm sum.cases(1)}
+  | mk_sum_casesN 2 2 = @{thm sum.cases(2)}
+  | mk_sum_casesN n k = trans OF [@{thm sum_case_step(2)}, mk_sum_casesN (n - 1) (k - 1)];
+
+fun mk_sum_step base step thm =
+  if Thm.eq_thm_prop (thm, refl) then base else trans OF [step, thm];
+
+fun mk_sum_casesN_balanced 1 1 = refl
+  | mk_sum_casesN_balanced n k =
+    Balanced_Tree.access {left = mk_sum_step @{thm sum.cases(1)} @{thm sum_case_step(1)},
+      right = mk_sum_step @{thm sum.cases(2)} @{thm sum_case_step(2)}, init = refl} n k;
+
+fun mk_rel_xtor_co_induct_thm fp pre_rels pre_phis rels phis xs ys xtors xtor's tac lthy =
+  let
+    val pre_relphis = map (fn rel => Term.list_comb (rel, phis @ pre_phis)) pre_rels;
+    val relphis = map (fn rel => Term.list_comb (rel, phis)) rels;
+    fun mk_xtor fp' xtor x = if fp = fp' then xtor $ x else x;
+    val dtor = mk_xtor Greatest_FP;
+    val ctor = mk_xtor Least_FP;
+    fun flip f x y = if fp = Greatest_FP then f y x else f x y;
+
+    fun mk_prem pre_relphi phi x y xtor xtor' =
+      HOLogic.mk_Trueprop (list_all_free [x, y] (flip (curry HOLogic.mk_imp)
+        (pre_relphi $ (dtor xtor x) $ (dtor xtor' y)) (phi $ (ctor xtor x) $ (ctor xtor' y))));
+    val prems = map6 mk_prem pre_relphis pre_phis xs ys xtors xtor's;
+
+    val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+      (map2 (flip mk_leq) relphis pre_phis));
+  in
+    Goal.prove_sorry lthy (map (fst o dest_Free) (phis @ pre_phis)) prems concl tac
+    |> Thm.close_derivation
+    |> (fn thm => thm OF (replicate (length pre_rels) @{thm allI[OF allI[OF impI]]}))
+  end;
+
+fun mk_un_fold_transfer_thms fp pre_rels pre_phis rels phis un_folds un_folds' tac lthy =
+  let
+    val pre_relphis = map (fn rel => Term.list_comb (rel, phis @ pre_phis)) pre_rels;
+    val relphis = map (fn rel => Term.list_comb (rel, phis)) rels;
+    fun flip f x y = if fp = Greatest_FP then f y x else f x y;
+
+    val arg_rels = map2 (flip mk_fun_rel) pre_relphis pre_phis;
+    fun mk_transfer relphi pre_phi un_fold un_fold' =
+      fold_rev mk_fun_rel arg_rels (flip mk_fun_rel relphi pre_phi) $ un_fold $ un_fold';
+    val transfers = map4 mk_transfer relphis pre_phis un_folds un_folds';
+
+    val goal = fold_rev Logic.all (phis @ pre_phis)
+      (HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj transfers));
+  in
+    Goal.prove_sorry lthy [] [] goal tac
+    |> Thm.close_derivation
+    |> split_conj_thm
+  end;
+
+fun mk_xtor_un_fold_o_map_thms fp is_rec m un_fold_unique xtor_maps xtor_un_folds sym_map_comps
+    map_cong0s =
+  let
+    val n = length sym_map_comps;
+    val rewrite_comp_comp2 = fp_case fp @{thm rewriteR_comp_comp2} @{thm rewriteL_comp_comp2};
+    val rewrite_comp_comp = fp_case fp @{thm rewriteR_comp_comp} @{thm rewriteL_comp_comp};
+    val map_cong_passive_args1 = replicate m (fp_case fp @{thm id_o} @{thm o_id} RS fun_cong);
+    val map_cong_active_args1 = replicate n (if is_rec
+      then fp_case fp @{thm convol_o} @{thm o_sum_case} RS fun_cong
+      else refl);
+    val map_cong_passive_args2 = replicate m (fp_case fp @{thm o_id} @{thm id_o} RS fun_cong);
+    val map_cong_active_args2 = replicate n (if is_rec
+      then fp_case fp @{thm map_pair_o_convol_id} @{thm sum_case_o_sum_map_id}
+      else fp_case fp @{thm id_o} @{thm o_id} RS fun_cong);
+    fun mk_map_congs passive active = map (fn thm => thm OF (passive @ active) RS ext) map_cong0s;
+    val map_cong1s = mk_map_congs map_cong_passive_args1 map_cong_active_args1;
+    val map_cong2s = mk_map_congs map_cong_passive_args2 map_cong_active_args2;
+    
+    fun mk_rewrites map_congs = map2 (fn sym_map_comp => fn map_cong =>
+      mk_trans sym_map_comp map_cong RS rewrite_comp_comp) sym_map_comps map_congs;
+    val rewrite1s = mk_rewrites map_cong1s;
+    val rewrite2s = mk_rewrites map_cong2s;
+    val unique_prems =
+      map4 (fn xtor_map => fn un_fold => fn rewrite1 => fn rewrite2 =>
+        mk_trans (rewrite_comp_comp2 OF [xtor_map, un_fold])
+          (mk_trans rewrite1 (mk_sym rewrite2)))
+      xtor_maps xtor_un_folds rewrite1s rewrite2s;
+  in
+    split_conj_thm (un_fold_unique OF map (fp_case fp I mk_sym) unique_prems)
+  end;
+
+fun mk_strong_coinduct_thm coind rel_eqs rel_monos ctxt =
+  let
+    val n = Thm.nprems_of coind;
+    val m = Thm.nprems_of (hd rel_monos) - n;
+    fun mk_inst phi = (phi, mk_union (phi, HOLogic.eq_const (fst (dest_pred2T (fastype_of phi)))))
+      |> pairself (certify ctxt);
+    val insts = Term.add_vars (Thm.prop_of coind) [] |> rev |> take n |> map (mk_inst o Var);
+    fun mk_unfold rel_eq rel_mono =
+      let
+        val eq = iffD2 OF [rel_eq RS @{thm predicate2_eqD}, refl];
+        val mono = rel_mono OF (replicate m @{thm order_refl} @ replicate n @{thm eq_subset});
+      in eq RS (mono RS @{thm predicate2D}) RS @{thm eqTrueI} end;
+    val unfolds = map2 mk_unfold rel_eqs rel_monos @ @{thms sup_fun_def sup_bool_def
+      imp_disjL all_conj_distrib subst_eq_imp simp_thms(18,21,35)};
+  in
+    Thm.instantiate ([], insts) coind
+    |> unfold_thms ctxt unfolds
+  end;
+
+fun fp_bnf construct_fp bs resBs fp_eqs lthy =
+  let
+    val time = time lthy;
+    val timer = time (Timer.startRealTimer ());
+    val (Xs, rhsXs) = split_list fp_eqs;
+
+    (* FIXME: because of "@ Xs", the output could contain type variables that are not in the
+       input; also, "fp_sort" should put the "resBs" first and in the order in which they appear *)
+    fun fp_sort Ass =
+      subtract (op =) Xs (filter (fn T => exists (fn Ts => member (op =) Ts T) Ass) resBs) @ Xs;
+
+    fun raw_qualify base_b =
+      let val (_, qs, n) = Binding.dest base_b;
+      in
+        Binding.prefix_name rawN
+        #> fold_rev (fn (s, mand) => Binding.qualify mand s) (qs @ [(n, true)])
+        #> Binding.conceal
+      end;
+
+    val ((bnfs, (deadss, livess)), (unfold_set, lthy)) = apfst (apsnd split_list o split_list)
+      (fold_map2 (fn b => bnf_of_typ Smart_Inline (raw_qualify b) fp_sort Xs) bs rhsXs
+        (empty_unfolds, lthy));
+
+    fun norm_qualify i = Binding.qualify true (Binding.name_of (nth bs (Int.max (0, i - 1))))
+      #> Binding.conceal;
+
+    val Ass = map (map dest_TFree) livess;
+    val resDs = fold (subtract (op =)) Ass resBs;
+    val Ds = fold (fold Term.add_tfreesT) deadss [];
+
+    val timer = time (timer "Construction of BNFs");
+
+    val ((kill_poss, _), (bnfs', (unfold_set', lthy'))) =
+      normalize_bnfs norm_qualify Ass Ds fp_sort bnfs unfold_set lthy;
+
+    val Dss = map3 (append oo map o nth) livess kill_poss deadss;
+
+    fun pre_qualify b = Binding.qualify false (Binding.name_of b)
+      #> Config.get lthy' bnf_note_all = false ? Binding.conceal;
+
+    val ((pre_bnfs, deadss), lthy'') =
+      fold_map3 (fn b => seal_bnf (pre_qualify b) unfold_set' (Binding.prefix_name preN b))
+        bs Dss bnfs' lthy'
+      |>> split_list;
+
+    val timer = time (timer "Normalization & sealing of BNFs");
+
+    val res = construct_fp bs resBs (map TFree resDs, deadss) pre_bnfs lthy'';
+
+    val timer = time (timer "FP construction in total");
+  in
+    timer; (pre_bnfs, res)
+  end;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_gfp.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,2827 @@
+(*  Title:      HOL/BNF/Tools/bnf_gfp.ML
+    Author:     Dmitriy Traytel, TU Muenchen
+    Author:     Andrei Popescu, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2012
+
+Codatatype construction.
+*)
+
+signature BNF_GFP =
+sig
+  val construct_gfp: mixfix list -> binding list -> binding list -> binding list list ->
+    binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list ->
+    local_theory -> BNF_FP_Util.fp_result * local_theory
+end;
+
+structure BNF_GFP : BNF_GFP =
+struct
+
+open BNF_Def
+open BNF_Util
+open BNF_Tactics
+open BNF_Comp
+open BNF_FP_Util
+open BNF_FP_Def_Sugar
+open BNF_GFP_Rec_Sugar
+open BNF_GFP_Util
+open BNF_GFP_Tactics
+
+datatype wit_tree = Wit_Leaf of int | Wit_Node of (int * int * int list) * wit_tree list;
+
+fun mk_tree_args (I, T) (I', Ts) = (sort_distinct int_ord (I @ I'), T :: Ts);
+
+fun finish Iss m seen i (nwit, I) =
+  let
+    val treess = map (fn j =>
+        if j < m orelse member (op =) seen j then [([j], Wit_Leaf j)]
+        else
+          map_index (finish Iss m (insert (op =) j seen) j) (nth Iss (j - m))
+          |> flat
+          |> minimize_wits)
+      I;
+  in
+    map (fn (I, t) => (I, Wit_Node ((i - m, nwit, filter (fn i => i < m) I), t)))
+      (fold_rev (map_product mk_tree_args) treess [([], [])])
+    |> minimize_wits
+  end;
+
+fun tree_to_ctor_wit vars _ _ (Wit_Leaf j) = ([j], nth vars j)
+  | tree_to_ctor_wit vars ctors witss (Wit_Node ((i, nwit, I), subtrees)) =
+     (I, nth ctors i $ (Term.list_comb (snd (nth (nth witss i) nwit),
+       map (snd o tree_to_ctor_wit vars ctors witss) subtrees)));
+
+fun tree_to_coind_wits _ (Wit_Leaf _) = []
+  | tree_to_coind_wits lwitss (Wit_Node ((i, nwit, I), subtrees)) =
+     ((i, I), nth (nth lwitss i) nwit) :: maps (tree_to_coind_wits lwitss) subtrees;
+
+(*all BNFs have the same lives*)
+fun construct_gfp mixfixes map_bs rel_bs set_bss0 bs resBs (resDs, Dss) bnfs lthy =
+  let
+    val time = time lthy;
+    val timer = time (Timer.startRealTimer ());
+
+    val live = live_of_bnf (hd bnfs);
+    val n = length bnfs; (*active*)
+    val ks = 1 upto n;
+    val m = live - n; (*passive, if 0 don't generate a new BNF*)
+    val ls = 1 upto m;
+
+    val note_all = Config.get lthy bnf_note_all;
+    val b_names = map Binding.name_of bs;
+    val b_name = mk_common_name b_names;
+    val b = Binding.name b_name;
+    val mk_internal_b = Binding.name #> Binding.prefix true b_name #> Binding.conceal;
+    fun mk_internal_bs name =
+      map (fn b =>
+        Binding.prefix true b_name (Binding.prefix_name (name ^ "_") b) |> Binding.conceal) bs;
+    val external_bs = map2 (Binding.prefix false) b_names bs
+      |> note_all = false ? map Binding.conceal;
+
+    (* TODO: check if m, n, etc., are sane *)
+
+    val deads = fold (union (op =)) Dss resDs;
+    val names_lthy = fold Variable.declare_typ deads lthy;
+    val passives = map fst (subtract (op = o apsnd TFree) deads resBs);
+
+    (* tvars *)
+    val ((((((passiveAs, activeAs), passiveBs), activeBs), passiveCs), activeCs), idxT) = names_lthy
+      |> variant_tfrees passives
+      ||>> mk_TFrees n
+      ||>> variant_tfrees passives
+      ||>> mk_TFrees n
+      ||>> mk_TFrees m
+      ||>> mk_TFrees n
+      ||> fst o mk_TFrees 1
+      ||> the_single;
+
+    val allAs = passiveAs @ activeAs;
+    val allBs' = passiveBs @ activeBs;
+    val Ass = replicate n allAs;
+    val allBs = passiveAs @ activeBs;
+    val Bss = replicate n allBs;
+    val allCs = passiveAs @ activeCs;
+    val allCs' = passiveBs @ activeCs;
+    val Css' = replicate n allCs';
+
+    (* types *)
+    val dead_poss =
+      map (fn x => if member (op =) deads (TFree x) then SOME (TFree x) else NONE) resBs;
+    fun mk_param NONE passive = (hd passive, tl passive)
+      | mk_param (SOME a) passive = (a, passive);
+    val mk_params = fold_map mk_param dead_poss #> fst;
+
+    fun mk_FTs Ts = map2 (fn Ds => mk_T_of_bnf Ds Ts) Dss bnfs;
+    val (params, params') = `(map Term.dest_TFree) (mk_params passiveAs);
+    val (dead_params, dead_params') = `(map Term.dest_TFree) (subtract (op =) passiveAs params');
+    val FTsAs = mk_FTs allAs;
+    val FTsBs = mk_FTs allBs;
+    val FTsCs = mk_FTs allCs;
+    val ATs = map HOLogic.mk_setT passiveAs;
+    val BTs = map HOLogic.mk_setT activeAs;
+    val B'Ts = map HOLogic.mk_setT activeBs;
+    val B''Ts = map HOLogic.mk_setT activeCs;
+    val sTs = map2 (fn T => fn U => T --> U) activeAs FTsAs;
+    val s'Ts = map2 (fn T => fn U => T --> U) activeBs FTsBs;
+    val s''Ts = map2 (fn T => fn U => T --> U) activeCs FTsCs;
+    val fTs = map2 (fn T => fn U => T --> U) activeAs activeBs;
+    val self_fTs = map (fn T => T --> T) activeAs;
+    val gTs = map2 (fn T => fn U => T --> U) activeBs activeCs;
+    val all_gTs = map2 (fn T => fn U => T --> U) allBs allCs';
+    val RTs = map2 (fn T => fn U => HOLogic.mk_prodT (T, U)) activeAs activeBs;
+    val sRTs = map2 (fn T => fn U => HOLogic.mk_prodT (T, U)) activeAs activeAs;
+    val R'Ts = map2 (fn T => fn U => HOLogic.mk_prodT (T, U)) activeBs activeCs;
+    val setsRTs = map HOLogic.mk_setT sRTs;
+    val setRTs = map HOLogic.mk_setT RTs;
+    val all_sbisT = HOLogic.mk_tupleT setsRTs;
+    val setR'Ts = map HOLogic.mk_setT R'Ts;
+    val FRTs = mk_FTs (passiveAs @ RTs);
+    val sumBsAs = map2 (curry mk_sumT) activeBs activeAs;
+    val sumFTs = mk_FTs (passiveAs @ sumBsAs);
+    val sum_sTs = map2 (fn T => fn U => T --> U) activeAs sumFTs;
+
+    (* terms *)
+    val mapsAsAs = map4 mk_map_of_bnf Dss Ass Ass bnfs;
+    val mapsAsBs = map4 mk_map_of_bnf Dss Ass Bss bnfs;
+    val mapsBsCs' = map4 mk_map_of_bnf Dss Bss Css' bnfs;
+    val mapsAsCs' = map4 mk_map_of_bnf Dss Ass Css' bnfs;
+    val map_Inls = map4 mk_map_of_bnf Dss Bss (replicate n (passiveAs @ sumBsAs)) bnfs;
+    val map_Inls_rev = map4 mk_map_of_bnf Dss (replicate n (passiveAs @ sumBsAs)) Bss bnfs;
+    val map_fsts = map4 mk_map_of_bnf Dss (replicate n (passiveAs @ RTs)) Ass bnfs;
+    val map_snds = map4 mk_map_of_bnf Dss (replicate n (passiveAs @ RTs)) Bss bnfs;
+    fun mk_setss Ts = map3 mk_sets_of_bnf (map (replicate live) Dss)
+      (map (replicate live) (replicate n Ts)) bnfs;
+    val setssAs = mk_setss allAs;
+    val setssAs' = transpose setssAs;
+    val bis_setss = mk_setss (passiveAs @ RTs);
+    val relsAsBs = map4 mk_rel_of_bnf Dss Ass Bss bnfs;
+    val bds = map3 mk_bd_of_bnf Dss Ass bnfs;
+    val sum_bd = Library.foldr1 (uncurry mk_csum) bds;
+    val sum_bdT = fst (dest_relT (fastype_of sum_bd));
+
+    val emptys = map (fn T => HOLogic.mk_set T []) passiveAs;
+    val Zeros = map (fn empty =>
+     HOLogic.mk_tuple (map (fn U => absdummy U empty) activeAs)) emptys;
+    val hrecTs = map fastype_of Zeros;
+    val hsetTs = map (fn hrecT => Library.foldr (op -->) (sTs, HOLogic.natT --> hrecT)) hrecTs;
+
+    val (((((((((((((((((((((((((((((((((((zs, zs'), zs_copy), zs_copy2), z's), (ys, ys')),
+      As), Bs), Bs_copy), B's), B''s), ss), sum_ss), s's), s''s), fs), fs_copy),
+      self_fs), gs), all_gs), xFs), yFs), yFs_copy), RFs), (Rtuple, Rtuple')), (hrecs, hrecs')),
+      (nat, nat')), Rs), Rs_copy), R's), sRs), (idx, idx')), Idx), Ris), Kss), names_lthy) = lthy
+      |> mk_Frees' "b" activeAs
+      ||>> mk_Frees "b" activeAs
+      ||>> mk_Frees "b" activeAs
+      ||>> mk_Frees "b" activeBs
+      ||>> mk_Frees' "y" passiveAs
+      ||>> mk_Frees "A" ATs
+      ||>> mk_Frees "B" BTs
+      ||>> mk_Frees "B" BTs
+      ||>> mk_Frees "B'" B'Ts
+      ||>> mk_Frees "B''" B''Ts
+      ||>> mk_Frees "s" sTs
+      ||>> mk_Frees "sums" sum_sTs
+      ||>> mk_Frees "s'" s'Ts
+      ||>> mk_Frees "s''" s''Ts
+      ||>> mk_Frees "f" fTs
+      ||>> mk_Frees "f" fTs
+      ||>> mk_Frees "f" self_fTs
+      ||>> mk_Frees "g" gTs
+      ||>> mk_Frees "g" all_gTs
+      ||>> mk_Frees "x" FTsAs
+      ||>> mk_Frees "y" FTsBs
+      ||>> mk_Frees "y" FTsBs
+      ||>> mk_Frees "x" FRTs
+      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "Rtuple") all_sbisT
+      ||>> mk_Frees' "rec" hrecTs
+      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "n") HOLogic.natT
+      ||>> mk_Frees "R" setRTs
+      ||>> mk_Frees "R" setRTs
+      ||>> mk_Frees "R'" setR'Ts
+      ||>> mk_Frees "R" setsRTs
+      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "i") idxT
+      ||>> yield_singleton (mk_Frees "I") (HOLogic.mk_setT idxT)
+      ||>> mk_Frees "Ri" (map (fn T => idxT --> T) setRTs)
+      ||>> mk_Freess "K" (map (fn AT => map (fn T => T --> AT) activeAs) ATs);
+
+    val passive_UNIVs = map HOLogic.mk_UNIV passiveAs;
+    val passive_Id_ons = map mk_Id_on As;
+    val active_UNIVs = map HOLogic.mk_UNIV activeAs;
+    val sum_UNIVs = map HOLogic.mk_UNIV sumBsAs;
+    val passive_ids = map HOLogic.id_const passiveAs;
+    val active_ids = map HOLogic.id_const activeAs;
+    val Inls = map2 Inl_const activeBs activeAs;
+    val fsts = map fst_const RTs;
+    val snds = map snd_const RTs;
+
+    (* thms *)
+    val bd_card_orders = map bd_card_order_of_bnf bnfs;
+    val bd_card_order = hd bd_card_orders
+    val bd_Card_orders = map bd_Card_order_of_bnf bnfs;
+    val bd_Card_order = hd bd_Card_orders;
+    val bd_Cinfinites = map bd_Cinfinite_of_bnf bnfs;
+    val bd_Cinfinite = hd bd_Cinfinites;
+    val in_monos = map in_mono_of_bnf bnfs;
+    val map_comp0s = map map_comp0_of_bnf bnfs;
+    val sym_map_comps = map mk_sym map_comp0s;
+    val map_comps = map map_comp_of_bnf bnfs;
+    val map_cong0s = map map_cong0_of_bnf bnfs;
+    val map_id0s = map map_id0_of_bnf bnfs;
+    val map_ids = map map_id_of_bnf bnfs;
+    val set_bdss = map set_bd_of_bnf bnfs;
+    val set_mapss = map set_map_of_bnf bnfs;
+    val rel_congs = map rel_cong_of_bnf bnfs;
+    val rel_converseps = map rel_conversep_of_bnf bnfs;
+    val rel_Grps = map rel_Grp_of_bnf bnfs;
+    val rel_OOs = map rel_OO_of_bnf bnfs;
+    val rel_OO_Grps = map rel_OO_Grp_of_bnf bnfs;
+
+    val timer = time (timer "Extracted terms & thms");
+
+    (* derived thms *)
+
+    (*map g1 ... gm g(m+1) ... g(m+n) (map id ... id f(m+1) ... f(m+n) x) =
+      map g1 ... gm (g(m+1) o f(m+1)) ... (g(m+n) o f(m+n)) x*)
+    fun mk_map_comp_id x mapAsBs mapBsCs mapAsCs map_comp0 =
+      let
+        val lhs = Term.list_comb (mapBsCs, all_gs) $
+          (Term.list_comb (mapAsBs, passive_ids @ fs) $ x);
+        val rhs =
+          Term.list_comb (mapAsCs, take m all_gs @ map HOLogic.mk_comp (drop m all_gs ~~ fs)) $ x;
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (x :: fs @ all_gs) (mk_Trueprop_eq (lhs, rhs)))
+          (K (mk_map_comp_id_tac map_comp0))
+        |> Thm.close_derivation
+      end;
+
+    val map_comp_id_thms = map5 mk_map_comp_id xFs mapsAsBs mapsBsCs' mapsAsCs' map_comps;
+
+    (*forall a : set(m+1) x. f(m+1) a = a; ...; forall a : set(m+n) x. f(m+n) a = a ==>
+      map id ... id f(m+1) ... f(m+n) x = x*)
+    fun mk_map_cong0L x mapAsAs sets map_cong0 map_id =
+      let
+        fun mk_prem set f z z' =
+          HOLogic.mk_Trueprop
+            (mk_Ball (set $ x) (Term.absfree z' (HOLogic.mk_eq (f $ z, z))));
+        val prems = map4 mk_prem (drop m sets) self_fs zs zs';
+        val goal = mk_Trueprop_eq (Term.list_comb (mapAsAs, passive_ids @ self_fs) $ x, x);
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (x :: self_fs) (Logic.list_implies (prems, goal)))
+          (K (mk_map_cong0L_tac m map_cong0 map_id))
+        |> Thm.close_derivation
+      end;
+
+    val map_cong0L_thms = map5 mk_map_cong0L xFs mapsAsAs setssAs map_cong0s map_ids;
+    val in_mono'_thms = map (fn thm =>
+      (thm OF (replicate m subset_refl)) RS @{thm set_mp}) in_monos;
+
+    val map_arg_cong_thms =
+      let
+        val prems = map2 (curry mk_Trueprop_eq) yFs yFs_copy;
+        val maps = map (fn mapx => Term.list_comb (mapx, all_gs)) mapsBsCs';
+        val concls =
+          map3 (fn x => fn y => fn mapx => mk_Trueprop_eq (mapx $ x, mapx $ y)) yFs yFs_copy maps;
+        val goals =
+          map4 (fn prem => fn concl => fn x => fn y =>
+            fold_rev Logic.all (x :: y :: all_gs) (Logic.mk_implies (prem, concl)))
+          prems concls yFs yFs_copy;
+      in
+        map (fn goal => Goal.prove_sorry lthy [] [] goal
+          (K ((hyp_subst_tac lthy THEN' rtac refl) 1)) |> Thm.close_derivation) goals
+      end;
+
+    val timer = time (timer "Derived simple theorems");
+
+    (* coalgebra *)
+
+    val coalg_bind = mk_internal_b (coN ^ algN) ;
+    val coalg_name = Binding.name_of coalg_bind;
+    val coalg_def_bind = (Thm.def_binding coalg_bind, []);
+
+    (*forall i = 1 ... n: (\<forall>x \<in> Bi. si \<in> Fi_in A1 .. Am B1 ... Bn)*)
+    val coalg_spec =
+      let
+        val coalgT = Library.foldr (op -->) (ATs @ BTs @ sTs, HOLogic.boolT);
+
+        val ins = map3 mk_in (replicate n (As @ Bs)) setssAs FTsAs;
+        fun mk_coalg_conjunct B s X z z' =
+          mk_Ball B (Term.absfree z' (HOLogic.mk_mem (s $ z, X)));
+
+        val lhs = Term.list_comb (Free (coalg_name, coalgT), As @ Bs @ ss);
+        val rhs = Library.foldr1 HOLogic.mk_conj (map5 mk_coalg_conjunct Bs ss ins zs zs')
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((coalg_free, (_, coalg_def_free)), (lthy, lthy_old)) =
+      lthy
+      |> Specification.definition (SOME (coalg_bind, NONE, NoSyn), (coalg_def_bind, coalg_spec))
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+    val coalg = fst (Term.dest_Const (Morphism.term phi coalg_free));
+    val coalg_def = Morphism.thm phi coalg_def_free;
+
+    fun mk_coalg As Bs ss =
+      let
+        val args = As @ Bs @ ss;
+        val Ts = map fastype_of args;
+        val coalgT = Library.foldr (op -->) (Ts, HOLogic.boolT);
+      in
+        Term.list_comb (Const (coalg, coalgT), args)
+      end;
+
+    val coalg_prem = HOLogic.mk_Trueprop (mk_coalg As Bs ss);
+
+    val coalg_in_thms = map (fn i =>
+      coalg_def RS iffD1 RS mk_conjunctN n i RS bspec) ks
+
+    val coalg_set_thmss =
+      let
+        val coalg_prem = HOLogic.mk_Trueprop (mk_coalg As Bs ss);
+        fun mk_prem x B = HOLogic.mk_Trueprop (HOLogic.mk_mem (x, B));
+        fun mk_concl s x B set = HOLogic.mk_Trueprop (mk_leq (set $ (s $ x)) B);
+        val prems = map2 mk_prem zs Bs;
+        val conclss = map3 (fn s => fn x => fn sets => map2 (mk_concl s x) (As @ Bs) sets)
+          ss zs setssAs;
+        val goalss = map3 (fn x => fn prem => fn concls => map (fn concl =>
+          fold_rev Logic.all (x :: As @ Bs @ ss)
+            (Logic.list_implies (coalg_prem :: [prem], concl))) concls) zs prems conclss;
+      in
+        map (fn goals => map (fn goal => Goal.prove_sorry lthy [] [] goal
+          (K (mk_coalg_set_tac coalg_def)) |> Thm.close_derivation) goals) goalss
+      end;
+
+    fun mk_tcoalg ATs BTs = mk_coalg (map HOLogic.mk_UNIV ATs) (map HOLogic.mk_UNIV BTs);
+
+    val tcoalg_thm =
+      let
+        val goal = fold_rev Logic.all ss
+          (HOLogic.mk_Trueprop (mk_tcoalg passiveAs activeAs ss))
+      in
+        Goal.prove_sorry lthy [] [] goal
+          (K (stac coalg_def 1 THEN CONJ_WRAP
+            (K (EVERY' [rtac ballI, rtac CollectI,
+              CONJ_WRAP' (K (EVERY' [rtac @{thm subset_UNIV}])) allAs] 1)) ss))
+        |> Thm.close_derivation
+      end;
+
+    val timer = time (timer "Coalgebra definition & thms");
+
+    (* morphism *)
+
+    val mor_bind = mk_internal_b morN;
+    val mor_name = Binding.name_of mor_bind;
+    val mor_def_bind = (Thm.def_binding mor_bind, []);
+
+    (*fbetw) forall i = 1 ... n: (\<forall>x \<in> Bi. fi x \<in> B'i)*)
+    (*mor) forall i = 1 ... n: (\<forall>x \<in> Bi.
+       Fi_map id ... id f1 ... fn (si x) = si' (fi x)*)
+    val mor_spec =
+      let
+        val morT = Library.foldr (op -->) (BTs @ sTs @ B'Ts @ s'Ts @ fTs, HOLogic.boolT);
+
+        fun mk_fbetw f B1 B2 z z' =
+          mk_Ball B1 (Term.absfree z' (HOLogic.mk_mem (f $ z, B2)));
+        fun mk_mor B mapAsBs f s s' z z' =
+          mk_Ball B (Term.absfree z' (HOLogic.mk_eq
+            (Term.list_comb (mapAsBs, passive_ids @ fs @ [s $ z]), s' $ (f $ z))));
+        val lhs = Term.list_comb (Free (mor_name, morT), Bs @ ss @ B's @ s's @ fs);
+        val rhs = HOLogic.mk_conj
+          (Library.foldr1 HOLogic.mk_conj (map5 mk_fbetw fs Bs B's zs zs'),
+           Library.foldr1 HOLogic.mk_conj (map7 mk_mor Bs mapsAsBs fs ss s's zs zs'))
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((mor_free, (_, mor_def_free)), (lthy, lthy_old)) =
+      lthy
+      |> Specification.definition (SOME (mor_bind, NONE, NoSyn), (mor_def_bind, mor_spec))
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+    val mor = fst (Term.dest_Const (Morphism.term phi mor_free));
+    val mor_def = Morphism.thm phi mor_def_free;
+
+    fun mk_mor Bs1 ss1 Bs2 ss2 fs =
+      let
+        val args = Bs1 @ ss1 @ Bs2 @ ss2 @ fs;
+        val Ts = map fastype_of (Bs1 @ ss1 @ Bs2 @ ss2 @ fs);
+        val morT = Library.foldr (op -->) (Ts, HOLogic.boolT);
+      in
+        Term.list_comb (Const (mor, morT), args)
+      end;
+
+    val mor_prem = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs);
+
+    val (mor_image_thms, morE_thms) =
+      let
+        val prem = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs);
+        fun mk_image_goal f B1 B2 = fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs)
+          (Logic.mk_implies (prem, HOLogic.mk_Trueprop (mk_leq (mk_image f $ B1) B2)));
+        val image_goals = map3 mk_image_goal fs Bs B's;
+        fun mk_elim_goal B mapAsBs f s s' x =
+          fold_rev Logic.all (x :: Bs @ ss @ B's @ s's @ fs)
+            (Logic.list_implies ([prem, HOLogic.mk_Trueprop (HOLogic.mk_mem (x, B))],
+              mk_Trueprop_eq (Term.list_comb (mapAsBs, passive_ids @ fs @ [s $ x]), s' $ (f $ x))));
+        val elim_goals = map6 mk_elim_goal Bs mapsAsBs fs ss s's zs;
+        fun prove goal =
+          Goal.prove_sorry lthy [] [] goal (K (mk_mor_elim_tac mor_def))
+          |> Thm.close_derivation;
+      in
+        (map prove image_goals, map prove elim_goals)
+      end;
+
+    val mor_image'_thms = map (fn thm => @{thm set_mp} OF [thm, imageI]) mor_image_thms;
+
+    val mor_incl_thm =
+      let
+        val prems = map2 (HOLogic.mk_Trueprop oo mk_leq) Bs Bs_copy;
+        val concl = HOLogic.mk_Trueprop (mk_mor Bs ss Bs_copy ss active_ids);
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (Bs @ ss @ Bs_copy) (Logic.list_implies (prems, concl)))
+          (K (mk_mor_incl_tac mor_def map_ids))
+        |> Thm.close_derivation
+      end;
+
+    val mor_id_thm = mor_incl_thm OF (replicate n subset_refl);
+
+    val mor_comp_thm =
+      let
+        val prems =
+          [HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs),
+           HOLogic.mk_Trueprop (mk_mor B's s's B''s s''s gs)];
+        val concl =
+          HOLogic.mk_Trueprop (mk_mor Bs ss B''s s''s (map2 (curry HOLogic.mk_comp) gs fs));
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (Bs @ ss @ B's @ s's @ B''s @ s''s @ fs @ gs)
+            (Logic.list_implies (prems, concl)))
+          (K (mk_mor_comp_tac mor_def mor_image'_thms morE_thms map_comp_id_thms))
+        |> Thm.close_derivation
+      end;
+
+    val mor_cong_thm =
+      let
+        val prems = map HOLogic.mk_Trueprop
+         (map2 (curry HOLogic.mk_eq) fs_copy fs @ [mk_mor Bs ss B's s's fs])
+        val concl = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs_copy);
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs @ fs_copy)
+            (Logic.list_implies (prems, concl)))
+          (K ((hyp_subst_tac lthy THEN' atac) 1))
+        |> Thm.close_derivation
+      end;
+
+    val mor_UNIV_thm =
+      let
+        fun mk_conjunct mapAsBs f s s' = HOLogic.mk_eq
+            (HOLogic.mk_comp (Term.list_comb (mapAsBs, passive_ids @ fs), s),
+            HOLogic.mk_comp (s', f));
+        val lhs = mk_mor active_UNIVs ss (map HOLogic.mk_UNIV activeBs) s's fs;
+        val rhs = Library.foldr1 HOLogic.mk_conj (map4 mk_conjunct mapsAsBs fs ss s's);
+      in
+        Goal.prove_sorry lthy [] [] (fold_rev Logic.all (ss @ s's @ fs) (mk_Trueprop_eq (lhs, rhs)))
+          (K (mk_mor_UNIV_tac morE_thms mor_def))
+        |> Thm.close_derivation
+      end;
+
+    val mor_str_thm =
+      let
+        val maps = map2 (fn Ds => fn bnf => Term.list_comb
+          (mk_map_of_bnf Ds allAs (passiveAs @ FTsAs) bnf, passive_ids @ ss)) Dss bnfs;
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all ss (HOLogic.mk_Trueprop
+            (mk_mor active_UNIVs ss (map HOLogic.mk_UNIV FTsAs) maps ss)))
+          (K (mk_mor_str_tac ks mor_UNIV_thm))
+        |> Thm.close_derivation
+      end;
+
+    val mor_sum_case_thm =
+      let
+        val maps = map3 (fn s => fn sum_s => fn mapx =>
+          mk_sum_case (HOLogic.mk_comp (Term.list_comb (mapx, passive_ids @ Inls), s), sum_s))
+          s's sum_ss map_Inls;
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (s's @ sum_ss) (HOLogic.mk_Trueprop
+            (mk_mor (map HOLogic.mk_UNIV activeBs) s's sum_UNIVs maps Inls)))
+          (K (mk_mor_sum_case_tac ks mor_UNIV_thm))
+        |> Thm.close_derivation
+      end;
+
+    val timer = time (timer "Morphism definition & thms");
+
+    fun hset_rec_bind j = mk_internal_b (hset_recN ^ (if m = 1 then "" else string_of_int j));
+    val hset_rec_name = Binding.name_of o hset_rec_bind;
+    val hset_rec_def_bind = rpair [] o Thm.def_binding o hset_rec_bind;
+
+    fun hset_rec_spec j Zero hsetT hrec hrec' =
+      let
+        fun mk_Suc s setsAs z z' =
+          let
+            val (set, sets) = apfst (fn xs => nth xs (j - 1)) (chop m setsAs);
+            fun mk_UN set k = mk_UNION (set $ (s $ z)) (mk_nthN n hrec k);
+          in
+            Term.absfree z'
+              (mk_union (set $ (s $ z), Library.foldl1 mk_union (map2 mk_UN sets ks)))
+          end;
+
+        val Suc = Term.absdummy HOLogic.natT (Term.absfree hrec'
+          (HOLogic.mk_tuple (map4 mk_Suc ss setssAs zs zs')));
+
+        val lhs = Term.list_comb (Free (hset_rec_name j, hsetT), ss);
+        val rhs = mk_nat_rec Zero Suc;
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((hset_rec_frees, (_, hset_rec_def_frees)), (lthy, lthy_old)) =
+      lthy
+      |> fold_map5 (fn j => fn Zero => fn hsetT => fn hrec => fn hrec' => Specification.definition
+        (SOME (hset_rec_bind j, NONE, NoSyn),
+          (hset_rec_def_bind j, hset_rec_spec j Zero hsetT hrec hrec')))
+        ls Zeros hsetTs hrecs hrecs'
+      |>> apsnd split_list o split_list
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+
+    val hset_rec_defs = map (Morphism.thm phi) hset_rec_def_frees;
+    val hset_recs = map (fst o Term.dest_Const o Morphism.term phi) hset_rec_frees;
+
+    fun mk_hset_rec ss nat i j T =
+      let
+        val args = ss @ [nat];
+        val Ts = map fastype_of ss;
+        val bTs = map domain_type Ts;
+        val hrecT = HOLogic.mk_tupleT (map (fn U => U --> HOLogic.mk_setT T) bTs)
+        val hset_recT = Library.foldr (op -->) (Ts, HOLogic.natT --> hrecT);
+      in
+        mk_nthN n (Term.list_comb (Const (nth hset_recs (j - 1), hset_recT), args)) i
+      end;
+
+    val hset_rec_0ss = mk_rec_simps n @{thm nat_rec_0} hset_rec_defs;
+    val hset_rec_Sucss = mk_rec_simps n @{thm nat_rec_Suc} hset_rec_defs;
+    val hset_rec_0ss' = transpose hset_rec_0ss;
+    val hset_rec_Sucss' = transpose hset_rec_Sucss;
+
+    fun hset_binds j = mk_internal_bs (hsetN ^ (if m = 1 then "" else string_of_int j))
+    fun hset_bind i j = nth (hset_binds j) (i - 1);
+    val hset_name = Binding.name_of oo hset_bind;
+    val hset_def_bind = rpair [] o Thm.def_binding oo hset_bind;
+
+    fun hset_spec i j =
+      let
+        val U = nth activeAs (i - 1);
+        val z = nth zs (i - 1);
+        val T = nth passiveAs (j - 1);
+        val setT = HOLogic.mk_setT T;
+        val hsetT = Library.foldr (op -->) (sTs, U --> setT);
+
+        val lhs = Term.list_comb (Free (hset_name i j, hsetT), ss @ [z]);
+        val rhs = mk_UNION (HOLogic.mk_UNIV HOLogic.natT)
+          (Term.absfree nat' (mk_hset_rec ss nat i j T $ z));
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((hset_frees, (_, hset_def_frees)), (lthy, lthy_old)) =
+      lthy
+      |> fold_map (fn i => fold_map (fn j => Specification.definition
+        (SOME (hset_bind i j, NONE, NoSyn), (hset_def_bind i j, hset_spec i j))) ls) ks
+      |>> map (apsnd split_list o split_list)
+      |>> apsnd split_list o split_list
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+
+    val hset_defss = map (map (Morphism.thm phi)) hset_def_frees;
+    val hset_defss' = transpose hset_defss;
+    val hset_namess = map (map (fst o Term.dest_Const o Morphism.term phi)) hset_frees;
+
+    fun mk_hset ss i j T =
+      let
+        val Ts = map fastype_of ss;
+        val bTs = map domain_type Ts;
+        val hsetT = Library.foldr (op -->) (Ts, nth bTs (i - 1) --> HOLogic.mk_setT T);
+      in
+        Term.list_comb (Const (nth (nth hset_namess (i - 1)) (j - 1), hsetT), ss)
+      end;
+
+    val hsetssAs = map (fn i => map2 (mk_hset ss i) ls passiveAs) ks;
+
+    val (set_incl_hset_thmss, set_hset_incl_hset_thmsss) =
+      let
+        fun mk_set_incl_hset s x set hset = fold_rev Logic.all (x :: ss)
+          (HOLogic.mk_Trueprop (mk_leq (set $ (s $ x)) (hset $ x)));
+
+        fun mk_set_hset_incl_hset s x y set hset1 hset2 =
+          fold_rev Logic.all (x :: y :: ss)
+            (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (x, set $ (s $ y))),
+            HOLogic.mk_Trueprop (mk_leq (hset1 $ x) (hset2 $ y))));
+
+        val set_incl_hset_goalss =
+          map4 (fn s => fn x => fn sets => fn hsets =>
+            map2 (mk_set_incl_hset s x) (take m sets) hsets)
+          ss zs setssAs hsetssAs;
+
+        (*xk : F(i)set(m+k) (si yi) ==> F(k)_hset(j) s1 ... sn xk <= F(i)_hset(j) s1 ... sn yi*)
+        val set_hset_incl_hset_goalsss =
+          map4 (fn si => fn yi => fn sets => fn hsetsi =>
+            map3 (fn xk => fn set => fn hsetsk =>
+              map2 (mk_set_hset_incl_hset si xk yi set) hsetsk hsetsi)
+            zs_copy (drop m sets) hsetssAs)
+          ss zs setssAs hsetssAs;
+      in
+        (map3 (fn goals => fn defs => fn rec_Sucs =>
+          map3 (fn goal => fn def => fn rec_Suc =>
+            Goal.prove_sorry lthy [] [] goal (K (mk_set_incl_hset_tac def rec_Suc))
+            |> Thm.close_derivation)
+          goals defs rec_Sucs)
+        set_incl_hset_goalss hset_defss hset_rec_Sucss,
+        map3 (fn goalss => fn defsi => fn rec_Sucs =>
+          map3 (fn k => fn goals => fn defsk =>
+            map4 (fn goal => fn defk => fn defi => fn rec_Suc =>
+              Goal.prove_sorry lthy [] [] goal
+                (K (mk_set_hset_incl_hset_tac n [defk, defi] rec_Suc k))
+              |> Thm.close_derivation)
+            goals defsk defsi rec_Sucs)
+          ks goalss hset_defss)
+        set_hset_incl_hset_goalsss hset_defss hset_rec_Sucss)
+      end;
+
+    val set_incl_hset_thmss' = transpose set_incl_hset_thmss;
+    val set_hset_incl_hset_thmsss' = transpose (map transpose set_hset_incl_hset_thmsss);
+    val set_hset_thmss = map (map (fn thm => thm RS @{thm set_mp})) set_incl_hset_thmss;
+    val set_hset_hset_thmsss = map (map (map (fn thm => thm RS @{thm set_mp})))
+      set_hset_incl_hset_thmsss;
+    val set_hset_thmss' = transpose set_hset_thmss;
+    val set_hset_hset_thmsss' = transpose (map transpose set_hset_hset_thmsss);
+
+    val hset_minimal_thms =
+      let
+        fun mk_passive_prem set s x K =
+          Logic.all x (HOLogic.mk_Trueprop (mk_leq (set $ (s $ x)) (K $ x)));
+
+        fun mk_active_prem s x1 K1 set x2 K2 =
+          fold_rev Logic.all [x1, x2]
+            (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (x2, set $ (s $ x1))),
+              HOLogic.mk_Trueprop (mk_leq (K2 $ x2) (K1 $ x1))));
+
+        val premss = map2 (fn j => fn Ks =>
+          map4 mk_passive_prem (map (fn xs => nth xs (j - 1)) setssAs) ss zs Ks @
+            flat (map4 (fn sets => fn s => fn x1 => fn K1 =>
+              map3 (mk_active_prem s x1 K1) (drop m sets) zs_copy Ks) setssAs ss zs Ks))
+          ls Kss;
+
+        val hset_rec_minimal_thms =
+          let
+            fun mk_conjunct j T i K x = mk_leq (mk_hset_rec ss nat i j T $ x) (K $ x);
+            fun mk_concl j T Ks = list_all_free zs
+              (Library.foldr1 HOLogic.mk_conj (map3 (mk_conjunct j T) ks Ks zs));
+            val concls = map3 mk_concl ls passiveAs Kss;
+
+            val goals = map2 (fn prems => fn concl =>
+              Logic.list_implies (prems, HOLogic.mk_Trueprop concl)) premss concls
+
+            val ctss =
+              map (fn phi => map (SOME o certify lthy) [Term.absfree nat' phi, nat]) concls;
+          in
+            map4 (fn goal => fn cts => fn hset_rec_0s => fn hset_rec_Sucs =>
+              singleton (Proof_Context.export names_lthy lthy)
+                (Goal.prove_sorry lthy [] [] goal
+                  (mk_hset_rec_minimal_tac m cts hset_rec_0s hset_rec_Sucs))
+              |> Thm.close_derivation)
+            goals ctss hset_rec_0ss' hset_rec_Sucss'
+          end;
+
+        fun mk_conjunct j T i K x = mk_leq (mk_hset ss i j T $ x) (K $ x);
+        fun mk_concl j T Ks = Library.foldr1 HOLogic.mk_conj (map3 (mk_conjunct j T) ks Ks zs);
+        val concls = map3 mk_concl ls passiveAs Kss;
+
+        val goals = map3 (fn Ks => fn prems => fn concl =>
+          fold_rev Logic.all (Ks @ ss @ zs)
+            (Logic.list_implies (prems, HOLogic.mk_Trueprop concl))) Kss premss concls;
+      in
+        map3 (fn goal => fn hset_defs => fn hset_rec_minimal =>
+          Goal.prove_sorry lthy [] [] goal
+            (mk_hset_minimal_tac n hset_defs hset_rec_minimal)
+          |> Thm.close_derivation)
+        goals hset_defss' hset_rec_minimal_thms
+      end;
+
+    val timer = time (timer "Hereditary sets");
+
+    (* bisimulation *)
+
+    val bis_bind = mk_internal_b bisN;
+    val bis_name = Binding.name_of bis_bind;
+    val bis_def_bind = (Thm.def_binding bis_bind, []);
+
+    fun mk_bis_le_conjunct R B1 B2 = mk_leq R (mk_Times (B1, B2));
+    val bis_le = Library.foldr1 HOLogic.mk_conj (map3 mk_bis_le_conjunct Rs Bs B's)
+
+    val bis_spec =
+      let
+        val bisT = Library.foldr (op -->) (ATs @ BTs @ sTs @ B'Ts @ s'Ts @ setRTs, HOLogic.boolT);
+
+        val fst_args = passive_ids @ fsts;
+        val snd_args = passive_ids @ snds;
+        fun mk_bis R s s' b1 b2 RF map1 map2 sets =
+          list_all_free [b1, b2] (HOLogic.mk_imp
+            (HOLogic.mk_mem (HOLogic.mk_prod (b1, b2), R),
+            mk_Bex (mk_in (As @ Rs) sets (snd (dest_Free RF))) (Term.absfree (dest_Free RF)
+              (HOLogic.mk_conj
+                (HOLogic.mk_eq (Term.list_comb (map1, fst_args) $ RF, s $ b1),
+                HOLogic.mk_eq (Term.list_comb (map2, snd_args) $ RF, s' $ b2))))));
+
+        val lhs = Term.list_comb (Free (bis_name, bisT), As @ Bs @ ss @ B's @ s's @ Rs);
+        val rhs = HOLogic.mk_conj
+          (bis_le, Library.foldr1 HOLogic.mk_conj
+            (map9 mk_bis Rs ss s's zs z's RFs map_fsts map_snds bis_setss))
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((bis_free, (_, bis_def_free)), (lthy, lthy_old)) =
+      lthy
+      |> Specification.definition (SOME (bis_bind, NONE, NoSyn), (bis_def_bind, bis_spec))
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+    val bis = fst (Term.dest_Const (Morphism.term phi bis_free));
+    val bis_def = Morphism.thm phi bis_def_free;
+
+    fun mk_bis As Bs1 ss1 Bs2 ss2 Rs =
+      let
+        val args = As @ Bs1 @ ss1 @ Bs2 @ ss2 @ Rs;
+        val Ts = map fastype_of args;
+        val bisT = Library.foldr (op -->) (Ts, HOLogic.boolT);
+      in
+        Term.list_comb (Const (bis, bisT), args)
+      end;
+
+    val bis_cong_thm =
+      let
+        val prems = map HOLogic.mk_Trueprop
+         (mk_bis As Bs ss B's s's Rs :: map2 (curry HOLogic.mk_eq) Rs_copy Rs)
+        val concl = HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's Rs_copy);
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ Rs @ Rs_copy)
+            (Logic.list_implies (prems, concl)))
+          (K ((hyp_subst_tac lthy THEN' atac) 1))
+        |> Thm.close_derivation
+      end;
+
+    val bis_rel_thm =
+      let
+        fun mk_conjunct R s s' b1 b2 rel =
+          list_all_free [b1, b2] (HOLogic.mk_imp
+            (HOLogic.mk_mem (HOLogic.mk_prod (b1, b2), R),
+            Term.list_comb (rel, map mk_in_rel (passive_Id_ons @ Rs)) $ (s $ b1) $ (s' $ b2)));
+
+        val rhs = HOLogic.mk_conj
+          (bis_le, Library.foldr1 HOLogic.mk_conj
+            (map6 mk_conjunct Rs ss s's zs z's relsAsBs))
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ Rs)
+            (mk_Trueprop_eq (mk_bis As Bs ss B's s's Rs, rhs)))
+          (K (mk_bis_rel_tac lthy m bis_def rel_OO_Grps map_comps map_cong0s set_mapss))
+        |> Thm.close_derivation
+      end;
+
+    val bis_converse_thm =
+      Goal.prove_sorry lthy [] []
+        (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ Rs)
+          (Logic.mk_implies
+            (HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's Rs),
+            HOLogic.mk_Trueprop (mk_bis As B's s's Bs ss (map mk_converse Rs)))))
+        (K (mk_bis_converse_tac m bis_rel_thm rel_congs rel_converseps))
+      |> Thm.close_derivation;
+
+    val bis_O_thm =
+      let
+        val prems =
+          [HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's Rs),
+           HOLogic.mk_Trueprop (mk_bis As B's s's B''s s''s R's)];
+        val concl =
+          HOLogic.mk_Trueprop (mk_bis As Bs ss B''s s''s (map2 (curry mk_rel_comp) Rs R's));
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ B''s @ s''s @ Rs @ R's)
+            (Logic.list_implies (prems, concl)))
+          (K (mk_bis_O_tac lthy m bis_rel_thm rel_congs rel_OOs))
+        |> Thm.close_derivation
+      end;
+
+    val bis_Gr_thm =
+      let
+        val concl =
+          HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's (map2 mk_Gr Bs fs));
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ fs)
+            (Logic.list_implies ([coalg_prem, mor_prem], concl)))
+          (mk_bis_Gr_tac bis_rel_thm rel_Grps mor_image_thms morE_thms coalg_in_thms)
+        |> Thm.close_derivation
+      end;
+
+    val bis_image2_thm = bis_cong_thm OF
+      ((bis_O_thm OF [bis_Gr_thm RS bis_converse_thm, bis_Gr_thm]) ::
+      replicate n @{thm image2_Gr});
+
+    val bis_Id_on_thm = bis_cong_thm OF ((mor_id_thm RSN (2, bis_Gr_thm)) ::
+      replicate n @{thm Id_on_Gr});
+
+    val bis_Union_thm =
+      let
+        val prem =
+          HOLogic.mk_Trueprop (mk_Ball Idx
+            (Term.absfree idx' (mk_bis As Bs ss B's s's (map (fn R => R $ idx) Ris))));
+        val concl =
+          HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's (map (mk_UNION Idx) Ris));
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (Idx :: As @ Bs @ ss @ B's @ s's @ Ris)
+            (Logic.mk_implies (prem, concl)))
+          (mk_bis_Union_tac bis_def in_mono'_thms)
+        |> Thm.close_derivation
+      end;
+
+    (* self-bisimulation *)
+
+    fun mk_sbis As Bs ss Rs = mk_bis As Bs ss Bs ss Rs;
+
+    val sbis_prem = HOLogic.mk_Trueprop (mk_sbis As Bs ss sRs);
+
+    (* largest self-bisimulation *)
+
+    val lsbis_binds = mk_internal_bs lsbisN;
+    fun lsbis_bind i = nth lsbis_binds (i - 1);
+    val lsbis_name = Binding.name_of o lsbis_bind;
+    val lsbis_def_bind = rpair [] o Thm.def_binding o lsbis_bind;
+
+    val all_sbis = HOLogic.mk_Collect (fst Rtuple', snd Rtuple', list_exists_free sRs
+      (HOLogic.mk_conj (HOLogic.mk_eq (Rtuple, HOLogic.mk_tuple sRs), mk_sbis As Bs ss sRs)));
+
+    fun lsbis_spec i RT =
+      let
+        fun mk_lsbisT RT =
+          Library.foldr (op -->) (map fastype_of (As @ Bs @ ss), RT);
+        val lhs = Term.list_comb (Free (lsbis_name i, mk_lsbisT RT), As @ Bs @ ss);
+        val rhs = mk_UNION all_sbis (Term.absfree Rtuple' (mk_nthN n Rtuple i));
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((lsbis_frees, (_, lsbis_def_frees)), (lthy, lthy_old)) =
+      lthy
+      |> fold_map2 (fn i => fn RT => Specification.definition
+        (SOME (lsbis_bind i, NONE, NoSyn), (lsbis_def_bind i, lsbis_spec i RT))) ks setsRTs
+      |>> apsnd split_list o split_list
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+
+    val lsbis_defs = map (Morphism.thm phi) lsbis_def_frees;
+    val lsbiss = map (fst o Term.dest_Const o Morphism.term phi) lsbis_frees;
+
+    fun mk_lsbis As Bs ss i =
+      let
+        val args = As @ Bs @ ss;
+        val Ts = map fastype_of args;
+        val RT = mk_relT (`I (HOLogic.dest_setT (fastype_of (nth Bs (i - 1)))));
+        val lsbisT = Library.foldr (op -->) (Ts, RT);
+      in
+        Term.list_comb (Const (nth lsbiss (i - 1), lsbisT), args)
+      end;
+
+    val sbis_lsbis_thm =
+      Goal.prove_sorry lthy [] []
+        (fold_rev Logic.all (As @ Bs @ ss)
+          (HOLogic.mk_Trueprop (mk_sbis As Bs ss (map (mk_lsbis As Bs ss) ks))))
+        (K (mk_sbis_lsbis_tac lthy lsbis_defs bis_Union_thm bis_cong_thm))
+      |> Thm.close_derivation;
+
+    val lsbis_incl_thms = map (fn i => sbis_lsbis_thm RS
+      (bis_def RS iffD1 RS conjunct1 RS mk_conjunctN n i)) ks;
+    val lsbisE_thms = map (fn i => (mk_specN 2 (sbis_lsbis_thm RS
+      (bis_def RS iffD1 RS conjunct2 RS mk_conjunctN n i))) RS mp) ks;
+
+    val incl_lsbis_thms =
+      let
+        fun mk_concl i R = HOLogic.mk_Trueprop (mk_leq R (mk_lsbis As Bs ss i));
+        val goals = map2 (fn i => fn R => fold_rev Logic.all (As @ Bs @ ss @ sRs)
+          (Logic.mk_implies (sbis_prem, mk_concl i R))) ks sRs;
+      in
+        map3 (fn goal => fn i => fn def => Goal.prove_sorry lthy [] [] goal
+          (K (mk_incl_lsbis_tac n i def)) |> Thm.close_derivation) goals ks lsbis_defs
+      end;
+
+    val equiv_lsbis_thms =
+      let
+        fun mk_concl i B = HOLogic.mk_Trueprop (mk_equiv B (mk_lsbis As Bs ss i));
+        val goals = map2 (fn i => fn B => fold_rev Logic.all (As @ Bs @ ss)
+          (Logic.mk_implies (coalg_prem, mk_concl i B))) ks Bs;
+      in
+        map3 (fn goal => fn l_incl => fn incl_l =>
+          Goal.prove_sorry lthy [] [] goal
+            (K (mk_equiv_lsbis_tac sbis_lsbis_thm l_incl incl_l
+              bis_Id_on_thm bis_converse_thm bis_O_thm))
+          |> Thm.close_derivation)
+        goals lsbis_incl_thms incl_lsbis_thms
+      end;
+
+    val timer = time (timer "Bisimulations");
+
+    (* bounds *)
+
+    val (lthy, sbd, sbdT,
+      sbd_card_order, sbd_Cinfinite, sbd_Card_order, set_sbdss) =
+      if n = 1
+      then (lthy, sum_bd, sum_bdT, bd_card_order, bd_Cinfinite, bd_Card_order, set_bdss)
+      else
+        let
+          val sbdT_bind = mk_internal_b sum_bdTN;
+
+          val ((sbdT_name, (sbdT_glob_info, sbdT_loc_info)), lthy) =
+            typedef (sbdT_bind, dead_params, NoSyn)
+              (HOLogic.mk_UNIV sum_bdT) NONE (EVERY' [rtac exI, rtac UNIV_I] 1) lthy;
+
+          val sbdT = Type (sbdT_name, dead_params');
+          val Abs_sbdT = Const (#Abs_name sbdT_glob_info, sum_bdT --> sbdT);
+
+          val sbd_bind = mk_internal_b sum_bdN;
+          val sbd_name = Binding.name_of sbd_bind;
+          val sbd_def_bind = (Thm.def_binding sbd_bind, []);
+
+          val sbd_spec = HOLogic.mk_Trueprop
+            (HOLogic.mk_eq (Free (sbd_name, mk_relT (`I sbdT)), mk_dir_image sum_bd Abs_sbdT));
+
+          val ((sbd_free, (_, sbd_def_free)), (lthy, lthy_old)) =
+            lthy
+            |> Specification.definition (SOME (sbd_bind, NONE, NoSyn), (sbd_def_bind, sbd_spec))
+            ||> `Local_Theory.restore;
+
+          val phi = Proof_Context.export_morphism lthy_old lthy;
+
+          val sbd_def = Morphism.thm phi sbd_def_free;
+          val sbd = Const (fst (Term.dest_Const (Morphism.term phi sbd_free)), mk_relT (`I sbdT));
+
+          val Abs_sbdT_inj = mk_Abs_inj_thm (#Abs_inject sbdT_loc_info);
+          val Abs_sbdT_bij = mk_Abs_bij_thm lthy Abs_sbdT_inj (#Abs_cases sbdT_loc_info);
+
+          fun mk_sum_Cinfinite [thm] = thm
+            | mk_sum_Cinfinite (thm :: thms) =
+              @{thm Cinfinite_csum_strong} OF [thm, mk_sum_Cinfinite thms];
+
+          val sum_Cinfinite = mk_sum_Cinfinite bd_Cinfinites;
+          val sum_Card_order = sum_Cinfinite RS conjunct2;
+
+          fun mk_sum_card_order [thm] = thm
+            | mk_sum_card_order (thm :: thms) =
+              @{thm card_order_csum} OF [thm, mk_sum_card_order thms];
+
+          val sum_card_order = mk_sum_card_order bd_card_orders;
+
+          val sbd_ordIso = fold_thms lthy [sbd_def]
+            (@{thm dir_image} OF [Abs_sbdT_inj, sum_Card_order]);
+          val sbd_card_order =  fold_thms lthy [sbd_def]
+            (@{thm card_order_dir_image} OF [Abs_sbdT_bij, sum_card_order]);
+          val sbd_Cinfinite = @{thm Cinfinite_cong} OF [sbd_ordIso, sum_Cinfinite];
+          val sbd_Card_order = sbd_Cinfinite RS conjunct2;
+
+          fun mk_set_sbd i bd_Card_order bds =
+            map (fn thm => @{thm ordLeq_ordIso_trans} OF
+              [bd_Card_order RS mk_ordLeq_csum n i thm, sbd_ordIso]) bds;
+          val set_sbdss = map3 mk_set_sbd ks bd_Card_orders set_bdss;
+       in
+         (lthy, sbd, sbdT, sbd_card_order, sbd_Cinfinite, sbd_Card_order, set_sbdss)
+       end;
+
+    val sbdTs = replicate n sbdT;
+    val sum_sbd = Library.foldr1 (uncurry mk_csum) (replicate n sbd);
+    val sum_sbdT = mk_sumTN sbdTs;
+    val sum_sbd_listT = HOLogic.listT sum_sbdT;
+    val sum_sbd_list_setT = HOLogic.mk_setT sum_sbd_listT;
+    val bdTs = passiveAs @ replicate n sbdT;
+    val to_sbd_maps = map4 mk_map_of_bnf Dss Ass (replicate n bdTs) bnfs;
+    val bdFTs = mk_FTs bdTs;
+    val sbdFT = mk_sumTN bdFTs;
+    val treeT = HOLogic.mk_prodT (sum_sbd_list_setT, sum_sbd_listT --> sbdFT);
+    val treeQT = HOLogic.mk_setT treeT;
+    val treeTs = passiveAs @ replicate n treeT;
+    val treeQTs = passiveAs @ replicate n treeQT;
+    val treeFTs = mk_FTs treeTs;
+    val tree_maps = map4 mk_map_of_bnf Dss (replicate n bdTs) (replicate n treeTs) bnfs;
+    val final_maps = map4 mk_map_of_bnf Dss (replicate n treeTs) (replicate n treeQTs) bnfs;
+    val isNode_setss = mk_setss (passiveAs @ replicate n sbdT);
+
+    val root = HOLogic.mk_set sum_sbd_listT [HOLogic.mk_list sum_sbdT []];
+    val Zero = HOLogic.mk_tuple (map (fn U => absdummy U root) activeAs);
+    val Lev_recT = fastype_of Zero;
+    val LevT = Library.foldr (op -->) (sTs, HOLogic.natT --> Lev_recT);
+
+    val Nil = HOLogic.mk_tuple (map3 (fn i => fn z => fn z'=>
+      Term.absfree z' (mk_InN activeAs z i)) ks zs zs');
+    val rv_recT = fastype_of Nil;
+    val rvT = Library.foldr (op -->) (sTs, sum_sbd_listT --> rv_recT);
+
+    val (((((((((((sumx, sumx'), (kks, kks')), (kl, kl')), (kl_copy, kl'_copy)), (Kl, Kl')),
+      (lab, lab')), (Kl_lab, Kl_lab')), xs), (Lev_rec, Lev_rec')), (rv_rec, rv_rec')),
+      names_lthy) = names_lthy
+      |> yield_singleton (apfst (op ~~) oo mk_Frees' "sumx") sum_sbdT
+      ||>> mk_Frees' "k" sbdTs
+      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "kl") sum_sbd_listT
+      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "kl") sum_sbd_listT
+      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "Kl") sum_sbd_list_setT
+      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "lab") (sum_sbd_listT --> sbdFT)
+      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "Kl_lab") treeT
+      ||>> mk_Frees "x" bdFTs
+      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "rec") Lev_recT
+      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "rec") rv_recT;
+
+    val (k, k') = (hd kks, hd kks')
+
+    val timer = time (timer "Bounds");
+
+    (* tree coalgebra *)
+
+    val isNode_binds = mk_internal_bs isNodeN;
+    fun isNode_bind i = nth isNode_binds (i - 1);
+    val isNode_name = Binding.name_of o isNode_bind;
+    val isNode_def_bind = rpair [] o Thm.def_binding o isNode_bind;
+
+    val isNodeT =
+      Library.foldr (op -->) (map fastype_of (As @ [Kl, lab, kl]), HOLogic.boolT);
+
+    val Succs = map3 (fn i => fn k => fn k' =>
+      HOLogic.mk_Collect (fst k', snd k', HOLogic.mk_mem (mk_InN sbdTs k i, mk_Succ Kl kl)))
+      ks kks kks';
+
+    fun isNode_spec sets x i =
+      let
+        val (passive_sets, active_sets) = chop m (map (fn set => set $ x) sets);
+        val lhs = Term.list_comb (Free (isNode_name i, isNodeT), As @ [Kl, lab, kl]);
+        val rhs = list_exists_free [x]
+          (Library.foldr1 HOLogic.mk_conj (HOLogic.mk_eq (lab $ kl, mk_InN bdFTs x i) ::
+          map2 mk_leq passive_sets As @ map2 (curry HOLogic.mk_eq) active_sets Succs));
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((isNode_frees, (_, isNode_def_frees)), (lthy, lthy_old)) =
+      lthy
+      |> fold_map3 (fn i => fn x => fn sets => Specification.definition
+        (SOME (isNode_bind i, NONE, NoSyn), (isNode_def_bind i, isNode_spec sets x i)))
+        ks xs isNode_setss
+      |>> apsnd split_list o split_list
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+
+    val isNode_defs = map (Morphism.thm phi) isNode_def_frees;
+    val isNodes = map (fst o Term.dest_Const o Morphism.term phi) isNode_frees;
+
+    fun mk_isNode As kl i =
+      Term.list_comb (Const (nth isNodes (i - 1), isNodeT), As @ [Kl, lab, kl]);
+
+    val isTree =
+      let
+        val empty = HOLogic.mk_mem (HOLogic.mk_list sum_sbdT [], Kl);
+        val Field = mk_leq Kl (mk_Field (mk_clists sum_sbd));
+        val prefCl = mk_prefCl Kl;
+
+        val tree = mk_Ball Kl (Term.absfree kl'
+          (HOLogic.mk_conj
+            (Library.foldr1 HOLogic.mk_disj (map (mk_isNode As kl) ks),
+            Library.foldr1 HOLogic.mk_conj (map4 (fn Succ => fn i => fn k => fn k' =>
+              mk_Ball Succ (Term.absfree k' (mk_isNode As
+                (mk_append (kl, HOLogic.mk_list sum_sbdT [mk_InN sbdTs k i])) i)))
+            Succs ks kks kks'))));
+
+        val undef = list_all_free [kl] (HOLogic.mk_imp
+          (HOLogic.mk_not (HOLogic.mk_mem (kl, Kl)),
+          HOLogic.mk_eq (lab $ kl, mk_undefined sbdFT)));
+      in
+        Library.foldr1 HOLogic.mk_conj [empty, Field, prefCl, tree, undef]
+      end;
+
+    val carT_binds = mk_internal_bs carTN;
+    fun carT_bind i = nth carT_binds (i - 1);
+    val carT_name = Binding.name_of o carT_bind;
+    val carT_def_bind = rpair [] o Thm.def_binding o carT_bind;
+
+    fun carT_spec i =
+      let
+        val carTT = Library.foldr (op -->) (ATs, HOLogic.mk_setT treeT);
+
+        val lhs = Term.list_comb (Free (carT_name i, carTT), As);
+        val rhs = HOLogic.mk_Collect (fst Kl_lab', snd Kl_lab', list_exists_free [Kl, lab]
+          (HOLogic.mk_conj (HOLogic.mk_eq (Kl_lab, HOLogic.mk_prod (Kl, lab)),
+            HOLogic.mk_conj (isTree, mk_isNode As (HOLogic.mk_list sum_sbdT []) i))));
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((carT_frees, (_, carT_def_frees)), (lthy, lthy_old)) =
+      lthy
+      |> fold_map (fn i => Specification.definition
+        (SOME (carT_bind i, NONE, NoSyn), (carT_def_bind i, carT_spec i))) ks
+      |>> apsnd split_list o split_list
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+
+    val carT_defs = map (Morphism.thm phi) carT_def_frees;
+    val carTs = map (fst o Term.dest_Const o Morphism.term phi) carT_frees;
+
+    fun mk_carT As i = Term.list_comb
+      (Const (nth carTs (i - 1),
+         Library.foldr (op -->) (map fastype_of As, HOLogic.mk_setT treeT)), As);
+
+    val strT_binds = mk_internal_bs strTN;
+    fun strT_bind i = nth strT_binds (i - 1);
+    val strT_name = Binding.name_of o strT_bind;
+    val strT_def_bind = rpair [] o Thm.def_binding o strT_bind;
+
+    fun strT_spec mapFT FT i =
+      let
+        val strTT = treeT --> FT;
+
+        fun mk_f i k k' =
+          let val in_k = mk_InN sbdTs k i;
+          in Term.absfree k' (HOLogic.mk_prod (mk_Shift Kl in_k, mk_shift lab in_k)) end;
+
+        val f = Term.list_comb (mapFT, passive_ids @ map3 mk_f ks kks kks');
+        val (fTs1, fTs2) = apsnd tl (chop (i - 1) (map (fn T => T --> FT) bdFTs));
+        val fs = map mk_undefined fTs1 @ (f :: map mk_undefined fTs2);
+        val lhs = Free (strT_name i, strTT);
+        val rhs = HOLogic.mk_split (Term.absfree Kl' (Term.absfree lab'
+          (mk_sum_caseN fs $ (lab $ HOLogic.mk_list sum_sbdT []))));
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((strT_frees, (_, strT_def_frees)), (lthy, lthy_old)) =
+      lthy
+      |> fold_map3 (fn i => fn mapFT => fn FT => Specification.definition
+        (SOME (strT_bind i, NONE, NoSyn), (strT_def_bind i, strT_spec mapFT FT i)))
+        ks tree_maps treeFTs
+      |>> apsnd split_list o split_list
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+
+    val strT_defs = map ((fn def => trans OF [def RS fun_cong, @{thm prod.cases}]) o
+      Morphism.thm phi) strT_def_frees;
+    val strTs = map (fst o Term.dest_Const o Morphism.term phi) strT_frees;
+
+    fun mk_strT FT i = Const (nth strTs (i - 1), treeT --> FT);
+
+    val carTAs = map (mk_carT As) ks;
+    val strTAs = map2 mk_strT treeFTs ks;
+
+    val coalgT_thm =
+      Goal.prove_sorry lthy [] []
+        (fold_rev Logic.all As (HOLogic.mk_Trueprop (mk_coalg As carTAs strTAs)))
+        (mk_coalgT_tac m (coalg_def :: isNode_defs @ carT_defs) strT_defs set_mapss)
+      |> Thm.close_derivation;
+
+    val timer = time (timer "Tree coalgebra");
+
+    fun mk_to_sbd s x i i' =
+      mk_toCard (nth (nth setssAs (i - 1)) (m + i' - 1) $ (s $ x)) sbd;
+    fun mk_from_sbd s x i i' =
+      mk_fromCard (nth (nth setssAs (i - 1)) (m + i' - 1) $ (s $ x)) sbd;
+
+    fun mk_to_sbd_thmss thm = map (map (fn set_sbd =>
+      thm OF [set_sbd, sbd_Card_order]) o drop m) set_sbdss;
+
+    val to_sbd_inj_thmss = mk_to_sbd_thmss @{thm toCard_inj};
+    val to_sbd_thmss = mk_to_sbd_thmss @{thm toCard};
+    val from_to_sbd_thmss = mk_to_sbd_thmss @{thm fromCard_toCard};
+
+    val Lev_bind = mk_internal_b LevN;
+    val Lev_name = Binding.name_of Lev_bind;
+    val Lev_def_bind = rpair [] (Thm.def_binding Lev_bind);
+
+    val Lev_spec =
+      let
+        fun mk_Suc i s setsAs a a' =
+          let
+            val sets = drop m setsAs;
+            fun mk_set i' set b =
+              let
+                val Cons = HOLogic.mk_eq (kl_copy,
+                  mk_Cons (mk_InN sbdTs (mk_to_sbd s a i i' $ b) i') kl)
+                val b_set = HOLogic.mk_mem (b, set $ (s $ a));
+                val kl_rec = HOLogic.mk_mem (kl, mk_nthN n Lev_rec i' $ b);
+              in
+                HOLogic.mk_Collect (fst kl'_copy, snd kl'_copy, list_exists_free [b, kl]
+                  (HOLogic.mk_conj (Cons, HOLogic.mk_conj (b_set, kl_rec))))
+              end;
+          in
+            Term.absfree a' (Library.foldl1 mk_union (map3 mk_set ks sets zs_copy))
+          end;
+
+        val Suc = Term.absdummy HOLogic.natT (Term.absfree Lev_rec'
+          (HOLogic.mk_tuple (map5 mk_Suc ks ss setssAs zs zs')));
+
+        val lhs = Term.list_comb (Free (Lev_name, LevT), ss);
+        val rhs = mk_nat_rec Zero Suc;
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((Lev_free, (_, Lev_def_free)), (lthy, lthy_old)) =
+      lthy
+      |> Specification.definition (SOME (Lev_bind, NONE, NoSyn), (Lev_def_bind, Lev_spec))
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+
+    val Lev_def = Morphism.thm phi Lev_def_free;
+    val Lev = fst (Term.dest_Const (Morphism.term phi Lev_free));
+
+    fun mk_Lev ss nat i =
+      let
+        val Ts = map fastype_of ss;
+        val LevT = Library.foldr (op -->) (Ts, HOLogic.natT -->
+          HOLogic.mk_tupleT (map (fn U => domain_type U --> sum_sbd_list_setT) Ts));
+      in
+        mk_nthN n (Term.list_comb (Const (Lev, LevT), ss) $ nat) i
+      end;
+
+    val Lev_0s = flat (mk_rec_simps n @{thm nat_rec_0} [Lev_def]);
+    val Lev_Sucs = flat (mk_rec_simps n @{thm nat_rec_Suc} [Lev_def]);
+
+    val rv_bind = mk_internal_b rvN;
+    val rv_name = Binding.name_of rv_bind;
+    val rv_def_bind = rpair [] (Thm.def_binding rv_bind);
+
+    val rv_spec =
+      let
+        fun mk_Cons i s b b' =
+          let
+            fun mk_case i' =
+              Term.absfree k' (mk_nthN n rv_rec i' $ (mk_from_sbd s b i i' $ k));
+          in
+            Term.absfree b' (mk_sum_caseN (map mk_case ks) $ sumx)
+          end;
+
+        val Cons = Term.absfree sumx' (Term.absdummy sum_sbd_listT (Term.absfree rv_rec'
+          (HOLogic.mk_tuple (map4 mk_Cons ks ss zs zs'))));
+
+        val lhs = Term.list_comb (Free (rv_name, rvT), ss);
+        val rhs = mk_list_rec Nil Cons;
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((rv_free, (_, rv_def_free)), (lthy, lthy_old)) =
+      lthy
+      |> Specification.definition (SOME (rv_bind, NONE, NoSyn), (rv_def_bind, rv_spec))
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+
+    val rv_def = Morphism.thm phi rv_def_free;
+    val rv = fst (Term.dest_Const (Morphism.term phi rv_free));
+
+    fun mk_rv ss kl i =
+      let
+        val Ts = map fastype_of ss;
+        val As = map domain_type Ts;
+        val rvT = Library.foldr (op -->) (Ts, fastype_of kl -->
+          HOLogic.mk_tupleT (map (fn U => U --> mk_sumTN As) As));
+      in
+        mk_nthN n (Term.list_comb (Const (rv, rvT), ss) $ kl) i
+      end;
+
+    val rv_Nils = flat (mk_rec_simps n @{thm list_rec_Nil} [rv_def]);
+    val rv_Conss = flat (mk_rec_simps n @{thm list_rec_Cons} [rv_def]);
+
+    val beh_binds = mk_internal_bs behN;
+    fun beh_bind i = nth beh_binds (i - 1);
+    val beh_name = Binding.name_of o beh_bind;
+    val beh_def_bind = rpair [] o Thm.def_binding o beh_bind;
+
+    fun beh_spec i z =
+      let
+        val mk_behT = Library.foldr (op -->) (map fastype_of (ss @ [z]), treeT);
+
+        fun mk_case i to_sbd_map s k k' =
+          Term.absfree k' (mk_InN bdFTs
+            (Term.list_comb (to_sbd_map, passive_ids @ map (mk_to_sbd s k i) ks) $ (s $ k)) i);
+
+        val Lab = Term.absfree kl' (mk_If
+          (HOLogic.mk_mem (kl, mk_Lev ss (mk_size kl) i $ z))
+          (mk_sum_caseN (map5 mk_case ks to_sbd_maps ss zs zs') $ (mk_rv ss kl i $ z))
+          (mk_undefined sbdFT));
+
+        val lhs = Term.list_comb (Free (beh_name i, mk_behT), ss) $ z;
+        val rhs = HOLogic.mk_prod (mk_UNION (HOLogic.mk_UNIV HOLogic.natT)
+          (Term.absfree nat' (mk_Lev ss nat i $ z)), Lab);
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((beh_frees, (_, beh_def_frees)), (lthy, lthy_old)) =
+      lthy
+      |> fold_map2 (fn i => fn z => Specification.definition
+        (SOME (beh_bind i, NONE, NoSyn), (beh_def_bind i, beh_spec i z))) ks zs
+      |>> apsnd split_list o split_list
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+
+    val beh_defs = map (Morphism.thm phi) beh_def_frees;
+    val behs = map (fst o Term.dest_Const o Morphism.term phi) beh_frees;
+
+    fun mk_beh ss i =
+      let
+        val Ts = map fastype_of ss;
+        val behT = Library.foldr (op -->) (Ts, nth activeAs (i - 1) --> treeT);
+      in
+        Term.list_comb (Const (nth behs (i - 1), behT), ss)
+      end;
+
+    val Lev_sbd_thms =
+      let
+        fun mk_conjunct i z = mk_leq (mk_Lev ss nat i $ z) (mk_Field (mk_clists sum_sbd));
+        val goal = list_all_free zs
+          (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs));
+
+        val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
+
+        val Lev_sbd = singleton (Proof_Context.export names_lthy lthy)
+          (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
+            (K (mk_Lev_sbd_tac lthy cts Lev_0s Lev_Sucs to_sbd_thmss))
+          |> Thm.close_derivation);
+
+        val Lev_sbd' = mk_specN n Lev_sbd;
+      in
+        map (fn i => Lev_sbd' RS mk_conjunctN n i) ks
+      end;
+
+    val (length_Lev_thms, length_Lev'_thms) =
+      let
+        fun mk_conjunct i z = HOLogic.mk_imp (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z),
+          HOLogic.mk_eq (mk_size kl, nat));
+        val goal = list_all_free (kl :: zs)
+          (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs));
+
+        val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
+
+        val length_Lev = singleton (Proof_Context.export names_lthy lthy)
+          (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
+            (K (mk_length_Lev_tac lthy cts Lev_0s Lev_Sucs))
+          |> Thm.close_derivation);
+
+        val length_Lev' = mk_specN (n + 1) length_Lev;
+        val length_Levs = map (fn i => length_Lev' RS mk_conjunctN n i RS mp) ks;
+
+        fun mk_goal i z = fold_rev Logic.all (z :: kl :: nat :: ss) (Logic.mk_implies
+            (HOLogic.mk_Trueprop (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z)),
+            HOLogic.mk_Trueprop (HOLogic.mk_mem (kl, mk_Lev ss (mk_size kl) i $ z))));
+        val goals = map2 mk_goal ks zs;
+
+        val length_Levs' = map2 (fn goal => fn length_Lev =>
+          Goal.prove_sorry lthy [] [] goal (K (mk_length_Lev'_tac length_Lev))
+          |> Thm.close_derivation) goals length_Levs;
+      in
+        (length_Levs, length_Levs')
+      end;
+
+    val prefCl_Lev_thms =
+      let
+        fun mk_conjunct i z = HOLogic.mk_imp
+          (HOLogic.mk_conj (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z), mk_prefixeq kl_copy kl),
+          HOLogic.mk_mem (kl_copy, mk_Lev ss (mk_size kl_copy) i $ z));
+        val goal = list_all_free (kl :: kl_copy :: zs)
+          (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs));
+
+        val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
+
+        val prefCl_Lev = singleton (Proof_Context.export names_lthy lthy)
+          (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
+            (K (mk_prefCl_Lev_tac lthy cts Lev_0s Lev_Sucs)))
+          |> Thm.close_derivation;
+
+        val prefCl_Lev' = mk_specN (n + 2) prefCl_Lev;
+      in
+        map (fn i => prefCl_Lev' RS mk_conjunctN n i RS mp) ks
+      end;
+
+    val rv_last_thmss =
+      let
+        fun mk_conjunct i z i' z_copy = list_exists_free [z_copy]
+          (HOLogic.mk_eq
+            (mk_rv ss (mk_append (kl, HOLogic.mk_list sum_sbdT [mk_InN sbdTs k i'])) i $ z,
+            mk_InN activeAs z_copy i'));
+        val goal = list_all_free (k :: zs)
+          (Library.foldr1 HOLogic.mk_conj (map2 (fn i => fn z =>
+            Library.foldr1 HOLogic.mk_conj
+              (map2 (mk_conjunct i z) ks zs_copy)) ks zs));
+
+        val cTs = [SOME (certifyT lthy sum_sbdT)];
+        val cts = map (SOME o certify lthy) [Term.absfree kl' goal, kl];
+
+        val rv_last = singleton (Proof_Context.export names_lthy lthy)
+          (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
+            (K (mk_rv_last_tac cTs cts rv_Nils rv_Conss)))
+          |> Thm.close_derivation;
+
+        val rv_last' = mk_specN (n + 1) rv_last;
+      in
+        map (fn i => map (fn i' => rv_last' RS mk_conjunctN n i RS mk_conjunctN n i') ks) ks
+      end;
+
+    val set_rv_Lev_thmsss = if m = 0 then replicate n (replicate n []) else
+      let
+        fun mk_case s sets z z_free = Term.absfree z_free (Library.foldr1 HOLogic.mk_conj
+          (map2 (fn set => fn A => mk_leq (set $ (s $ z)) A) (take m sets) As));
+
+        fun mk_conjunct i z B = HOLogic.mk_imp
+          (HOLogic.mk_conj (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z), HOLogic.mk_mem (z, B)),
+          mk_sum_caseN (map4 mk_case ss setssAs zs zs') $ (mk_rv ss kl i $ z));
+
+        val goal = list_all_free (kl :: zs)
+          (Library.foldr1 HOLogic.mk_conj (map3 mk_conjunct ks zs Bs));
+
+        val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
+
+        val set_rv_Lev = singleton (Proof_Context.export names_lthy lthy)
+          (Goal.prove_sorry lthy [] []
+            (Logic.mk_implies (coalg_prem, HOLogic.mk_Trueprop goal))
+            (K (mk_set_rv_Lev_tac lthy m cts Lev_0s Lev_Sucs rv_Nils rv_Conss
+              coalg_set_thmss from_to_sbd_thmss)))
+          |> Thm.close_derivation;
+
+        val set_rv_Lev' = mk_specN (n + 1) set_rv_Lev;
+      in
+        map (fn i => map (fn i' =>
+          split_conj_thm (if n = 1 then set_rv_Lev' RS mk_conjunctN n i RS mp
+            else set_rv_Lev' RS mk_conjunctN n i RS mp RSN
+              (2, @{thm sum_case_weak_cong} RS iffD1) RS
+              (mk_sum_casesN n i' RS iffD1))) ks) ks
+      end;
+
+    val set_Lev_thmsss =
+      let
+        fun mk_conjunct i z =
+          let
+            fun mk_conjunct' i' sets s z' =
+              let
+                fun mk_conjunct'' i'' set z'' = HOLogic.mk_imp
+                  (HOLogic.mk_mem (z'', set $ (s $ z')),
+                    HOLogic.mk_mem (mk_append (kl,
+                      HOLogic.mk_list sum_sbdT [mk_InN sbdTs (mk_to_sbd s z' i' i'' $ z'') i'']),
+                      mk_Lev ss (HOLogic.mk_Suc nat) i $ z));
+              in
+                HOLogic.mk_imp (HOLogic.mk_eq (mk_rv ss kl i $ z, mk_InN activeAs z' i'),
+                  (Library.foldr1 HOLogic.mk_conj (map3 mk_conjunct'' ks (drop m sets) zs_copy2)))
+              end;
+          in
+            HOLogic.mk_imp (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z),
+              Library.foldr1 HOLogic.mk_conj (map4 mk_conjunct' ks setssAs ss zs_copy))
+          end;
+
+        val goal = list_all_free (kl :: zs @ zs_copy @ zs_copy2)
+          (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs));
+
+        val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
+
+        val set_Lev = singleton (Proof_Context.export names_lthy lthy)
+          (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
+            (K (mk_set_Lev_tac lthy cts Lev_0s Lev_Sucs rv_Nils rv_Conss from_to_sbd_thmss)))
+          |> Thm.close_derivation;
+
+        val set_Lev' = mk_specN (3 * n + 1) set_Lev;
+      in
+        map (fn i => map (fn i' => map (fn i'' => set_Lev' RS
+          mk_conjunctN n i RS mp RS
+          mk_conjunctN n i' RS mp RS
+          mk_conjunctN n i'' RS mp) ks) ks) ks
+      end;
+
+    val set_image_Lev_thmsss =
+      let
+        fun mk_conjunct i z =
+          let
+            fun mk_conjunct' i' sets =
+              let
+                fun mk_conjunct'' i'' set s z'' = HOLogic.mk_imp
+                  (HOLogic.mk_eq (mk_rv ss kl i $ z, mk_InN activeAs z'' i''),
+                  HOLogic.mk_mem (k, mk_image (mk_to_sbd s z'' i'' i') $ (set $ (s $ z''))));
+              in
+                HOLogic.mk_imp (HOLogic.mk_mem
+                  (mk_append (kl, HOLogic.mk_list sum_sbdT [mk_InN sbdTs k i']),
+                    mk_Lev ss (HOLogic.mk_Suc nat) i $ z),
+                  (Library.foldr1 HOLogic.mk_conj (map4 mk_conjunct'' ks sets ss zs_copy)))
+              end;
+          in
+            HOLogic.mk_imp (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z),
+              Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct' ks (drop m setssAs')))
+          end;
+
+        val goal = list_all_free (kl :: k :: zs @ zs_copy)
+          (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs));
+
+        val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
+
+        val set_image_Lev = singleton (Proof_Context.export names_lthy lthy)
+          (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
+            (K (mk_set_image_Lev_tac lthy cts Lev_0s Lev_Sucs rv_Nils rv_Conss
+              from_to_sbd_thmss to_sbd_inj_thmss)))
+          |> Thm.close_derivation;
+
+        val set_image_Lev' = mk_specN (2 * n + 2) set_image_Lev;
+      in
+        map (fn i => map (fn i' => map (fn i'' => set_image_Lev' RS
+          mk_conjunctN n i RS mp RS
+          mk_conjunctN n i'' RS mp RS
+          mk_conjunctN n i' RS mp) ks) ks) ks
+      end;
+
+    val mor_beh_thm =
+      Goal.prove_sorry lthy [] []
+        (fold_rev Logic.all (As @ Bs @ ss) (Logic.mk_implies (coalg_prem,
+          HOLogic.mk_Trueprop (mk_mor Bs ss carTAs strTAs (map (mk_beh ss) ks)))))
+        (mk_mor_beh_tac m mor_def mor_cong_thm
+          beh_defs carT_defs strT_defs isNode_defs
+          to_sbd_inj_thmss from_to_sbd_thmss Lev_0s Lev_Sucs rv_Nils rv_Conss Lev_sbd_thms
+          length_Lev_thms length_Lev'_thms prefCl_Lev_thms rv_last_thmss
+          set_rv_Lev_thmsss set_Lev_thmsss set_image_Lev_thmsss
+          set_mapss coalg_set_thmss map_comp_id_thms map_cong0s map_arg_cong_thms)
+      |> Thm.close_derivation;
+
+    val timer = time (timer "Behavioral morphism");
+
+    fun mk_LSBIS As i = mk_lsbis As (map (mk_carT As) ks) strTAs i;
+    fun mk_car_final As i =
+      mk_quotient (mk_carT As i) (mk_LSBIS As i);
+    fun mk_str_final As i =
+      mk_univ (HOLogic.mk_comp (Term.list_comb (nth final_maps (i - 1),
+        passive_ids @ map (mk_proj o mk_LSBIS As) ks), nth strTAs (i - 1)));
+
+    val car_finalAs = map (mk_car_final As) ks;
+    val str_finalAs = map (mk_str_final As) ks;
+    val car_finals = map (mk_car_final passive_UNIVs) ks;
+    val str_finals = map (mk_str_final passive_UNIVs) ks;
+
+    val coalgT_set_thmss = map (map (fn thm => coalgT_thm RS thm)) coalg_set_thmss;
+    val equiv_LSBIS_thms = map (fn thm => coalgT_thm RS thm) equiv_lsbis_thms;
+
+    val congruent_str_final_thms =
+      let
+        fun mk_goal R final_map strT =
+          fold_rev Logic.all As (HOLogic.mk_Trueprop
+            (mk_congruent R (HOLogic.mk_comp
+              (Term.list_comb (final_map, passive_ids @ map (mk_proj o mk_LSBIS As) ks), strT))));
+
+        val goals = map3 mk_goal (map (mk_LSBIS As) ks) final_maps strTAs;
+      in
+        map4 (fn goal => fn lsbisE => fn map_comp_id => fn map_cong0 =>
+          Goal.prove_sorry lthy [] [] goal
+            (K (mk_congruent_str_final_tac m lsbisE map_comp_id map_cong0 equiv_LSBIS_thms))
+          |> Thm.close_derivation)
+        goals lsbisE_thms map_comp_id_thms map_cong0s
+      end;
+
+    val coalg_final_thm = Goal.prove_sorry lthy [] [] (fold_rev Logic.all As
+      (HOLogic.mk_Trueprop (mk_coalg As car_finalAs str_finalAs)))
+      (K (mk_coalg_final_tac m coalg_def congruent_str_final_thms equiv_LSBIS_thms
+        set_mapss coalgT_set_thmss))
+      |> Thm.close_derivation;
+
+    val mor_T_final_thm = Goal.prove_sorry lthy [] [] (fold_rev Logic.all As
+      (HOLogic.mk_Trueprop (mk_mor carTAs strTAs car_finalAs str_finalAs
+        (map (mk_proj o mk_LSBIS As) ks))))
+      (K (mk_mor_T_final_tac mor_def congruent_str_final_thms equiv_LSBIS_thms))
+      |> Thm.close_derivation;
+
+    val mor_final_thm = mor_comp_thm OF [mor_beh_thm, mor_T_final_thm];
+    val in_car_final_thms = map (fn mor_image' => mor_image' OF
+      [tcoalg_thm RS mor_final_thm, UNIV_I]) mor_image'_thms;
+
+    val timer = time (timer "Final coalgebra");
+
+    val ((T_names, (T_glob_infos, T_loc_infos)), lthy) =
+      lthy
+      |> fold_map4 (fn b => fn mx => fn car_final => fn in_car_final =>
+        typedef (Binding.conceal b, params, mx) car_final NONE
+          (EVERY' [rtac exI, rtac in_car_final] 1)) bs mixfixes car_finals in_car_final_thms
+      |>> apsnd split_list o split_list;
+
+    val Ts = map (fn name => Type (name, params')) T_names;
+    fun mk_Ts passive = map (Term.typ_subst_atomic (passiveAs ~~ passive)) Ts;
+    val Ts' = mk_Ts passiveBs;
+    val Rep_Ts = map2 (fn info => fn T => Const (#Rep_name info, T --> treeQT)) T_glob_infos Ts;
+    val Abs_Ts = map2 (fn info => fn T => Const (#Abs_name info, treeQT --> T)) T_glob_infos Ts;
+
+    val Reps = map #Rep T_loc_infos;
+    val Rep_injects = map #Rep_inject T_loc_infos;
+    val Abs_inverses = map #Abs_inverse T_loc_infos;
+
+    val timer = time (timer "THE TYPEDEFs & Rep/Abs thms");
+
+    val UNIVs = map HOLogic.mk_UNIV Ts;
+    val FTs = mk_FTs (passiveAs @ Ts);
+    val FTs' = mk_FTs (passiveBs @ Ts);
+    val prodTs = map (HOLogic.mk_prodT o `I) Ts;
+    val prodFTs = mk_FTs (passiveAs @ prodTs);
+    val FTs_setss = mk_setss (passiveAs @ Ts);
+    val prodFT_setss = mk_setss (passiveAs @ prodTs);
+    val map_FTs = map2 (fn Ds => mk_map_of_bnf Ds treeQTs (passiveAs @ Ts)) Dss bnfs;
+    val map_FT_nths = map2 (fn Ds =>
+      mk_map_of_bnf Ds (passiveAs @ prodTs) (passiveAs @ Ts)) Dss bnfs;
+    val fstsTs = map fst_const prodTs;
+    val sndsTs = map snd_const prodTs;
+    val dtorTs = map2 (curry op -->) Ts FTs;
+    val ctorTs = map2 (curry op -->) FTs Ts;
+    val unfold_fTs = map2 (curry op -->) activeAs Ts;
+    val corec_sTs = map (Term.typ_subst_atomic (activeBs ~~ Ts)) sum_sTs;
+    val corec_maps = map (Term.subst_atomic_types (activeBs ~~ Ts)) map_Inls;
+    val corec_maps_rev = map (Term.subst_atomic_types (activeBs ~~ Ts)) map_Inls_rev;
+    val corec_Inls = map (Term.subst_atomic_types (activeBs ~~ Ts)) Inls;
+    val corec_UNIVs = map2 (HOLogic.mk_UNIV oo curry mk_sumT) Ts activeAs;
+
+    val (((((((((((((Jzs, Jzs'), Jz's), Jzs_copy), Jz's_copy), Jzs1), Jzs2),
+      FJzs), TRs), unfold_fs), corec_ss), phis), dtor_set_induct_phiss),
+      names_lthy) = names_lthy
+      |> mk_Frees' "z" Ts
+      ||>> mk_Frees "y" Ts'
+      ||>> mk_Frees "z'" Ts
+      ||>> mk_Frees "y'" Ts'
+      ||>> mk_Frees "z1" Ts
+      ||>> mk_Frees "z2" Ts
+      ||>> mk_Frees "x" prodFTs
+      ||>> mk_Frees "r" (map (mk_relT o `I) Ts)
+      ||>> mk_Frees "f" unfold_fTs
+      ||>> mk_Frees "s" corec_sTs
+      ||>> mk_Frees "P" (map2 mk_pred2T Ts Ts)
+      ||>> mk_Freess "P" (map (fn A => map (mk_pred2T A) Ts) passiveAs);
+
+    fun dtor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtorN ^ "_");
+    val dtor_name = Binding.name_of o dtor_bind;
+    val dtor_def_bind = rpair [] o Binding.conceal o Thm.def_binding o dtor_bind;
+
+    fun dtor_spec i rep str map_FT dtorT Jz Jz' =
+      let
+        val lhs = Free (dtor_name i, dtorT);
+        val rhs = Term.absfree Jz'
+          (Term.list_comb (map_FT, map HOLogic.id_const passiveAs @ Abs_Ts) $
+            (str $ (rep $ Jz)));
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((dtor_frees, (_, dtor_def_frees)), (lthy, lthy_old)) =
+      lthy
+      |> fold_map7 (fn i => fn rep => fn str => fn mapx => fn dtorT => fn Jz => fn Jz' =>
+        Specification.definition (SOME (dtor_bind i, NONE, NoSyn),
+          (dtor_def_bind i, dtor_spec i rep str mapx dtorT Jz Jz')))
+        ks Rep_Ts str_finals map_FTs dtorTs Jzs Jzs'
+      |>> apsnd split_list o split_list
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+    fun mk_dtors passive =
+      map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ (mk_params passive)) o
+        Morphism.term phi) dtor_frees;
+    val dtors = mk_dtors passiveAs;
+    val dtor's = mk_dtors passiveBs;
+    val dtor_defs = map ((fn thm => thm RS fun_cong) o Morphism.thm phi) dtor_def_frees;
+
+    val coalg_final_set_thmss = map (map (fn thm => coalg_final_thm RS thm)) coalg_set_thmss;
+    val (mor_Rep_thm, mor_Abs_thm) =
+      let
+        val mor_Rep =
+          Goal.prove_sorry lthy [] []
+            (HOLogic.mk_Trueprop (mk_mor UNIVs dtors car_finals str_finals Rep_Ts))
+            (mk_mor_Rep_tac m (mor_def :: dtor_defs) Reps Abs_inverses coalg_final_set_thmss
+              map_comp_id_thms map_cong0L_thms)
+          |> Thm.close_derivation;
+
+        val mor_Abs =
+          Goal.prove_sorry lthy [] []
+            (HOLogic.mk_Trueprop (mk_mor car_finals str_finals UNIVs dtors Abs_Ts))
+            (mk_mor_Abs_tac (mor_def :: dtor_defs) Abs_inverses)
+          |> Thm.close_derivation;
+      in
+        (mor_Rep, mor_Abs)
+      end;
+
+    val timer = time (timer "dtor definitions & thms");
+
+    fun unfold_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtor_unfoldN ^ "_");
+    val unfold_name = Binding.name_of o unfold_bind;
+    val unfold_def_bind = rpair [] o Binding.conceal o Thm.def_binding o unfold_bind;
+
+    fun unfold_spec i T AT abs f z z' =
+      let
+        val unfoldT = Library.foldr (op -->) (sTs, AT --> T);
+
+        val lhs = Term.list_comb (Free (unfold_name i, unfoldT), ss);
+        val rhs = Term.absfree z' (abs $ (f $ z));
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((unfold_frees, (_, unfold_def_frees)), (lthy, lthy_old)) =
+      lthy
+      |> fold_map7 (fn i => fn T => fn AT => fn abs => fn f => fn z => fn z' =>
+        Specification.definition
+          (SOME (unfold_bind i, NONE, NoSyn), (unfold_def_bind i, unfold_spec i T AT abs f z z')))
+          ks Ts activeAs Abs_Ts (map (fn i => HOLogic.mk_comp
+            (mk_proj (mk_LSBIS passive_UNIVs i), mk_beh ss i)) ks) zs zs'
+      |>> apsnd split_list o split_list
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+    val unfolds = map (Morphism.term phi) unfold_frees;
+    val unfold_names = map (fst o dest_Const) unfolds;
+    fun mk_unfolds passives actives =
+      map3 (fn name => fn T => fn active =>
+        Const (name, Library.foldr (op -->)
+          (map2 (curry op -->) actives (mk_FTs (passives @ actives)), active --> T)))
+      unfold_names (mk_Ts passives) actives;
+    fun mk_unfold Ts ss i = Term.list_comb (Const (nth unfold_names (i - 1), Library.foldr (op -->)
+      (map fastype_of ss, domain_type (fastype_of (nth ss (i - 1))) --> nth Ts (i - 1))), ss);
+    val unfold_defs = map ((fn thm => thm RS fun_cong) o Morphism.thm phi) unfold_def_frees;
+
+    val mor_unfold_thm =
+      let
+        val Abs_inverses' = map2 (curry op RS) in_car_final_thms Abs_inverses;
+        val morEs' = map (fn thm =>
+          (thm OF [tcoalg_thm RS mor_final_thm, UNIV_I]) RS sym) morE_thms;
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all ss
+            (HOLogic.mk_Trueprop (mk_mor active_UNIVs ss UNIVs dtors (map (mk_unfold Ts ss) ks))))
+          (K (mk_mor_unfold_tac m mor_UNIV_thm dtor_defs unfold_defs Abs_inverses' morEs'
+            map_comp_id_thms map_cong0s))
+        |> Thm.close_derivation
+      end;
+    val dtor_unfold_thms = map (fn thm => (thm OF [mor_unfold_thm, UNIV_I]) RS sym) morE_thms;
+
+    val (raw_coind_thms, raw_coind_thm) =
+      let
+        val prem = HOLogic.mk_Trueprop (mk_sbis passive_UNIVs UNIVs dtors TRs);
+        val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+          (map2 (fn R => fn T => mk_leq R (Id_const T)) TRs Ts));
+        val goal = fold_rev Logic.all TRs (Logic.mk_implies (prem, concl));
+      in
+        `split_conj_thm (Goal.prove_sorry lthy [] [] goal
+          (K (mk_raw_coind_tac bis_def bis_cong_thm bis_O_thm bis_converse_thm bis_Gr_thm
+            tcoalg_thm coalgT_thm mor_T_final_thm sbis_lsbis_thm
+            lsbis_incl_thms incl_lsbis_thms equiv_LSBIS_thms mor_Rep_thm Rep_injects))
+          |> Thm.close_derivation)
+      end;
+
+    val (unfold_unique_mor_thms, unfold_unique_mor_thm) =
+      let
+        val prem = HOLogic.mk_Trueprop (mk_mor active_UNIVs ss UNIVs dtors unfold_fs);
+        fun mk_fun_eq f i = HOLogic.mk_eq (f, mk_unfold Ts ss i);
+        val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+          (map2 mk_fun_eq unfold_fs ks));
+
+        val bis_thm = tcoalg_thm RSN (2, tcoalg_thm RS bis_image2_thm);
+        val mor_thm = mor_comp_thm OF [tcoalg_thm RS mor_final_thm, mor_Abs_thm];
+
+        val unique_mor = Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (ss @ unfold_fs) (Logic.mk_implies (prem, unique)))
+          (K (mk_unfold_unique_mor_tac raw_coind_thms bis_thm mor_thm unfold_defs))
+          |> Thm.close_derivation;
+      in
+        `split_conj_thm unique_mor
+      end;
+
+    val (dtor_unfold_unique_thms, dtor_unfold_unique_thm) = `split_conj_thm (split_conj_prems n
+      (mor_UNIV_thm RS iffD2 RS unfold_unique_mor_thm));
+
+    val unfold_dtor_thms = map (fn thm => mor_id_thm RS thm RS sym) unfold_unique_mor_thms;
+
+    val unfold_o_dtor_thms =
+      let
+        val mor = mor_comp_thm OF [mor_str_thm, mor_unfold_thm];
+      in
+        map2 (fn unique => fn unfold_ctor =>
+          trans OF [mor RS unique, unfold_ctor]) unfold_unique_mor_thms unfold_dtor_thms
+      end;
+
+    val timer = time (timer "unfold definitions & thms");
+
+    val map_dtors = map2 (fn Ds => fn bnf =>
+      Term.list_comb (mk_map_of_bnf Ds (passiveAs @ Ts) (passiveAs @ FTs) bnf,
+        map HOLogic.id_const passiveAs @ dtors)) Dss bnfs;
+
+    fun ctor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctorN ^ "_");
+    val ctor_name = Binding.name_of o ctor_bind;
+    val ctor_def_bind = rpair [] o Binding.conceal o Thm.def_binding o ctor_bind;
+
+    fun ctor_spec i ctorT =
+      let
+        val lhs = Free (ctor_name i, ctorT);
+        val rhs = mk_unfold Ts map_dtors i;
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((ctor_frees, (_, ctor_def_frees)), (lthy, lthy_old)) =
+      lthy
+      |> fold_map2 (fn i => fn ctorT =>
+        Specification.definition
+          (SOME (ctor_bind i, NONE, NoSyn), (ctor_def_bind i, ctor_spec i ctorT))) ks ctorTs
+      |>> apsnd split_list o split_list
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+    fun mk_ctors params =
+      map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ params) o Morphism.term phi)
+        ctor_frees;
+    val ctors = mk_ctors params';
+    val ctor_defs = map (Morphism.thm phi) ctor_def_frees;
+
+    val ctor_o_dtor_thms = map2 (fold_thms lthy o single) ctor_defs unfold_o_dtor_thms;
+
+    val dtor_o_ctor_thms =
+      let
+        fun mk_goal dtor ctor FT =
+         mk_Trueprop_eq (HOLogic.mk_comp (dtor, ctor), HOLogic.id_const FT);
+        val goals = map3 mk_goal dtors ctors FTs;
+      in
+        map5 (fn goal => fn ctor_def => fn unfold => fn map_comp_id => fn map_cong0L =>
+          Goal.prove_sorry lthy [] [] goal
+            (mk_dtor_o_ctor_tac ctor_def unfold map_comp_id map_cong0L unfold_o_dtor_thms)
+          |> Thm.close_derivation)
+          goals ctor_defs dtor_unfold_thms map_comp_id_thms map_cong0L_thms
+      end;
+
+    val dtor_ctor_thms = map (fn thm => thm RS @{thm pointfree_idE}) dtor_o_ctor_thms;
+    val ctor_dtor_thms = map (fn thm => thm RS @{thm pointfree_idE}) ctor_o_dtor_thms;
+
+    val bij_dtor_thms =
+      map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) ctor_o_dtor_thms dtor_o_ctor_thms;
+    val inj_dtor_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_dtor_thms;
+    val surj_dtor_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_dtor_thms;
+    val dtor_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_dtor_thms;
+    val dtor_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_dtor_thms;
+    val dtor_exhaust_thms = map (fn thm => thm RS exE) dtor_nchotomy_thms;
+
+    val bij_ctor_thms =
+      map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) dtor_o_ctor_thms ctor_o_dtor_thms;
+    val inj_ctor_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_ctor_thms;
+    val surj_ctor_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_ctor_thms;
+    val ctor_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_ctor_thms;
+    val ctor_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_ctor_thms;
+    val ctor_exhaust_thms = map (fn thm => thm RS exE) ctor_nchotomy_thms;
+
+    val timer = time (timer "ctor definitions & thms");
+
+    val corec_Inl_sum_thms =
+      let
+        val mor = mor_comp_thm OF [mor_sum_case_thm, mor_unfold_thm];
+      in
+        map2 (fn unique => fn unfold_dtor =>
+          trans OF [mor RS unique, unfold_dtor]) unfold_unique_mor_thms unfold_dtor_thms
+      end;
+
+    fun corec_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtor_corecN ^ "_");
+    val corec_name = Binding.name_of o corec_bind;
+    val corec_def_bind = rpair [] o Binding.conceal o Thm.def_binding o corec_bind;
+
+    val corec_strs =
+      map3 (fn dtor => fn sum_s => fn mapx =>
+        mk_sum_case
+          (HOLogic.mk_comp (Term.list_comb (mapx, passive_ids @ corec_Inls), dtor), sum_s))
+      dtors corec_ss corec_maps;
+
+    fun corec_spec i T AT =
+      let
+        val corecT = Library.foldr (op -->) (corec_sTs, AT --> T);
+
+        val lhs = Term.list_comb (Free (corec_name i, corecT), corec_ss);
+        val rhs = HOLogic.mk_comp (mk_unfold Ts corec_strs i, Inr_const T AT);
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((corec_frees, (_, corec_def_frees)), (lthy, lthy_old)) =
+      lthy
+      |> fold_map3 (fn i => fn T => fn AT =>
+        Specification.definition
+          (SOME (corec_bind i, NONE, NoSyn), (corec_def_bind i, corec_spec i T AT)))
+          ks Ts activeAs
+      |>> apsnd split_list o split_list
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+    val corecs = map (Morphism.term phi) corec_frees;
+    val corec_names = map (fst o dest_Const) corecs;
+    fun mk_corec ss i = Term.list_comb (Const (nth corec_names (i - 1), Library.foldr (op -->)
+      (map fastype_of ss, domain_type (fastype_of (nth ss (i - 1))) --> nth Ts (i - 1))), ss);
+    val corec_defs = map (Morphism.thm phi) corec_def_frees;
+
+    val sum_cases =
+      map2 (fn T => fn i => mk_sum_case (HOLogic.id_const T, mk_corec corec_ss i)) Ts ks;
+    val dtor_corec_thms =
+      let
+        fun mk_goal i corec_s corec_map dtor z =
+          let
+            val lhs = dtor $ (mk_corec corec_ss i $ z);
+            val rhs = Term.list_comb (corec_map, passive_ids @ sum_cases) $ (corec_s $ z);
+          in
+            fold_rev Logic.all (z :: corec_ss) (mk_Trueprop_eq (lhs, rhs))
+          end;
+        val goals = map5 mk_goal ks corec_ss corec_maps_rev dtors zs;
+      in
+        map3 (fn goal => fn unfold => fn map_cong0 =>
+          Goal.prove_sorry lthy [] [] goal
+            (mk_corec_tac m corec_defs unfold map_cong0 corec_Inl_sum_thms)
+          |> Thm.close_derivation)
+        goals dtor_unfold_thms map_cong0s
+      end;
+
+    val corec_unique_mor_thm =
+      let
+        val id_fs = map2 (fn T => fn f => mk_sum_case (HOLogic.id_const T, f)) Ts unfold_fs;
+        val prem = HOLogic.mk_Trueprop (mk_mor corec_UNIVs corec_strs UNIVs dtors id_fs);
+        fun mk_fun_eq f i = HOLogic.mk_eq (f, mk_corec corec_ss i);
+        val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+          (map2 mk_fun_eq unfold_fs ks));
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (corec_ss @ unfold_fs) (Logic.mk_implies (prem, unique)))
+          (mk_corec_unique_mor_tac corec_defs corec_Inl_sum_thms unfold_unique_mor_thm)
+          |> Thm.close_derivation
+      end;
+
+    val map_id0s_o_id =
+      map (fn thm =>
+        mk_trans (thm RS @{thm arg_cong2[of _ _ _ _ "op o", OF _ refl]}) @{thm id_o})
+      map_id0s;
+
+    val (dtor_corec_unique_thms, dtor_corec_unique_thm) =
+      `split_conj_thm (split_conj_prems n
+        (mor_UNIV_thm RS iffD2 RS corec_unique_mor_thm)
+        |> Local_Defs.unfold lthy (@{thms o_sum_case o_id id_o o_assoc sum_case_o_inj(1)} @
+           map_id0s_o_id @ sym_map_comps)
+        OF replicate n @{thm arg_cong2[of _ _ _ _ sum_case, OF refl]});
+
+    val timer = time (timer "corec definitions & thms");
+
+    val (dtor_map_coinduct_thm, coinduct_params, dtor_coinduct_thm) =
+      let
+        val zs = Jzs1 @ Jzs2;
+        val frees = phis @ zs;
+
+        val rels = map (Term.subst_atomic_types ((activeAs ~~ Ts) @ (activeBs ~~ Ts))) relsAsBs;
+
+        fun mk_concl phi z1 z2 = HOLogic.mk_imp (phi $ z1 $ z2, HOLogic.mk_eq (z1, z2));
+        val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+          (map3 mk_concl phis Jzs1 Jzs2));
+
+        fun mk_rel_prem phi dtor rel Jz Jz_copy =
+          let
+            val concl = Term.list_comb (rel, map HOLogic.eq_const passiveAs @ phis) $
+              (dtor $ Jz) $ (dtor $ Jz_copy);
+          in
+            HOLogic.mk_Trueprop
+              (list_all_free [Jz, Jz_copy] (HOLogic.mk_imp (phi $ Jz $ Jz_copy, concl)))
+          end;
+
+        val rel_prems = map5 mk_rel_prem phis dtors rels Jzs Jzs_copy;
+        val dtor_coinduct_goal =
+          fold_rev Logic.all frees (Logic.list_implies (rel_prems, concl));
+
+        val dtor_coinduct =
+          Goal.prove_sorry lthy [] [] dtor_coinduct_goal
+            (K (mk_dtor_coinduct_tac m raw_coind_thm bis_rel_thm rel_congs))
+          |> Thm.close_derivation;
+
+        fun mk_prem phi dtor map_nth sets Jz Jz_copy FJz =
+          let
+            val xs = [Jz, Jz_copy];
+
+            fun mk_map_conjunct nths x =
+              HOLogic.mk_eq (Term.list_comb (map_nth, passive_ids @ nths) $ FJz, dtor $ x);
+
+            fun mk_set_conjunct set phi z1 z2 =
+              list_all_free [z1, z2]
+                (HOLogic.mk_imp (HOLogic.mk_mem (HOLogic.mk_prod (z1, z2), set $ FJz),
+                  phi $ z1 $ z2));
+
+            val concl = list_exists_free [FJz] (HOLogic.mk_conj
+              (Library.foldr1 HOLogic.mk_conj (map2 mk_map_conjunct [fstsTs, sndsTs] xs),
+              Library.foldr1 HOLogic.mk_conj
+                (map4 mk_set_conjunct (drop m sets) phis Jzs1 Jzs2)));
+          in
+            fold_rev Logic.all xs (Logic.mk_implies
+              (HOLogic.mk_Trueprop (Term.list_comb (phi, xs)), HOLogic.mk_Trueprop concl))
+          end;
+
+        val prems = map7 mk_prem phis dtors map_FT_nths prodFT_setss Jzs Jzs_copy FJzs;
+
+        val dtor_map_coinduct_goal = fold_rev Logic.all frees (Logic.list_implies (prems, concl));
+        val dtor_map_coinduct =
+          Goal.prove_sorry lthy [] [] dtor_map_coinduct_goal
+            (K (mk_dtor_map_coinduct_tac m ks raw_coind_thm bis_def))
+          |> Thm.close_derivation;
+      in
+        (dtor_map_coinduct, rev (Term.add_tfrees dtor_map_coinduct_goal []), dtor_coinduct)
+      end;
+
+    val timer = time (timer "coinduction");
+
+    val setss_by_bnf = map (fn i => map2 (mk_hset dtors i) ls passiveAs) ks;
+    val setss_by_range = transpose setss_by_bnf;
+
+    val (hset_dtor_incl_thmss, hset_hset_dtor_incl_thmsss, dtor_hset_induct_thms) =
+      let
+        fun tinst_of dtor =
+          map (SOME o certify lthy) (dtor :: remove (op =) dtor dtors);
+        fun tinst_of' dtor = case tinst_of dtor of t :: ts => t :: NONE :: ts;
+        val Tinst = map (pairself (certifyT lthy))
+          (map Logic.varifyT_global (deads @ allAs) ~~ (deads @ passiveAs @ Ts));
+        val set_incl_thmss =
+          map2 (fn dtor => map (singleton (Proof_Context.export names_lthy lthy) o
+            Drule.instantiate' [] (tinst_of' dtor) o
+            Thm.instantiate (Tinst, []) o Drule.zero_var_indexes))
+          dtors set_incl_hset_thmss;
+
+        val tinst = splice (map (SOME o certify lthy) dtors) (replicate n NONE)
+        val set_minimal_thms =
+          map (Drule.instantiate' [] tinst o Thm.instantiate (Tinst, []) o
+            Drule.zero_var_indexes)
+          hset_minimal_thms;
+
+        val set_set_incl_thmsss =
+          map2 (fn dtor => map (map (singleton (Proof_Context.export names_lthy lthy) o
+            Drule.instantiate' [] (NONE :: tinst_of' dtor) o
+            Thm.instantiate (Tinst, []) o Drule.zero_var_indexes)))
+          dtors set_hset_incl_hset_thmsss;
+
+        val set_set_incl_thmsss' = transpose (map transpose set_set_incl_thmsss);
+
+        val incls =
+          maps (map (fn thm => thm RS @{thm subset_Collect_iff})) set_incl_thmss @
+            @{thms subset_Collect_iff[OF subset_refl]};
+
+        fun mk_induct_tinst phis jsets y y' =
+          map4 (fn phi => fn jset => fn Jz => fn Jz' =>
+            SOME (certify lthy (Term.absfree Jz' (HOLogic.mk_Collect (fst y', snd y',
+              HOLogic.mk_conj (HOLogic.mk_mem (y, jset $ Jz), phi $ y $ Jz))))))
+          phis jsets Jzs Jzs';
+        val dtor_set_induct_thms =
+          map6 (fn set_minimal => fn set_set_inclss => fn jsets => fn y => fn y' => fn phis =>
+            ((set_minimal
+              |> Drule.instantiate' [] (mk_induct_tinst phis jsets y y')
+              |> unfold_thms lthy incls) OF
+              (replicate n ballI @
+                maps (map (fn thm => thm RS @{thm subset_CollectI})) set_set_inclss))
+            |> singleton (Proof_Context.export names_lthy lthy)
+            |> rule_by_tactic lthy (ALLGOALS (TRY o etac asm_rl)))
+          set_minimal_thms set_set_incl_thmsss' setss_by_range ys ys' dtor_set_induct_phiss
+      in
+        (set_incl_thmss, set_set_incl_thmsss, dtor_set_induct_thms)
+      end;
+
+    fun mk_dtor_map_DEADID_thm dtor_inject map_id0 =
+      trans OF [iffD2 OF [dtor_inject, id_apply], map_id0 RS sym];
+
+    fun mk_dtor_Jrel_DEADID_thm dtor_inject bnf =
+      trans OF [rel_eq_of_bnf bnf RS @{thm predicate2_eqD}, dtor_inject] RS sym;
+
+    val JphiTs = map2 mk_pred2T passiveAs passiveBs;
+    val Jpsi1Ts = map2 mk_pred2T passiveAs passiveCs;
+    val Jpsi2Ts = map2 mk_pred2T passiveCs passiveBs;
+    val prodTsTs' = map2 (curry HOLogic.mk_prodT) Ts Ts';
+    val fstsTsTs' = map fst_const prodTsTs';
+    val sndsTsTs' = map snd_const prodTsTs';
+    val activephiTs = map2 mk_pred2T activeAs activeBs;
+    val activeJphiTs = map2 mk_pred2T Ts Ts';
+    val (((((Jphis, Jpsi1s), Jpsi2s), activephis), activeJphis), names_lthy) = names_lthy
+      |> mk_Frees "R" JphiTs
+      ||>> mk_Frees "R" Jpsi1Ts
+      ||>> mk_Frees "Q" Jpsi2Ts
+      ||>> mk_Frees "S" activephiTs
+      ||>> mk_Frees "JR" activeJphiTs;
+    val rels = map2 (fn Ds => mk_rel_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
+    val in_rels = map in_rel_of_bnf bnfs;
+
+    fun mk_Jrel_DEADID_coinduct_thm () = 
+      mk_rel_xtor_co_induct_thm Greatest_FP rels activeJphis (map HOLogic.eq_const Ts) Jphis
+        Jzs Jz's dtors dtor's (fn {context = ctxt, prems} =>
+          (unfold_thms_tac ctxt @{thms le_fun_def le_bool_def all_simps(1,2)[symmetric]} THEN
+          REPEAT_DETERM (rtac allI 1) THEN rtac (dtor_coinduct_thm OF prems) 1)) lthy;
+
+    (*register new codatatypes as BNFs*)
+    val (timer, Jbnfs, (dtor_Jmap_o_thms, dtor_Jmap_thms), dtor_Jset_thmss',
+      dtor_Jrel_thms, Jrel_coinduct_thm, Jbnf_notes, lthy) =
+      if m = 0 then
+        (timer, replicate n DEADID_bnf,
+        map_split (`(mk_pointfree lthy)) (map2 mk_dtor_map_DEADID_thm dtor_inject_thms map_ids),
+        replicate n [], map2 mk_dtor_Jrel_DEADID_thm dtor_inject_thms bnfs,
+        mk_Jrel_DEADID_coinduct_thm (), [], lthy)
+      else let
+        val fTs = map2 (curry op -->) passiveAs passiveBs;
+        val gTs = map2 (curry op -->) passiveBs passiveCs;
+        val uTs = map2 (curry op -->) Ts Ts';
+
+        val ((((((((fs, fs'), fs_copy), gs), us), (Jys, Jys')), (Jys_copy, Jys'_copy)),
+          (ys_copy, ys'_copy)), names_lthy) = names_lthy
+          |> mk_Frees' "f" fTs
+          ||>> mk_Frees "f" fTs
+          ||>> mk_Frees "g" gTs
+          ||>> mk_Frees "u" uTs
+          ||>> mk_Frees' "b" Ts'
+          ||>> mk_Frees' "b" Ts'
+          ||>> mk_Frees' "y" passiveAs;
+
+        val map_FTFT's = map2 (fn Ds =>
+          mk_map_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
+
+        fun mk_maps ATs BTs Ts mk_T =
+          map2 (fn Ds => mk_map_of_bnf Ds (ATs @ Ts) (BTs @ map mk_T Ts)) Dss bnfs;
+        fun mk_Fmap mk_const fs Ts Fmap = Term.list_comb (Fmap, fs @ map mk_const Ts);
+        fun mk_map mk_const mk_T Ts fs Ts' dtors mk_maps =
+          mk_unfold Ts' (map2 (fn dtor => fn Fmap =>
+            HOLogic.mk_comp (mk_Fmap mk_const fs Ts Fmap, dtor)) dtors (mk_maps Ts mk_T));
+        val mk_map_id = mk_map HOLogic.id_const I;
+        val mk_mapsAB = mk_maps passiveAs passiveBs;
+        val fs_maps = map (mk_map_id Ts fs Ts' dtors mk_mapsAB) ks;
+
+        val set_bss =
+          map (flat o map2 (fn B => fn b =>
+            if member (op =) resDs (TFree B) then [] else [b]) resBs) set_bss0;
+
+        fun close_wit I wit = (I, fold_rev Term.absfree (map (nth ys') I) wit);
+
+        val all_unitTs = replicate live HOLogic.unitT;
+        val unitTs = replicate n HOLogic.unitT;
+        val unit_funs = replicate n (Term.absdummy HOLogic.unitT HOLogic.unit);
+        fun mk_map_args I =
+          map (fn i =>
+            if member (op =) I i then Term.absdummy HOLogic.unitT (nth ys i)
+            else mk_undefined (HOLogic.unitT --> nth passiveAs i))
+          (0 upto (m - 1));
+
+        fun mk_nat_wit Ds bnf (I, wit) () =
+          let
+            val passiveI = filter (fn i => i < m) I;
+            val map_args = mk_map_args passiveI;
+          in
+            Term.absdummy HOLogic.unitT (Term.list_comb
+              (mk_map_of_bnf Ds all_unitTs (passiveAs @ unitTs) bnf, map_args @ unit_funs) $ wit)
+          end;
+
+        fun mk_dummy_wit Ds bnf I =
+          let
+            val map_args = mk_map_args I;
+          in
+            Term.absdummy HOLogic.unitT (Term.list_comb
+              (mk_map_of_bnf Ds all_unitTs (passiveAs @ unitTs) bnf, map_args @ unit_funs) $
+              mk_undefined (mk_T_of_bnf Ds all_unitTs bnf))
+          end;
+
+        val nat_witss =
+          map2 (fn Ds => fn bnf => mk_wits_of_bnf (replicate (nwits_of_bnf bnf) Ds)
+            (replicate (nwits_of_bnf bnf) (replicate live HOLogic.unitT)) bnf
+            |> map (fn (I, wit) =>
+              (I, Lazy.lazy (mk_nat_wit Ds bnf (I, Term.list_comb (wit, map (K HOLogic.unit) I))))))
+          Dss bnfs;
+
+        val nat_wit_thmss = map2 (curry op ~~) nat_witss (map wit_thmss_of_bnf bnfs)
+
+        val Iss = map (map fst) nat_witss;
+
+        fun filter_wits (I, wit) =
+          let val J = filter (fn i => i < m) I;
+          in (J, (length J < length I, wit)) end;
+
+        val wit_treess = map_index (fn (i, Is) =>
+          map_index (finish Iss m [i+m] (i+m)) Is) Iss
+          |> map (minimize_wits o map filter_wits o minimize_wits o flat);
+
+        val coind_wit_argsss =
+          map (map (tree_to_coind_wits nat_wit_thmss o snd o snd) o filter (fst o snd)) wit_treess;
+
+        val nonredundant_coind_wit_argsss =
+          fold (fn i => fn argsss =>
+            nth_map (i - 1) (filter_out (fn xs =>
+              exists (fn ys =>
+                let
+                  val xs' = (map (fst o fst) xs, snd (fst (hd xs)));
+                  val ys' = (map (fst o fst) ys, snd (fst (hd ys)));
+                in
+                  eq_pair (subset (op =)) (eq_set (op =)) (xs', ys') andalso not (fst xs' = fst ys')
+                end)
+              (flat argsss)))
+            argsss)
+          ks coind_wit_argsss;
+
+        fun prepare_args args =
+          let
+            val I = snd (fst (hd args));
+            val (dummys, args') =
+              map_split (fn i =>
+                (case find_first (fn arg => fst (fst arg) = i - 1) args of
+                  SOME (_, ((_, wit), thms)) => (NONE, (Lazy.force wit, thms))
+                | NONE =>
+                  (SOME (i - 1), (mk_dummy_wit (nth Dss (i - 1)) (nth bnfs (i - 1)) I, []))))
+              ks;
+          in
+            ((I, dummys), apsnd flat (split_list args'))
+          end;
+
+        fun mk_coind_wits ((I, dummys), (args, thms)) =
+          ((I, dummys), (map (fn i => mk_unfold Ts args i $ HOLogic.unit) ks, thms));
+
+        val coind_witss =
+          maps (map (mk_coind_wits o prepare_args)) nonredundant_coind_wit_argsss;
+
+        val witss = map2 (fn Ds => fn bnf => mk_wits_of_bnf
+          (replicate (nwits_of_bnf bnf) Ds)
+          (replicate (nwits_of_bnf bnf) (passiveAs @ Ts)) bnf) Dss bnfs;
+
+        val ctor_witss =
+          map (map (uncurry close_wit o tree_to_ctor_wit ys ctors witss o snd o snd) o
+            filter_out (fst o snd)) wit_treess;
+
+        fun mk_coind_wit_thms ((I, dummys), (wits, wit_thms)) =
+          let
+            fun mk_goal sets y y_copy y'_copy j =
+              let
+                fun mk_conjunct set z dummy wit =
+                  mk_Ball (set $ z) (Term.absfree y'_copy
+                    (if dummy = NONE orelse member (op =) I (j - 1) then
+                      HOLogic.mk_imp (HOLogic.mk_eq (z, wit),
+                        if member (op =) I (j - 1) then HOLogic.mk_eq (y_copy, y)
+                        else @{term False})
+                    else @{term True}));
+              in
+                fold_rev Logic.all (map (nth ys) I @ Jzs) (HOLogic.mk_Trueprop
+                  (Library.foldr1 HOLogic.mk_conj (map4 mk_conjunct sets Jzs dummys wits)))
+              end;
+            val goals = map5 mk_goal setss_by_range ys ys_copy ys'_copy ls;
+          in
+            map2 (fn goal => fn induct =>
+              Goal.prove_sorry lthy [] [] goal
+                (mk_coind_wit_tac induct dtor_unfold_thms (flat set_mapss) wit_thms)
+              |> Thm.close_derivation)
+            goals dtor_hset_induct_thms
+            |> map split_conj_thm
+            |> transpose
+            |> map (map_filter (try (fn thm => thm RS bspec RS mp)))
+            |> curry op ~~ (map_index Library.I (map (close_wit I) wits))
+            |> filter (fn (_, thms) => length thms = m)
+          end;
+
+        val coind_wit_thms = maps mk_coind_wit_thms coind_witss;
+
+        val (wit_thmss, all_witss) =
+          fold (fn ((i, wit), thms) => fn witss =>
+            nth_map i (fn (thms', wits) => (thms @ thms', wit :: wits)) witss)
+          coind_wit_thms (map (pair []) ctor_witss)
+          |> map (apsnd (map snd o minimize_wits))
+          |> split_list;
+
+        val (Jbnf_consts, lthy) =
+          fold_map8 (fn b => fn map_b => fn rel_b => fn set_bs => fn mapx => fn sets => fn wits =>
+              fn T => fn lthy =>
+            define_bnf_consts Dont_Inline (user_policy Note_Some lthy) (SOME deads)
+              map_b rel_b set_bs
+              ((((((b, T), fold_rev Term.absfree fs' mapx), sets), sbd), wits), NONE) lthy)
+          bs map_bs rel_bs set_bss fs_maps setss_by_bnf all_witss Ts lthy;
+
+        val (_, Jconsts, Jconst_defs, mk_Jconsts) = split_list4 Jbnf_consts;
+        val (_, Jsetss, Jbds_Ds, Jwitss_Ds, _) = split_list5 Jconsts;
+        val (Jmap_defs, Jset_defss, Jbd_defs, Jwit_defss, Jrel_defs) = split_list5 Jconst_defs;
+        val (mk_Jmaps_Ds, mk_Jt_Ds, _, mk_Jrels_Ds, _) = split_list5 mk_Jconsts;
+
+        val Jrel_unabs_defs = map (fn def => mk_unabs_def m (def RS meta_eq_to_obj_eq)) Jrel_defs;
+        val Jset_defs = flat Jset_defss;
+        val Jset_unabs_defs = map (fn def => def RS meta_eq_to_obj_eq RS fun_cong) Jset_defs;
+
+        fun mk_Jmaps As Bs = map (fn mk => mk deads As Bs) mk_Jmaps_Ds;
+        fun mk_Jsetss As = map2 (fn mk => fn Jsets => map (mk deads As) Jsets) mk_Jt_Ds Jsetss;
+        val Jbds = map2 (fn mk => mk deads passiveAs) mk_Jt_Ds Jbds_Ds;
+        val Jwitss =
+          map2 (fn mk => fn Jwits => map (mk deads passiveAs o snd) Jwits) mk_Jt_Ds Jwitss_Ds;
+        fun mk_Jrels As Bs = map (fn mk => mk deads As Bs) mk_Jrels_Ds;
+
+        val Jmaps = mk_Jmaps passiveAs passiveBs;
+        val fs_Jmaps = map (fn m => Term.list_comb (m, fs)) Jmaps;
+        val fs_copy_Jmaps = map (fn m => Term.list_comb (m, fs_copy)) Jmaps;
+        val gs_Jmaps = map (fn m => Term.list_comb (m, gs)) (mk_Jmaps passiveBs passiveCs);
+        val fgs_Jmaps = map (fn m => Term.list_comb (m, map2 (curry HOLogic.mk_comp) gs fs))
+          (mk_Jmaps passiveAs passiveCs);
+        val (Jsetss_by_range, Jsetss_by_bnf) = `transpose (mk_Jsetss passiveAs);
+
+        val timer = time (timer "bnf constants for the new datatypes");
+
+        val (dtor_Jmap_thms, Jmap_thms) =
+          let
+            fun mk_goal fs_Jmap map dtor dtor' = fold_rev Logic.all fs
+              (mk_Trueprop_eq (HOLogic.mk_comp (dtor', fs_Jmap),
+                HOLogic.mk_comp (Term.list_comb (map, fs @ fs_Jmaps), dtor)));
+            val goals = map4 mk_goal fs_Jmaps map_FTFT's dtors dtor's;
+            val cTs = map (SOME o certifyT lthy) FTs';
+            val maps =
+              map5 (fn goal => fn cT => fn unfold => fn map_comp => fn map_cong0 =>
+                Goal.prove_sorry lthy [] [] goal
+                  (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jmap_defs THEN
+                     mk_map_tac m n cT unfold map_comp map_cong0)
+                |> Thm.close_derivation)
+              goals cTs dtor_unfold_thms map_comps map_cong0s;
+          in
+            map_split (fn thm => (thm RS @{thm comp_eq_dest}, thm)) maps
+          end;
+
+        val dtor_Jmap_unique_thm =
+          let
+            fun mk_prem u map dtor dtor' =
+              mk_Trueprop_eq (HOLogic.mk_comp (dtor', u),
+                HOLogic.mk_comp (Term.list_comb (map, fs @ us), dtor));
+            val prems = map4 mk_prem us map_FTFT's dtors dtor's;
+            val goal =
+              HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+                (map2 (curry HOLogic.mk_eq) us fs_Jmaps));
+          in
+            Goal.prove_sorry lthy [] []
+              (fold_rev Logic.all (us @ fs) (Logic.list_implies (prems, goal)))
+                (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jmap_defs THEN
+                  mk_dtor_map_unique_tac dtor_unfold_unique_thm sym_map_comps ctxt)
+              |> Thm.close_derivation
+          end;
+
+        val Jmap_comp0_thms =
+          let
+            val goal = fold_rev Logic.all (fs @ gs)
+              (HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+                (map3 (fn fmap => fn gmap => fn fgmap =>
+                   HOLogic.mk_eq (HOLogic.mk_comp (gmap, fmap), fgmap))
+                fs_Jmaps gs_Jmaps fgs_Jmaps)))
+          in
+            split_conj_thm (Goal.prove_sorry lthy [] [] goal
+              (K (mk_map_comp0_tac Jmap_thms map_comp0s dtor_Jmap_unique_thm))
+              |> Thm.close_derivation)
+          end;
+
+        val timer = time (timer "map functions for the new codatatypes");
+
+        val (dtor_Jset_thmss', dtor_Jset_thmss) =
+          let
+            fun mk_simp_goal relate pas_set act_sets sets dtor z set =
+              relate (set $ z, mk_union (pas_set $ (dtor $ z),
+                 Library.foldl1 mk_union
+                   (map2 (fn X => mk_UNION (X $ (dtor $ z))) act_sets sets)));
+            fun mk_goals eq =
+              map2 (fn i => fn sets =>
+                map4 (fn Fsets =>
+                  mk_simp_goal eq (nth Fsets (i - 1)) (drop m Fsets) sets)
+                FTs_setss dtors Jzs sets)
+              ls Jsetss_by_range;
+
+            val le_goals = map
+              (fold_rev Logic.all Jzs o HOLogic.mk_Trueprop o Library.foldr1 HOLogic.mk_conj)
+              (mk_goals (uncurry mk_leq));
+            val set_le_thmss = map split_conj_thm
+              (map4 (fn goal => fn hset_minimal => fn set_hsets => fn set_hset_hsetss =>
+                Goal.prove_sorry lthy [] [] goal
+                  (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jset_defs THEN
+                    mk_set_le_tac n hset_minimal set_hsets set_hset_hsetss)
+                |> Thm.close_derivation)
+              le_goals hset_minimal_thms set_hset_thmss' set_hset_hset_thmsss');
+
+            val ge_goalss = map (map2 (fn z => fn goal =>
+                Logic.all z (HOLogic.mk_Trueprop goal)) Jzs)
+              (mk_goals (uncurry mk_leq o swap));
+            val set_ge_thmss = 
+              map3 (map3 (fn goal => fn set_incl_hset => fn set_hset_incl_hsets =>
+                Goal.prove_sorry lthy [] [] goal
+                  (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jset_defs THEN
+                    mk_set_ge_tac n set_incl_hset set_hset_incl_hsets)
+                |> Thm.close_derivation))
+              ge_goalss set_incl_hset_thmss' set_hset_incl_hset_thmsss'
+          in
+            map2 (map2 (fn le => fn ge => equalityI OF [le, ge])) set_le_thmss set_ge_thmss
+            |> `transpose
+          end;
+
+        val timer = time (timer "set functions for the new codatatypes");
+
+        val colss = map2 (fn j => fn T =>
+          map (fn i => mk_hset_rec dtors nat i j T) ks) ls passiveAs;
+        val colss' = map2 (fn j => fn T =>
+          map (fn i => mk_hset_rec dtor's nat i j T) ks) ls passiveBs;
+
+        val col_natural_thmss =
+          let
+            fun mk_col_natural f map z col col' =
+              HOLogic.mk_eq (mk_image f $ (col $ z), col' $ (map $ z));
+
+            fun mk_goal f cols cols' = list_all_free Jzs (Library.foldr1 HOLogic.mk_conj
+              (map4 (mk_col_natural f) fs_Jmaps Jzs cols cols'));
+
+            val goals = map3 mk_goal fs colss colss';
+
+            val ctss =
+              map (fn phi => map (SOME o certify lthy) [Term.absfree nat' phi, nat]) goals;
+
+            val thms =
+              map4 (fn goal => fn cts => fn rec_0s => fn rec_Sucs =>
+                singleton (Proof_Context.export names_lthy lthy)
+                  (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
+                    (mk_col_natural_tac cts rec_0s rec_Sucs dtor_Jmap_thms set_mapss))
+                |> Thm.close_derivation)
+              goals ctss hset_rec_0ss' hset_rec_Sucss';
+          in
+            map (split_conj_thm o mk_specN n) thms
+          end;
+
+        val col_bd_thmss =
+          let
+            fun mk_col_bd z col bd = mk_ordLeq (mk_card_of (col $ z)) bd;
+
+            fun mk_goal bds cols = list_all_free Jzs (Library.foldr1 HOLogic.mk_conj
+              (map3 mk_col_bd Jzs cols bds));
+
+            val goals = map (mk_goal Jbds) colss;
+
+            val ctss = map (fn phi => map (SOME o certify lthy) [Term.absfree nat' phi, nat])
+              (map (mk_goal (replicate n sbd)) colss);
+
+            val thms =
+              map5 (fn j => fn goal => fn cts => fn rec_0s => fn rec_Sucs =>
+                singleton (Proof_Context.export names_lthy lthy)
+                  (Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop goal)
+                    (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jbd_defs THEN
+                      mk_col_bd_tac m j cts rec_0s rec_Sucs sbd_Card_order sbd_Cinfinite set_sbdss))
+                |> Thm.close_derivation)
+              ls goals ctss hset_rec_0ss' hset_rec_Sucss';
+          in
+            map (split_conj_thm o mk_specN n) thms
+          end;
+
+        val map_cong0_thms =
+          let
+            val cTs = map (SOME o certifyT lthy o
+              Term.typ_subst_atomic (passiveAs ~~ passiveBs) o TFree) coinduct_params;
+
+            fun mk_prem z set f g y y' =
+              mk_Ball (set $ z) (Term.absfree y' (HOLogic.mk_eq (f $ y, g $ y)));
+
+            fun mk_prems sets z =
+              Library.foldr1 HOLogic.mk_conj (map5 (mk_prem z) sets fs fs_copy ys ys')
+
+            fun mk_map_cong0 sets z fmap gmap =
+              HOLogic.mk_imp (mk_prems sets z, HOLogic.mk_eq (fmap $ z, gmap $ z));
+
+            fun mk_coind_body sets (x, T) z fmap gmap y y_copy =
+              HOLogic.mk_conj
+                (HOLogic.mk_mem (z, HOLogic.mk_Collect (x, T, mk_prems sets z)),
+                  HOLogic.mk_conj (HOLogic.mk_eq (y, fmap $ z),
+                    HOLogic.mk_eq (y_copy, gmap $ z)))
+
+            fun mk_cphi sets (z' as (x, T)) z fmap gmap y' y y'_copy y_copy =
+              HOLogic.mk_exists (x, T, mk_coind_body sets z' z fmap gmap y y_copy)
+              |> Term.absfree y'_copy
+              |> Term.absfree y'
+              |> certify lthy;
+
+            val cphis = map9 mk_cphi
+              Jsetss_by_bnf Jzs' Jzs fs_Jmaps fs_copy_Jmaps Jys' Jys Jys'_copy Jys_copy;
+
+            val coinduct = unfold_thms lthy Jset_defs
+              (Drule.instantiate' cTs (map SOME cphis) dtor_map_coinduct_thm);
+
+            val goal =
+              HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+                (map4 mk_map_cong0 Jsetss_by_bnf Jzs fs_Jmaps fs_copy_Jmaps));
+
+            val thm = singleton (Proof_Context.export names_lthy lthy)
+              (Goal.prove_sorry lthy [] [] goal
+                (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jset_defs THEN
+                  mk_mcong_tac lthy m (rtac coinduct) map_comps dtor_Jmap_thms map_cong0s
+                    set_mapss set_hset_thmss set_hset_hset_thmsss))
+              |> Thm.close_derivation
+          in
+            split_conj_thm thm
+          end;
+
+        val in_Jrels = map (fn def => trans OF [def, @{thm OO_Grp_alt}] RS @{thm predicate2_eqD})
+            Jrel_unabs_defs;
+
+        val fold_Jsets = fold_thms lthy Jset_unabs_defs;
+        val dtor_Jset_incl_thmss = map (map fold_Jsets) hset_dtor_incl_thmss;
+        val dtor_set_Jset_incl_thmsss = map (map (map fold_Jsets)) hset_hset_dtor_incl_thmsss;
+        val dtor_Jset_induct_thms = map fold_Jsets dtor_hset_induct_thms;
+        val Jwit_thmss = map (map fold_Jsets) wit_thmss;
+
+        val Jrels = mk_Jrels passiveAs passiveBs;
+        val Jrelphis = map (fn rel => Term.list_comb (rel, Jphis)) Jrels;
+        val relphis = map (fn rel => Term.list_comb (rel, Jphis @ Jrelphis)) rels;
+        val Jrelpsi1s = map (fn rel => Term.list_comb (rel, Jpsi1s)) (mk_Jrels passiveAs passiveCs);
+        val Jrelpsi2s = map (fn rel => Term.list_comb (rel, Jpsi2s)) (mk_Jrels passiveCs passiveBs);
+        val Jrelpsi12s = map (fn rel =>
+            Term.list_comb (rel, map2 (curry mk_rel_compp) Jpsi1s Jpsi2s)) Jrels;
+
+        val dtor_Jrel_thms =
+          let
+            fun mk_goal Jz Jz' dtor dtor' Jrelphi relphi = fold_rev Logic.all (Jz :: Jz' :: Jphis)
+              (mk_Trueprop_eq (Jrelphi $ Jz $ Jz', relphi $ (dtor $ Jz) $ (dtor' $ Jz')));
+            val goals = map6 mk_goal Jzs Jz's dtors dtor's Jrelphis relphis;
+          in
+            map12 (fn i => fn goal => fn in_rel => fn map_comp0 => fn map_cong0 =>
+              fn dtor_map => fn dtor_sets => fn dtor_inject => fn dtor_ctor =>
+              fn set_map0s => fn dtor_set_incls => fn dtor_set_set_inclss =>
+              Goal.prove_sorry lthy [] [] goal
+                (K (mk_dtor_rel_tac lthy in_Jrels i in_rel map_comp0 map_cong0 dtor_map dtor_sets
+                  dtor_inject dtor_ctor set_map0s dtor_set_incls dtor_set_set_inclss))
+              |> Thm.close_derivation)
+            ks goals in_rels map_comps map_cong0s dtor_Jmap_thms dtor_Jset_thmss'
+              dtor_inject_thms dtor_ctor_thms set_mapss dtor_Jset_incl_thmss
+              dtor_set_Jset_incl_thmsss
+          end;
+
+      val passiveABs = map2 (curry HOLogic.mk_prodT) passiveAs passiveBs;
+      val zip_ranTs = passiveABs @ prodTsTs';
+      val allJphis = Jphis @ activeJphis;
+      val zipFTs = mk_FTs zip_ranTs;
+      val zipTs = map3 (fn T => fn T' => fn FT => T --> T' --> FT) Ts Ts' zipFTs;
+      val zip_zTs = mk_Ts passiveABs;
+      val (((zips, (abs, abs')), zip_zs), names_lthy) = names_lthy
+        |> mk_Frees "zip" zipTs
+        ||>> mk_Frees' "ab" passiveABs
+        ||>> mk_Frees "z" zip_zTs;
+
+      val Iphi_sets =
+        map2 (fn phi => fn T => HOLogic.Collect_const T $ HOLogic.mk_split phi) allJphis zip_ranTs;
+      val in_phis = map2 (mk_in Iphi_sets) (mk_setss zip_ranTs) zipFTs;
+      val fstABs = map fst_const passiveABs;
+      val all_fsts = fstABs @ fstsTsTs';
+      val map_all_fsts = map2 (fn Ds => fn bnf =>
+        Term.list_comb (mk_map_of_bnf Ds zip_ranTs (passiveAs @ Ts) bnf, all_fsts)) Dss bnfs;
+      val Jmap_fsts = map2 (fn map => fn T => if m = 0 then HOLogic.id_const T
+        else Term.list_comb (map, fstABs)) (mk_Jmaps passiveABs passiveAs) Ts;
+
+      val sndABs = map snd_const passiveABs;
+      val all_snds = sndABs @ sndsTsTs';
+      val map_all_snds = map2 (fn Ds => fn bnf =>
+        Term.list_comb (mk_map_of_bnf Ds zip_ranTs (passiveBs @ Ts') bnf, all_snds)) Dss bnfs;
+      val Jmap_snds = map2 (fn map => fn T => if m = 0 then HOLogic.id_const T
+        else Term.list_comb (map, sndABs)) (mk_Jmaps passiveABs passiveBs) Ts;
+      val zip_unfolds = map (mk_unfold zip_zTs (map HOLogic.mk_split zips)) ks;
+      val zip_setss = mk_Jsetss passiveABs |> transpose;
+
+      val Jrel_coinduct_tac =
+        let
+          fun mk_helper_prem phi in_phi zip x y map map' dtor dtor' =
+            let
+              val zipxy = zip $ x $ y;
+            in
+              HOLogic.mk_Trueprop (list_all_free [x, y]
+                (HOLogic.mk_imp (phi $ x $ y, HOLogic.mk_conj (HOLogic.mk_mem (zipxy, in_phi),
+                  HOLogic.mk_conj (HOLogic.mk_eq (map $ zipxy, dtor $ x),
+                    HOLogic.mk_eq (map' $ zipxy, dtor' $ y))))))
+            end;
+          val helper_prems = map9 mk_helper_prem
+            activeJphis in_phis zips Jzs Jz's map_all_fsts map_all_snds dtors dtor's;
+          fun mk_helper_coind_concl fst phi x alt y map zip_unfold =
+            HOLogic.mk_imp (list_exists_free [if fst then y else x] (HOLogic.mk_conj (phi $ x $ y,
+              HOLogic.mk_eq (alt, map $ (zip_unfold $ HOLogic.mk_prod (x, y))))),
+            HOLogic.mk_eq (alt, if fst then x else y));
+          val helper_coind1_concl =
+            HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+              (map6 (mk_helper_coind_concl true)
+              activeJphis Jzs Jzs_copy Jz's Jmap_fsts zip_unfolds));
+          val helper_coind2_concl =
+            HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+              (map6 (mk_helper_coind_concl false)
+              activeJphis Jzs Jz's_copy Jz's Jmap_snds zip_unfolds));
+          val helper_coind_tac = mk_rel_coinduct_coind_tac m dtor_map_coinduct_thm ks map_comps
+            map_cong0s map_arg_cong_thms set_mapss dtor_unfold_thms dtor_Jmap_thms;
+          fun mk_helper_coind_thms vars concl =
+            Goal.prove_sorry lthy [] []
+              (fold_rev Logic.all (Jphis @ activeJphis @ vars @ zips)
+                (Logic.list_implies (helper_prems, concl)))
+              helper_coind_tac
+            |> Thm.close_derivation
+            |> split_conj_thm;
+          val helper_coind1_thms = mk_helper_coind_thms (Jzs @ Jzs_copy) helper_coind1_concl;
+          val helper_coind2_thms = mk_helper_coind_thms (Jz's @ Jz's_copy) helper_coind2_concl;
+  
+          fun mk_helper_ind_concl phi ab' ab fst snd z active_phi x y zip_unfold set =
+            mk_Ball (set $ z) (Term.absfree ab' (list_all_free [x, y] (HOLogic.mk_imp
+              (HOLogic.mk_conj (active_phi $ x $ y,
+                 HOLogic.mk_eq (z, zip_unfold $ HOLogic.mk_prod (x, y))),
+              phi $ (fst $ ab) $ (snd $ ab)))));
+  
+          val mk_helper_ind_concls =
+            map6 (fn Jphi => fn ab' => fn ab => fn fst => fn snd => fn zip_sets =>
+              map6 (mk_helper_ind_concl Jphi ab' ab fst snd)
+              zip_zs activeJphis Jzs Jz's zip_unfolds zip_sets)
+            Jphis abs' abs fstABs sndABs zip_setss
+            |> map (HOLogic.mk_Trueprop o Library.foldr1 HOLogic.mk_conj);
+  
+          val helper_ind_thmss = if m = 0 then replicate n [] else
+            map3 (fn concl => fn j => fn set_induct =>
+              Goal.prove_sorry lthy [] []
+                (fold_rev Logic.all (Jphis @ activeJphis @ zip_zs @ zips)
+                  (Logic.list_implies (helper_prems, concl)))
+                (mk_rel_coinduct_ind_tac m ks dtor_unfold_thms set_mapss j set_induct)
+              |> Thm.close_derivation
+              |> split_conj_thm)
+            mk_helper_ind_concls ls dtor_Jset_induct_thms
+            |> transpose;
+        in
+          mk_rel_coinduct_tac in_rels in_Jrels
+            helper_ind_thmss helper_coind1_thms helper_coind2_thms
+        end;
+
+      val Jrel_coinduct_thm =
+        mk_rel_xtor_co_induct_thm Greatest_FP rels activeJphis Jrels Jphis Jzs Jz's dtors dtor's
+          Jrel_coinduct_tac lthy;
+
+        val le_Jrel_OO_thm =
+          let
+            fun mk_le_Jrel_OO Jrelpsi1 Jrelpsi2 Jrelpsi12 =
+              mk_leq (mk_rel_compp (Jrelpsi1, Jrelpsi2)) Jrelpsi12;
+            val goals = map3 mk_le_Jrel_OO Jrelpsi1s Jrelpsi2s Jrelpsi12s;
+
+            val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj goals);
+          in
+            singleton (Proof_Context.export names_lthy lthy)
+              (Goal.prove_sorry lthy [] [] goal
+                (K (mk_le_rel_OO_tac Jrel_coinduct_thm dtor_Jrel_thms rel_OOs)))
+              |> Thm.close_derivation
+          end;
+
+        val timer = time (timer "helpers for BNF properties");
+
+        val map_id0_tacs =
+          map2 (K oo mk_map_id0_tac Jmap_thms) dtor_unfold_unique_thms unfold_dtor_thms;
+        val map_comp0_tacs = map (fn thm => K (rtac (thm RS sym) 1)) Jmap_comp0_thms;
+        val map_cong0_tacs = map (mk_map_cong0_tac m) map_cong0_thms;
+        val set_nat_tacss =
+          map2 (map2 (fn def => fn col => fn {context = ctxt, prems = _} =>
+              unfold_thms_tac ctxt Jset_defs THEN mk_set_map0_tac def col))
+            hset_defss (transpose col_natural_thmss);
+
+        val Jbd_card_orders = map (fn def => fold_thms lthy [def] sbd_card_order) Jbd_defs;
+        val Jbd_Cinfinites = map (fn def => fold_thms lthy [def] sbd_Cinfinite) Jbd_defs;
+
+        val bd_co_tacs = map (fn thm => K (rtac thm 1)) Jbd_card_orders;
+        val bd_cinf_tacs = map (fn thm => K (rtac (thm RS conjunct1) 1)) Jbd_Cinfinites;
+
+        val set_bd_tacss =
+          map3 (fn Cinf => map2 (fn def => fn col => fn {context = ctxt, prems = _} =>
+            unfold_thms_tac ctxt Jset_defs THEN mk_set_bd_tac Cinf def col))
+          Jbd_Cinfinites hset_defss (transpose col_bd_thmss);
+
+        val le_rel_OO_tacs = map (fn i => K (rtac (le_Jrel_OO_thm RS mk_conjunctN n i) 1)) ks;
+
+        val rel_OO_Grp_tacs = map (fn def => K (rtac def 1)) Jrel_unabs_defs;
+
+        val tacss = map9 zip_axioms map_id0_tacs map_comp0_tacs map_cong0_tacs set_nat_tacss
+          bd_co_tacs bd_cinf_tacs set_bd_tacss le_rel_OO_tacs rel_OO_Grp_tacs;
+
+        fun wit_tac thms {context = ctxt, prems = _} = unfold_thms_tac ctxt (flat Jwit_defss) THEN
+          mk_wit_tac n dtor_ctor_thms (flat dtor_Jset_thmss) (maps wit_thms_of_bnf bnfs) thms ctxt;
+
+        val (Jbnfs, lthy) =
+          fold_map6 (fn tacs => fn map_b => fn rel_b => fn set_bs => fn Jwit_thms => fn consts =>
+              fn lthy =>
+            bnf_def Do_Inline (user_policy Note_Some) I tacs (wit_tac Jwit_thms) (SOME deads)
+              map_b rel_b set_bs consts lthy
+            |> register_bnf (Local_Theory.full_name lthy b))
+          tacss map_bs rel_bs set_bss Jwit_thmss
+          ((((((bs ~~ Ts) ~~ Jmaps) ~~ Jsetss_by_bnf) ~~ Jbds) ~~ Jwitss) ~~ map SOME Jrels)
+          lthy;
+
+        val timer = time (timer "registered new codatatypes as BNFs");
+
+        val ls' = if m = 1 then [0] else ls;
+
+        val Jbnf_common_notes =
+          [(dtor_map_uniqueN, [dtor_Jmap_unique_thm])] @
+          map2 (fn i => fn thm => (mk_dtor_set_inductN i, [thm])) ls' dtor_Jset_induct_thms
+          |> map (fn (thmN, thms) =>
+            ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]));
+
+        val Jbnf_notes =
+          [(dtor_mapN, map single dtor_Jmap_thms),
+          (dtor_relN, map single dtor_Jrel_thms),
+          (dtor_set_inclN, dtor_Jset_incl_thmss),
+          (dtor_set_set_inclN, map flat dtor_set_Jset_incl_thmsss)] @
+          map2 (fn i => fn thms => (mk_dtor_setN i, map single thms)) ls' dtor_Jset_thmss
+          |> maps (fn (thmN, thmss) =>
+            map2 (fn b => fn thms =>
+              ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]))
+            bs thmss)
+      in
+        (timer, Jbnfs, (Jmap_thms, dtor_Jmap_thms), dtor_Jset_thmss',
+          dtor_Jrel_thms, Jrel_coinduct_thm, Jbnf_common_notes @ Jbnf_notes, lthy)
+      end;
+
+      val dtor_unfold_o_Jmap_thms = mk_xtor_un_fold_o_map_thms Greatest_FP false m
+        dtor_unfold_unique_thm dtor_Jmap_o_thms (map (mk_pointfree lthy) dtor_unfold_thms)
+        sym_map_comps map_cong0s;
+      val dtor_corec_o_Jmap_thms = mk_xtor_un_fold_o_map_thms Greatest_FP true m
+        dtor_corec_unique_thm dtor_Jmap_o_thms (map (mk_pointfree lthy) dtor_corec_thms)
+        sym_map_comps map_cong0s;
+
+      val rels = map2 (fn Ds => mk_rel_of_bnf Ds allAs allBs') Dss bnfs;
+
+      val dtor_unfold_transfer_thms =
+        mk_un_fold_transfer_thms Greatest_FP rels activephis
+          (if m = 0 then map HOLogic.eq_const Ts
+            else map (mk_rel_of_bnf deads passiveAs passiveBs) Jbnfs) Jphis
+          (mk_unfolds passiveAs activeAs) (mk_unfolds passiveBs activeBs)
+          (mk_unfold_transfer_tac m Jrel_coinduct_thm (map map_transfer_of_bnf bnfs)
+            dtor_unfold_thms)
+          lthy;
+
+      val timer = time (timer "relator coinduction");
+
+      val common_notes =
+        [(dtor_coinductN, [dtor_coinduct_thm]),
+        (dtor_map_coinductN, [dtor_map_coinduct_thm]),
+        (rel_coinductN, [Jrel_coinduct_thm]),
+        (dtor_unfold_transferN, dtor_unfold_transfer_thms)]
+        |> map (fn (thmN, thms) =>
+          ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]));
+
+      val notes =
+        [(ctor_dtorN, ctor_dtor_thms),
+        (ctor_exhaustN, ctor_exhaust_thms),
+        (ctor_injectN, ctor_inject_thms),
+        (dtor_corecN, dtor_corec_thms),
+        (dtor_ctorN, dtor_ctor_thms),
+        (dtor_exhaustN, dtor_exhaust_thms),
+        (dtor_injectN, dtor_inject_thms),
+        (dtor_unfoldN, dtor_unfold_thms),
+        (dtor_unfold_uniqueN, dtor_unfold_unique_thms),
+        (dtor_corec_uniqueN, dtor_corec_unique_thms),
+        (dtor_unfold_o_mapN, dtor_unfold_o_Jmap_thms),
+        (dtor_corec_o_mapN, dtor_corec_o_Jmap_thms)]
+        |> map (apsnd (map single))
+        |> maps (fn (thmN, thmss) =>
+          map2 (fn b => fn thms =>
+            ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]))
+          bs thmss);
+
+    (*FIXME: once the package exports all the necessary high-level characteristic theorems,
+       those should not only be concealed but rather not noted at all*)
+    val maybe_conceal_notes = note_all = false ? map (apfst (apfst Binding.conceal));
+  in
+    timer;
+    ({Ts = Ts, bnfs = Jbnfs, ctors = ctors, dtors = dtors,
+      xtor_co_iterss = transpose [unfolds, corecs],
+      xtor_co_induct = dtor_coinduct_thm, dtor_ctors = dtor_ctor_thms,
+      ctor_dtors = ctor_dtor_thms,
+      ctor_injects = ctor_inject_thms, dtor_injects = dtor_inject_thms,
+      xtor_map_thms = dtor_Jmap_thms, xtor_set_thmss = dtor_Jset_thmss',
+      xtor_rel_thms = dtor_Jrel_thms,
+      xtor_co_iter_thmss = transpose [dtor_unfold_thms, dtor_corec_thms],
+      xtor_co_iter_o_map_thmss = transpose [dtor_unfold_o_Jmap_thms, dtor_corec_o_Jmap_thms],
+      rel_xtor_co_induct_thm = Jrel_coinduct_thm},
+     lthy |> Local_Theory.notes (maybe_conceal_notes (common_notes @ notes @ Jbnf_notes)) |> snd)
+  end;
+
+val _ =
+  Outer_Syntax.local_theory @{command_spec "codatatype"} "define coinductive datatypes"
+    (parse_co_datatype_cmd Greatest_FP construct_gfp);
+
+val option_parser = Parse.group (fn () => "option")
+  ((Parse.reserved "sequential" >> K Sequential_Option)
+  || (Parse.reserved "exhaustive" >> K Exhaustive_Option))
+
+val where_alt_specs_of_parser = Parse.where_ |-- Parse.!!! (Parse.enum1 "|"
+  (Parse_Spec.spec -- Scan.option (Parse.reserved "of" |-- Parse.const)));
+
+val _ = Outer_Syntax.local_theory_to_proof @{command_spec "primcorecursive"}
+  "define primitive corecursive functions"
+  ((Scan.optional (@{keyword "("} |--
+      Parse.!!! (Parse.list1 option_parser) --| @{keyword ")"}) []) --
+    (Parse.fixes -- where_alt_specs_of_parser) >> uncurry add_primcorecursive_cmd);
+
+val _ = Outer_Syntax.local_theory @{command_spec "primcorec"}
+  "define primitive corecursive functions"
+  ((Scan.optional (@{keyword "("} |--
+      Parse.!!! (Parse.list1 option_parser) --| @{keyword ")"}) []) --
+    (Parse.fixes -- where_alt_specs_of_parser) >> uncurry add_primcorec_cmd);
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_gfp_rec_sugar.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,1359 @@
+(*  Title:      HOL/BNF/Tools/bnf_gfp_rec_sugar.ML
+    Author:     Lorenz Panny, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2013
+
+Corecursor sugar.
+*)
+
+signature BNF_GFP_REC_SUGAR =
+sig
+  datatype primcorec_option = Sequential_Option | Exhaustive_Option
+
+  val add_primcorecursive_cmd: primcorec_option list ->
+    (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
+    Proof.context -> Proof.state
+  val add_primcorec_cmd: primcorec_option list ->
+    (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
+    local_theory -> local_theory
+end;
+
+structure BNF_GFP_Rec_Sugar : BNF_GFP_REC_SUGAR =
+struct
+
+open Ctr_Sugar_General_Tactics
+open Ctr_Sugar
+open BNF_Util
+open BNF_Def
+open BNF_FP_Util
+open BNF_FP_Def_Sugar
+open BNF_FP_N2M_Sugar
+open BNF_FP_Rec_Sugar_Util
+open BNF_GFP_Rec_Sugar_Tactics
+
+val codeN = "code"
+val ctrN = "ctr"
+val discN = "disc"
+val disc_iffN = "disc_iff"
+val excludeN = "exclude"
+val selN = "sel"
+
+val nitpicksimp_attrs = @{attributes [nitpick_simp]};
+val simp_attrs = @{attributes [simp]};
+val code_nitpicksimp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs;
+
+exception Primcorec_Error of string * term list;
+
+fun primcorec_error str = raise Primcorec_Error (str, []);
+fun primcorec_error_eqn str eqn = raise Primcorec_Error (str, [eqn]);
+fun primcorec_error_eqns str eqns = raise Primcorec_Error (str, eqns);
+
+datatype primcorec_option = Sequential_Option | Exhaustive_Option;
+
+datatype corec_call =
+  Dummy_No_Corec of int |
+  No_Corec of int |
+  Mutual_Corec of int * int * int |
+  Nested_Corec of int;
+
+type basic_corec_ctr_spec =
+  {ctr: term,
+   disc: term,
+   sels: term list};
+
+type corec_ctr_spec =
+  {ctr: term,
+   disc: term,
+   sels: term list,
+   pred: int option,
+   calls: corec_call list,
+   discI: thm,
+   sel_thms: thm list,
+   disc_excludess: thm list list,
+   collapse: thm,
+   corec_thm: thm,
+   disc_corec: thm,
+   sel_corecs: thm list};
+
+type corec_spec =
+  {corec: term,
+   disc_exhausts: thm list,
+   nested_maps: thm list,
+   nested_map_idents: thm list,
+   nested_map_comps: thm list,
+   ctr_specs: corec_ctr_spec list};
+
+exception AINT_NO_MAP of term;
+
+fun not_codatatype ctxt T =
+  error ("Not a codatatype: " ^ Syntax.string_of_typ ctxt T);
+fun ill_formed_corec_call ctxt t =
+  error ("Ill-formed corecursive call: " ^ quote (Syntax.string_of_term ctxt t));
+fun invalid_map ctxt t =
+  error ("Invalid map function in " ^ quote (Syntax.string_of_term ctxt t));
+fun unexpected_corec_call ctxt t =
+  error ("Unexpected corecursive call: " ^ quote (Syntax.string_of_term ctxt t));
+
+fun order_list_duplicates xs = map snd (sort (int_ord o pairself fst) xs)
+
+val mk_conjs = try (foldr1 HOLogic.mk_conj) #> the_default @{const True};
+val mk_disjs = try (foldr1 HOLogic.mk_disj) #> the_default @{const False};
+val mk_dnf = mk_disjs o map mk_conjs;
+
+val conjuncts_s = filter_out (curry (op aconv) @{const True}) o HOLogic.conjuncts;
+
+fun s_not @{const True} = @{const False}
+  | s_not @{const False} = @{const True}
+  | s_not (@{const Not} $ t) = t
+  | s_not (@{const conj} $ t $ u) = @{const disj} $ s_not t $ s_not u
+  | s_not (@{const disj} $ t $ u) = @{const conj} $ s_not t $ s_not u
+  | s_not t = @{const Not} $ t;
+
+val s_not_conj = conjuncts_s o s_not o mk_conjs;
+
+fun propagate_unit_pos u cs = if member (op aconv) cs u then [@{const False}] else cs;
+
+fun propagate_unit_neg not_u cs = remove (op aconv) not_u cs;
+
+fun propagate_units css =
+  (case List.partition (can the_single) css of
+     ([], _) => css
+   | ([u] :: uss, css') =>
+     [u] :: propagate_units (map (propagate_unit_neg (s_not u))
+       (map (propagate_unit_pos u) (uss @ css'))));
+
+fun s_conjs cs =
+  if member (op aconv) cs @{const False} then @{const False}
+  else mk_conjs (remove (op aconv) @{const True} cs);
+
+fun s_disjs ds =
+  if member (op aconv) ds @{const True} then @{const True}
+  else mk_disjs (remove (op aconv) @{const False} ds);
+
+fun s_dnf css0 =
+  let val css = propagate_units css0 in
+    if null css then
+      [@{const False}]
+    else if exists null css then
+      []
+    else
+      map (fn c :: cs => (c, cs)) css
+      |> AList.coalesce (op =)
+      |> map (fn (c, css) => c :: s_dnf css)
+      |> (fn [cs] => cs | css => [s_disjs (map s_conjs css)])
+  end;
+
+fun fold_rev_let_if_case ctxt f bound_Ts t =
+  let
+    val thy = Proof_Context.theory_of ctxt;
+
+    fun fld conds t =
+      (case Term.strip_comb t of
+        (Const (@{const_name Let}, _), [_, _]) => fld conds (unfold_let t)
+      | (Const (@{const_name If}, _), [cond, then_branch, else_branch]) =>
+        fld (conds @ conjuncts_s cond) then_branch o fld (conds @ s_not_conj [cond]) else_branch
+      | (Const (c, _), args as _ :: _ :: _) =>
+        let val n = num_binder_types (Sign.the_const_type thy c) - 1 in
+          if n >= 0 andalso n < length args then
+            (case fastype_of1 (bound_Ts, nth args n) of
+              Type (s, Ts) =>
+              (case dest_case ctxt s Ts t of
+                SOME (ctr_sugar as {sel_splits = _ :: _, ...}, conds', branches) =>
+                apfst (cons ctr_sugar) o fold_rev (uncurry fld)
+                  (map (append conds o conjuncts_s) conds' ~~ branches)
+              | _ => apsnd (f conds t))
+            | _ => apsnd (f conds t))
+          else
+            apsnd (f conds t)
+        end
+      | _ => apsnd (f conds t))
+  in
+    fld [] t o pair []
+  end;
+
+fun case_of ctxt s =
+  (case ctr_sugar_of ctxt s of
+    SOME {casex = Const (s', _), sel_splits = _ :: _, ...} => SOME s'
+  | _ => NONE);
+
+fun massage_let_if_case ctxt has_call massage_leaf =
+  let
+    val thy = Proof_Context.theory_of ctxt;
+
+    fun check_no_call t = if has_call t then unexpected_corec_call ctxt t else ();
+
+    fun massage_abs bound_Ts 0 t = massage_rec bound_Ts t
+      | massage_abs bound_Ts m (Abs (s, T, t)) = Abs (s, T, massage_abs (T :: bound_Ts) (m - 1) t)
+      | massage_abs bound_Ts m t =
+        let val T = domain_type (fastype_of1 (bound_Ts, t)) in
+          Abs (Name.uu, T, massage_abs (T :: bound_Ts) (m - 1) (incr_boundvars 1 t $ Bound 0))
+        end
+    and massage_rec bound_Ts t =
+      let val typof = curry fastype_of1 bound_Ts in
+        (case Term.strip_comb t of
+          (Const (@{const_name Let}, _), [_, _]) => massage_rec bound_Ts (unfold_let t)
+        | (Const (@{const_name If}, _), obj :: (branches as [_, _])) =>
+          let val branches' = map (massage_rec bound_Ts) branches in
+            Term.list_comb (If_const (typof (hd branches')) $ tap check_no_call obj, branches')
+          end
+        | (Const (c, _), args as _ :: _ :: _) =>
+          (case try strip_fun_type (Sign.the_const_type thy c) of
+            SOME (gen_branch_Ts, gen_body_fun_T) =>
+            let
+              val gen_branch_ms = map num_binder_types gen_branch_Ts;
+              val n = length gen_branch_ms;
+            in
+              if n < length args then
+                (case gen_body_fun_T of
+                  Type (_, [Type (T_name, _), _]) =>
+                  if case_of ctxt T_name = SOME c then
+                    let
+                      val (branches, obj_leftovers) = chop n args;
+                      val branches' = map2 (massage_abs bound_Ts) gen_branch_ms branches;
+                      val branch_Ts' = map typof branches';
+                      val body_T' = snd (strip_typeN (hd gen_branch_ms) (hd branch_Ts'));
+                      val casex' = Const (c, branch_Ts' ---> map typof obj_leftovers ---> body_T');
+                    in
+                      Term.list_comb (casex',
+                        branches' @ tap (List.app check_no_call) obj_leftovers)
+                    end
+                  else
+                    massage_leaf bound_Ts t
+                | _ => massage_leaf bound_Ts t)
+              else
+                massage_leaf bound_Ts t
+            end
+          | NONE => massage_leaf bound_Ts t)
+        | _ => massage_leaf bound_Ts t)
+      end
+  in
+    massage_rec
+  end;
+
+val massage_mutual_corec_call = massage_let_if_case;
+
+fun curried_type (Type (@{type_name fun}, [Type (@{type_name prod}, Ts), T])) = Ts ---> T;
+
+fun massage_nested_corec_call ctxt has_call raw_massage_call bound_Ts U t =
+  let
+    fun check_no_call t = if has_call t then unexpected_corec_call ctxt t else ();
+
+    val build_map_Inl = build_map ctxt (uncurry Inl_const o dest_sumT o snd);
+
+    fun massage_mutual_call bound_Ts U T t =
+      if has_call t then
+        (case try dest_sumT U of
+          SOME (U1, U2) => if U1 = T then raw_massage_call bound_Ts T U2 t else invalid_map ctxt t
+        | NONE => invalid_map ctxt t)
+      else
+        build_map_Inl (T, U) $ t;
+
+    fun massage_mutual_fun bound_Ts U T t =
+      (case t of
+        Const (@{const_name comp}, _) $ t1 $ t2 =>
+        mk_comp bound_Ts (massage_mutual_fun bound_Ts U T t1, tap check_no_call t2)
+      | _ =>
+        let
+          val var = Var ((Name.uu, Term.maxidx_of_term t + 1),
+            domain_type (fastype_of1 (bound_Ts, t)));
+        in
+          Term.lambda var (massage_mutual_call bound_Ts U T (t $ var))
+        end);
+
+    fun massage_map bound_Ts (Type (_, Us)) (Type (s, Ts)) t =
+        (case try (dest_map ctxt s) t of
+          SOME (map0, fs) =>
+          let
+            val Type (_, dom_Ts) = domain_type (fastype_of1 (bound_Ts, t));
+            val map' = mk_map (length fs) dom_Ts Us map0;
+            val fs' =
+              map_flattened_map_args ctxt s (map3 (massage_map_or_map_arg bound_Ts) Us Ts) fs;
+          in
+            Term.list_comb (map', fs')
+          end
+        | NONE => raise AINT_NO_MAP t)
+      | massage_map _ _ _ t = raise AINT_NO_MAP t
+    and massage_map_or_map_arg bound_Ts U T t =
+      if T = U then
+        tap check_no_call t
+      else
+        massage_map bound_Ts U T t
+        handle AINT_NO_MAP _ => massage_mutual_fun bound_Ts U T t;
+
+    fun massage_call bound_Ts U T =
+      massage_let_if_case ctxt has_call (fn bound_Ts => fn t =>
+        if has_call t then
+          (case U of
+            Type (s, Us) =>
+            (case try (dest_ctr ctxt s) t of
+              SOME (f, args) =>
+              let
+                val typof = curry fastype_of1 bound_Ts;
+                val f' = mk_ctr Us f
+                val f'_T = typof f';
+                val arg_Ts = map typof args;
+              in
+                Term.list_comb (f', map3 (massage_call bound_Ts) (binder_types f'_T) arg_Ts args)
+              end
+            | NONE =>
+              (case t of
+                Const (@{const_name prod_case}, _) $ t' =>
+                let
+                  val U' = curried_type U;
+                  val T' = curried_type T;
+                in
+                  Const (@{const_name prod_case}, U' --> U) $ massage_call bound_Ts U' T' t'
+                end
+              | t1 $ t2 =>
+                (if has_call t2 then
+                  massage_mutual_call bound_Ts U T t
+                else
+                  massage_map bound_Ts U T t1 $ t2
+                  handle AINT_NO_MAP _ => massage_mutual_call bound_Ts U T t)
+              | Abs (s, T', t') =>
+                Abs (s, T', massage_call (T' :: bound_Ts) (range_type U) (range_type T) t')
+              | _ => massage_mutual_call bound_Ts U T t))
+          | _ => ill_formed_corec_call ctxt t)
+        else
+          build_map_Inl (T, U) $ t) bound_Ts;
+
+    val T = fastype_of1 (bound_Ts, t);
+  in
+    if has_call t then massage_call bound_Ts U T t else build_map_Inl (T, U) $ t
+  end;
+
+fun expand_to_ctr_term ctxt s Ts t =
+  (case ctr_sugar_of ctxt s of
+    SOME {ctrs, casex, ...} =>
+    Term.list_comb (mk_case Ts (Type (s, Ts)) casex, map (mk_ctr Ts) ctrs) $ t
+  | NONE => raise Fail "expand_to_ctr_term");
+
+fun expand_corec_code_rhs ctxt has_call bound_Ts t =
+  (case fastype_of1 (bound_Ts, t) of
+    Type (s, Ts) =>
+    massage_let_if_case ctxt has_call (fn _ => fn t =>
+      if can (dest_ctr ctxt s) t then t else expand_to_ctr_term ctxt s Ts t) bound_Ts t
+  | _ => raise Fail "expand_corec_code_rhs");
+
+fun massage_corec_code_rhs ctxt massage_ctr =
+  massage_let_if_case ctxt (K false)
+    (fn bound_Ts => uncurry (massage_ctr bound_Ts) o Term.strip_comb);
+
+fun fold_rev_corec_code_rhs ctxt f =
+  snd ooo fold_rev_let_if_case ctxt (fn conds => uncurry (f conds) o Term.strip_comb);
+
+fun case_thms_of_term ctxt bound_Ts t =
+  let val (ctr_sugars, _) = fold_rev_let_if_case ctxt (K (K I)) bound_Ts t () in
+    (maps #distincts ctr_sugars, maps #discIs ctr_sugars, maps #disc_exhausts ctr_sugars,
+     maps #sel_splits ctr_sugars, maps #sel_split_asms ctr_sugars)
+  end;
+
+fun basic_corec_specs_of ctxt res_T =
+  (case res_T of
+    Type (T_name, _) =>
+    (case Ctr_Sugar.ctr_sugar_of ctxt T_name of
+      NONE => not_codatatype ctxt res_T
+    | SOME {ctrs, discs, selss, ...} =>
+      let
+        val thy = Proof_Context.theory_of ctxt;
+
+        val gfpT = body_type (fastype_of (hd ctrs));
+        val As_rho = tvar_subst thy [gfpT] [res_T];
+        val substA = Term.subst_TVars As_rho;
+
+        fun mk_spec ctr disc sels = {ctr = substA ctr, disc = substA disc, sels = map substA sels};
+      in
+        map3 mk_spec ctrs discs selss
+        handle ListPair.UnequalLengths => not_codatatype ctxt res_T
+      end)
+  | _ => not_codatatype ctxt res_T);
+
+fun map_thms_of_typ ctxt (Type (s, _)) =
+    (case fp_sugar_of ctxt s of
+      SOME {index, mapss, ...} => nth mapss index
+    | NONE => [])
+  | map_thms_of_typ _ _ = [];
+
+fun corec_specs_of bs arg_Ts res_Ts get_indices callssss0 lthy0 =
+  let
+    val thy = Proof_Context.theory_of lthy0;
+
+    val ((missing_res_Ts, perm0_kks,
+          fp_sugars as {nested_bnfs, fp_res = {xtor_co_iterss = dtor_coiters1 :: _, ...},
+            co_inducts = coinduct_thms, ...} :: _, (_, gfp_sugar_thms)), lthy) =
+      nested_to_mutual_fps Greatest_FP bs res_Ts get_indices callssss0 lthy0;
+
+    val perm_fp_sugars = sort (int_ord o pairself #index) fp_sugars;
+
+    val indices = map #index fp_sugars;
+    val perm_indices = map #index perm_fp_sugars;
+
+    val perm_ctrss = map (#ctrs o of_fp_sugar #ctr_sugars) perm_fp_sugars;
+    val perm_ctr_Tsss = map (map (binder_types o fastype_of)) perm_ctrss;
+    val perm_gfpTs = map (body_type o fastype_of o hd) perm_ctrss;
+
+    val nn0 = length res_Ts;
+    val nn = length perm_gfpTs;
+    val kks = 0 upto nn - 1;
+    val perm_ns = map length perm_ctr_Tsss;
+
+    val perm_Cs = map (domain_type o body_fun_type o fastype_of o co_rec_of o
+      of_fp_sugar (#xtor_co_iterss o #fp_res)) perm_fp_sugars;
+    val (perm_p_Tss, (perm_q_Tssss, _, perm_f_Tssss, _)) =
+      mk_coiter_fun_arg_types perm_ctr_Tsss perm_Cs perm_ns (co_rec_of dtor_coiters1);
+
+    val (perm_p_hss, h) = indexedd perm_p_Tss 0;
+    val (perm_q_hssss, h') = indexedddd perm_q_Tssss h;
+    val (perm_f_hssss, _) = indexedddd perm_f_Tssss h';
+
+    val fun_arg_hs =
+      flat (map3 flat_corec_preds_predsss_gettersss perm_p_hss perm_q_hssss perm_f_hssss);
+
+    fun unpermute0 perm0_xs = permute_like (op =) perm0_kks kks perm0_xs;
+    fun unpermute perm_xs = permute_like (op =) perm_indices indices perm_xs;
+
+    val coinduct_thmss = map (unpermute0 o conj_dests nn) coinduct_thms;
+
+    val p_iss = map (map (find_index_eq fun_arg_hs)) (unpermute perm_p_hss);
+    val q_issss = map (map (map (map (find_index_eq fun_arg_hs)))) (unpermute perm_q_hssss);
+    val f_issss = map (map (map (map (find_index_eq fun_arg_hs)))) (unpermute perm_f_hssss);
+
+    val f_Tssss = unpermute perm_f_Tssss;
+    val gfpTs = unpermute perm_gfpTs;
+    val Cs = unpermute perm_Cs;
+
+    val As_rho = tvar_subst thy (take nn0 gfpTs) res_Ts;
+    val Cs_rho = map (fst o dest_TVar) Cs ~~ pad_list HOLogic.unitT nn arg_Ts;
+
+    val substA = Term.subst_TVars As_rho;
+    val substAT = Term.typ_subst_TVars As_rho;
+    val substCT = Term.typ_subst_TVars Cs_rho;
+
+    val perm_Cs' = map substCT perm_Cs;
+
+    fun call_of nullary [] [g_i] [Type (@{type_name fun}, [_, T])] =
+        (if exists_subtype_in Cs T then Nested_Corec
+         else if nullary then Dummy_No_Corec
+         else No_Corec) g_i
+      | call_of _ [q_i] [g_i, g_i'] _ = Mutual_Corec (q_i, g_i, g_i');
+
+    fun mk_ctr_spec ctr disc sels p_io q_iss f_iss f_Tss discI sel_thms disc_excludess collapse
+        corec_thm disc_corec sel_corecs =
+      let val nullary = not (can dest_funT (fastype_of ctr)) in
+        {ctr = substA ctr, disc = substA disc, sels = map substA sels, pred = p_io,
+         calls = map3 (call_of nullary) q_iss f_iss f_Tss, discI = discI, sel_thms = sel_thms,
+         disc_excludess = disc_excludess, collapse = collapse, corec_thm = corec_thm,
+         disc_corec = disc_corec, sel_corecs = sel_corecs}
+      end;
+
+    fun mk_ctr_specs index ctr_sugars p_is q_isss f_isss f_Tsss coiter_thmsss disc_coitersss
+        sel_coiterssss =
+      let
+        val {ctrs, discs, selss, discIs, sel_thmss, disc_excludesss, collapses, ...} : ctr_sugar =
+          nth ctr_sugars index;
+        val p_ios = map SOME p_is @ [NONE];
+        val discIs = #discIs (nth ctr_sugars index);
+        val corec_thms = co_rec_of (nth coiter_thmsss index);
+        val disc_corecs = co_rec_of (nth disc_coitersss index);
+        val sel_corecss = co_rec_of (nth sel_coiterssss index);
+      in
+        map14 mk_ctr_spec ctrs discs selss p_ios q_isss f_isss f_Tsss discIs sel_thmss
+          disc_excludesss collapses corec_thms disc_corecs sel_corecss
+      end;
+
+    fun mk_spec ({T, index, ctr_sugars, co_iterss = coiterss, co_iter_thmsss = coiter_thmsss,
+          disc_co_itersss = disc_coitersss, sel_co_iterssss = sel_coiterssss, ...} : fp_sugar)
+        p_is q_isss f_isss f_Tsss =
+      {corec = mk_co_iter thy Greatest_FP (substAT T) perm_Cs' (co_rec_of (nth coiterss index)),
+       disc_exhausts = #disc_exhausts (nth ctr_sugars index),
+       nested_maps = maps (map_thms_of_typ lthy o T_of_bnf) nested_bnfs,
+       nested_map_idents = map (unfold_thms lthy @{thms id_def} o map_id0_of_bnf) nested_bnfs,
+       nested_map_comps = map map_comp_of_bnf nested_bnfs,
+       ctr_specs = mk_ctr_specs index ctr_sugars p_is q_isss f_isss f_Tsss coiter_thmsss
+         disc_coitersss sel_coiterssss};
+  in
+    ((is_some gfp_sugar_thms, map5 mk_spec fp_sugars p_iss q_issss f_issss f_Tssss, missing_res_Ts,
+      co_induct_of coinduct_thms, strong_co_induct_of coinduct_thms, co_induct_of coinduct_thmss,
+      strong_co_induct_of coinduct_thmss), lthy)
+  end;
+
+val undef_const = Const (@{const_name undefined}, dummyT);
+
+val abs_tuple = HOLogic.tupled_lambda o HOLogic.mk_tuple;
+
+fun abstract vs =
+  let fun a n (t $ u) = a n t $ a n u
+        | a n (Abs (v, T, b)) = Abs (v, T, a (n + 1) b)
+        | a n t = let val idx = find_index (curry (op =) t) vs in
+            if idx < 0 then t else Bound (n + idx) end
+  in a 0 end;
+
+fun mk_prod1 bound_Ts (t, u) =
+  HOLogic.pair_const (fastype_of1 (bound_Ts, t)) (fastype_of1 (bound_Ts, u)) $ t $ u;
+fun mk_tuple1 bound_Ts = the_default HOLogic.unit o try (foldr1 (mk_prod1 bound_Ts));
+
+type coeqn_data_disc = {
+  fun_name: string,
+  fun_T: typ,
+  fun_args: term list,
+  ctr: term,
+  ctr_no: int,
+  disc: term,
+  prems: term list,
+  auto_gen: bool,
+  ctr_rhs_opt: term option,
+  code_rhs_opt: term option,
+  eqn_pos: int,
+  user_eqn: term
+};
+
+type coeqn_data_sel = {
+  fun_name: string,
+  fun_T: typ,
+  fun_args: term list,
+  ctr: term,
+  sel: term,
+  rhs_term: term,
+  ctr_rhs_opt: term option,
+  code_rhs_opt: term option,
+  eqn_pos: int,
+  user_eqn: term
+};
+
+datatype coeqn_data =
+  Disc of coeqn_data_disc |
+  Sel of coeqn_data_sel;
+
+fun dissect_coeqn_disc fun_names sequentials (basic_ctr_specss : basic_corec_ctr_spec list list)
+    eqn_pos ctr_rhs_opt code_rhs_opt prems' concl matchedsss =
+  let
+    fun find_subterm p =
+      let (* FIXME \<exists>? *)
+        fun find (t as u $ v) = if p t then SOME t else merge_options (find u, find v)
+          | find t = if p t then SOME t else NONE;
+      in find end;
+
+    val applied_fun = concl
+      |> find_subterm (member (op = o apsnd SOME) fun_names o try (fst o dest_Free o head_of))
+      |> the
+      handle Option.Option => primcorec_error_eqn "malformed discriminator formula" concl;
+    val ((fun_name, fun_T), fun_args) = strip_comb applied_fun |>> dest_Free;
+    val SOME (sequential, basic_ctr_specs) =
+      AList.lookup (op =) (fun_names ~~ (sequentials ~~ basic_ctr_specss)) fun_name;
+
+    val discs = map #disc basic_ctr_specs;
+    val ctrs = map #ctr basic_ctr_specs;
+    val not_disc = head_of concl = @{term Not};
+    val _ = not_disc andalso length ctrs <> 2 andalso
+      primcorec_error_eqn "negated discriminator for a type with \<noteq> 2 constructors" concl;
+    val disc' = find_subterm (member (op =) discs o head_of) concl;
+    val eq_ctr0 = concl |> perhaps (try HOLogic.dest_not) |> try (HOLogic.dest_eq #> snd)
+        |> (fn SOME t => let val n = find_index (curry (op =) t) ctrs in
+          if n >= 0 then SOME n else NONE end | _ => NONE);
+    val _ = is_some disc' orelse is_some eq_ctr0 orelse
+      primcorec_error_eqn "no discriminator in equation" concl;
+    val ctr_no' =
+      if is_none disc' then the eq_ctr0 else find_index (curry (op =) (head_of (the disc'))) discs;
+    val ctr_no = if not_disc then 1 - ctr_no' else ctr_no';
+    val {ctr, disc, ...} = nth basic_ctr_specs ctr_no;
+
+    val catch_all = try (fst o dest_Free o the_single) prems' = SOME Name.uu_;
+    val matchedss = AList.lookup (op =) matchedsss fun_name |> the_default [];
+    val prems = map (abstract (List.rev fun_args)) prems';
+    val actual_prems =
+      (if catch_all orelse sequential then maps s_not_conj matchedss else []) @
+      (if catch_all then [] else prems);
+
+    val matchedsss' = AList.delete (op =) fun_name matchedsss
+      |> cons (fun_name, if sequential then matchedss @ [prems] else matchedss @ [actual_prems]);
+
+    val user_eqn =
+      (actual_prems, concl)
+      |>> map HOLogic.mk_Trueprop ||> HOLogic.mk_Trueprop o abstract (List.rev fun_args)
+      |> curry Logic.list_all (map dest_Free fun_args) o Logic.list_implies;
+  in
+    (Disc {
+      fun_name = fun_name,
+      fun_T = fun_T,
+      fun_args = fun_args,
+      ctr = ctr,
+      ctr_no = ctr_no,
+      disc = disc,
+      prems = actual_prems,
+      auto_gen = catch_all,
+      ctr_rhs_opt = ctr_rhs_opt,
+      code_rhs_opt = code_rhs_opt,
+      eqn_pos = eqn_pos,
+      user_eqn = user_eqn
+    }, matchedsss')
+  end;
+
+fun dissect_coeqn_sel fun_names (basic_ctr_specss : basic_corec_ctr_spec list list) eqn_pos
+    ctr_rhs_opt code_rhs_opt eqn0 of_spec_opt eqn =
+  let
+    val (lhs, rhs) = HOLogic.dest_eq eqn
+      handle TERM _ =>
+        primcorec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn;
+    val sel = head_of lhs;
+    val ((fun_name, fun_T), fun_args) = dest_comb lhs |> snd |> strip_comb |> apfst dest_Free
+      handle TERM _ =>
+        primcorec_error_eqn "malformed selector argument in left-hand side" eqn;
+    val basic_ctr_specs = the (AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name)
+      handle Option.Option =>
+        primcorec_error_eqn "malformed selector argument in left-hand side" eqn;
+    val {ctr, ...} =
+      (case of_spec_opt of
+        SOME of_spec => the (find_first (curry (op =) of_spec o #ctr) basic_ctr_specs)
+      | NONE => filter (exists (curry (op =) sel) o #sels) basic_ctr_specs |> the_single
+          handle List.Empty => primcorec_error_eqn "ambiguous selector - use \"of\"" eqn);
+    val user_eqn = drop_all eqn0;
+  in
+    Sel {
+      fun_name = fun_name,
+      fun_T = fun_T,
+      fun_args = fun_args,
+      ctr = ctr,
+      sel = sel,
+      rhs_term = rhs,
+      ctr_rhs_opt = ctr_rhs_opt,
+      code_rhs_opt = code_rhs_opt,
+      eqn_pos = eqn_pos,
+      user_eqn = user_eqn
+    }
+  end;
+
+fun dissect_coeqn_ctr fun_names sequentials (basic_ctr_specss : basic_corec_ctr_spec list list)
+    eqn_pos eqn0 code_rhs_opt prems concl matchedsss =
+  let
+    val (lhs, rhs) = HOLogic.dest_eq concl;
+    val (fun_name, fun_args) = strip_comb lhs |>> fst o dest_Free;
+    val SOME basic_ctr_specs = AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name;
+    val (ctr, ctr_args) = strip_comb (unfold_let rhs);
+    val {disc, sels, ...} = the (find_first (curry (op =) ctr o #ctr) basic_ctr_specs)
+      handle Option.Option => primcorec_error_eqn "not a constructor" ctr;
+
+    val disc_concl = betapply (disc, lhs);
+    val (eqn_data_disc_opt, matchedsss') =
+      if null (tl basic_ctr_specs) then
+        (NONE, matchedsss)
+      else
+        apfst SOME (dissect_coeqn_disc fun_names sequentials basic_ctr_specss eqn_pos
+          (SOME (abstract (List.rev fun_args) rhs)) code_rhs_opt prems disc_concl matchedsss);
+
+    val sel_concls = sels ~~ ctr_args
+      |> map (fn (sel, ctr_arg) => HOLogic.mk_eq (betapply (sel, lhs), ctr_arg));
+
+(*
+val _ = tracing ("reduced\n    " ^ Syntax.string_of_term @{context} concl ^ "\nto\n    \<cdot> " ^
+ (is_some eqn_data_disc_opt ? K (Syntax.string_of_term @{context} disc_concl ^ "\n    \<cdot> ")) "" ^
+ space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) sel_concls) ^
+ "\nfor premise(s)\n    \<cdot> " ^
+ space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) prems));
+*)
+
+    val eqns_data_sel =
+      map (dissect_coeqn_sel fun_names basic_ctr_specss eqn_pos
+        (SOME (abstract (List.rev fun_args) rhs)) code_rhs_opt eqn0 (SOME ctr)) sel_concls;
+  in
+    (the_list eqn_data_disc_opt @ eqns_data_sel, matchedsss')
+  end;
+
+fun dissect_coeqn_code lthy has_call fun_names basic_ctr_specss eqn_pos eqn0 concl matchedsss =
+  let
+    val (lhs, (rhs', rhs)) = HOLogic.dest_eq concl ||> `(expand_corec_code_rhs lthy has_call []);
+    val (fun_name, fun_args) = strip_comb lhs |>> fst o dest_Free;
+    val SOME basic_ctr_specs = AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name;
+
+    val cond_ctrs = fold_rev_corec_code_rhs lthy (fn cs => fn ctr => fn _ =>
+        if member (op = o apsnd #ctr) basic_ctr_specs ctr then cons (ctr, cs)
+        else primcorec_error_eqn "not a constructor" ctr) [] rhs' []
+      |> AList.group (op =);
+
+    val ctr_premss = (case cond_ctrs of [_] => [[]] | _ => map (s_dnf o snd) cond_ctrs);
+    val ctr_concls = cond_ctrs |> map (fn (ctr, _) =>
+        binder_types (fastype_of ctr)
+        |> map_index (fn (n, T) => massage_corec_code_rhs lthy (fn _ => fn ctr' => fn args =>
+          if ctr' = ctr then nth args n else Const (@{const_name undefined}, T)) [] rhs')
+        |> curry list_comb ctr
+        |> curry HOLogic.mk_eq lhs);
+
+    val sequentials = replicate (length fun_names) false;
+  in
+    fold_map2 (dissect_coeqn_ctr fun_names sequentials basic_ctr_specss eqn_pos eqn0
+        (SOME (abstract (List.rev fun_args) rhs)))
+      ctr_premss ctr_concls matchedsss
+  end;
+
+fun dissect_coeqn lthy has_call fun_names sequentials
+    (basic_ctr_specss : basic_corec_ctr_spec list list) (eqn_pos, eqn0) of_spec_opt matchedsss =
+  let
+    val eqn = drop_all eqn0
+      handle TERM _ => primcorec_error_eqn "malformed function equation" eqn0;
+    val (prems, concl) = Logic.strip_horn eqn
+      |> apfst (map HOLogic.dest_Trueprop) o apsnd HOLogic.dest_Trueprop;
+
+    val head = concl
+      |> perhaps (try HOLogic.dest_not) |> perhaps (try (fst o HOLogic.dest_eq))
+      |> head_of;
+
+    val rhs_opt = concl |> perhaps (try HOLogic.dest_not) |> try (snd o HOLogic.dest_eq);
+
+    val discs = maps (map #disc) basic_ctr_specss;
+    val sels = maps (maps #sels) basic_ctr_specss;
+    val ctrs = maps (map #ctr) basic_ctr_specss;
+  in
+    if member (op =) discs head orelse
+        is_some rhs_opt andalso
+          member (op =) (filter (null o binder_types o fastype_of) ctrs) (the rhs_opt) then
+      dissect_coeqn_disc fun_names sequentials basic_ctr_specss eqn_pos NONE NONE prems concl
+        matchedsss
+      |>> single
+    else if member (op =) sels head then
+      ([dissect_coeqn_sel fun_names basic_ctr_specss eqn_pos NONE NONE eqn0 of_spec_opt concl],
+       matchedsss)
+    else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) then
+      if member (op =) ctrs (head_of (unfold_let (the rhs_opt))) then
+        dissect_coeqn_ctr fun_names sequentials basic_ctr_specss eqn_pos eqn0
+          (if null prems then
+             SOME (snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (Logic.strip_assums_concl eqn0))))
+           else
+             NONE)
+          prems concl matchedsss
+      else if null prems then
+        dissect_coeqn_code lthy has_call fun_names basic_ctr_specss eqn_pos eqn0 concl matchedsss
+        |>> flat
+      else
+        primcorec_error_eqn "cannot mix constructor and code views (see manual for details)" eqn
+    else
+      primcorec_error_eqn "malformed function equation" eqn
+  end;
+
+fun build_corec_arg_disc (ctr_specs : corec_ctr_spec list)
+    ({fun_args, ctr_no, prems, ...} : coeqn_data_disc) =
+  if is_none (#pred (nth ctr_specs ctr_no)) then I else
+    s_conjs prems
+    |> curry subst_bounds (List.rev fun_args)
+    |> HOLogic.tupled_lambda (HOLogic.mk_tuple fun_args)
+    |> K |> nth_map (the (#pred (nth ctr_specs ctr_no)));
+
+fun build_corec_arg_no_call (sel_eqns : coeqn_data_sel list) sel =
+  find_first (curry (op =) sel o #sel) sel_eqns
+  |> try (fn SOME {fun_args, rhs_term, ...} => abs_tuple fun_args rhs_term)
+  |> the_default undef_const
+  |> K;
+
+fun build_corec_args_mutual_call lthy has_call (sel_eqns : coeqn_data_sel list) sel =
+  (case find_first (curry (op =) sel o #sel) sel_eqns of
+    NONE => (I, I, I)
+  | SOME {fun_args, rhs_term, ... } =>
+    let
+      val bound_Ts = List.rev (map fastype_of fun_args);
+      fun rewrite_stop _ t = if has_call t then @{term False} else @{term True};
+      fun rewrite_end _ t = if has_call t then undef_const else t;
+      fun rewrite_cont bound_Ts t =
+        if has_call t then mk_tuple1 bound_Ts (snd (strip_comb t)) else undef_const;
+      fun massage f _ = massage_mutual_corec_call lthy has_call f bound_Ts rhs_term
+        |> abs_tuple fun_args;
+    in
+      (massage rewrite_stop, massage rewrite_end, massage rewrite_cont)
+    end);
+
+fun build_corec_arg_nested_call lthy has_call (sel_eqns : coeqn_data_sel list) sel =
+  (case find_first (curry (op =) sel o #sel) sel_eqns of
+    NONE => I
+  | SOME {fun_args, rhs_term, ...} =>
+    let
+      val bound_Ts = List.rev (map fastype_of fun_args);
+      fun rewrite bound_Ts U T (Abs (v, V, b)) = Abs (v, V, rewrite (V :: bound_Ts) U T b)
+        | rewrite bound_Ts U T (t as _ $ _) =
+          let val (u, vs) = strip_comb t in
+            if is_Free u andalso has_call u then
+              Inr_const U T $ mk_tuple1 bound_Ts vs
+            else if try (fst o dest_Const) u = SOME @{const_name prod_case} then
+              map (rewrite bound_Ts U T) vs |> chop 1 |>> HOLogic.mk_split o the_single |> list_comb
+            else
+              list_comb (rewrite bound_Ts U T u, map (rewrite bound_Ts U T) vs)
+          end
+        | rewrite _ U T t =
+          if is_Free t andalso has_call t then Inr_const U T $ HOLogic.unit else t;
+      fun massage t =
+        rhs_term
+        |> massage_nested_corec_call lthy has_call rewrite bound_Ts (range_type (fastype_of t))
+        |> abs_tuple fun_args;
+    in
+      massage
+    end);
+
+fun build_corec_args_sel lthy has_call (all_sel_eqns : coeqn_data_sel list)
+    (ctr_spec : corec_ctr_spec) =
+  (case filter (curry (op =) (#ctr ctr_spec) o #ctr) all_sel_eqns of
+    [] => I
+  | sel_eqns =>
+    let
+      val sel_call_list = #sels ctr_spec ~~ #calls ctr_spec;
+      val no_calls' = map_filter (try (apsnd (fn No_Corec n => n))) sel_call_list;
+      val mutual_calls' = map_filter (try (apsnd (fn Mutual_Corec n => n))) sel_call_list;
+      val nested_calls' = map_filter (try (apsnd (fn Nested_Corec n => n))) sel_call_list;
+    in
+      I
+      #> fold (fn (sel, n) => nth_map n (build_corec_arg_no_call sel_eqns sel)) no_calls'
+      #> fold (fn (sel, (q, g, h)) =>
+        let val (fq, fg, fh) = build_corec_args_mutual_call lthy has_call sel_eqns sel in
+          nth_map q fq o nth_map g fg o nth_map h fh end) mutual_calls'
+      #> fold (fn (sel, n) => nth_map n
+        (build_corec_arg_nested_call lthy has_call sel_eqns sel)) nested_calls'
+    end);
+
+fun build_codefs lthy bs mxs has_call arg_Tss (corec_specs : corec_spec list)
+    (disc_eqnss : coeqn_data_disc list list) (sel_eqnss : coeqn_data_sel list list) =
+  let
+    val corecs = map #corec corec_specs;
+    val ctr_specss = map #ctr_specs corec_specs;
+    val corec_args = hd corecs
+      |> fst o split_last o binder_types o fastype_of
+      |> map (fn T => if range_type T = @{typ bool}
+          then Abs (Name.uu_, domain_type T, @{term False})
+          else Const (@{const_name undefined}, T))
+      |> fold2 (fold o build_corec_arg_disc) ctr_specss disc_eqnss
+      |> fold2 (fold o build_corec_args_sel lthy has_call) sel_eqnss ctr_specss;
+    fun currys [] t = t
+      | currys Ts t = t $ mk_tuple1 (List.rev Ts) (map Bound (length Ts - 1 downto 0))
+          |> fold_rev (Term.abs o pair Name.uu) Ts;
+
+(*
+val _ = tracing ("corecursor arguments:\n    \<cdot> " ^
+ space_implode "\n    \<cdot> " (map (Syntax.string_of_term lthy) corec_args));
+*)
+
+    val excludess' =
+      disc_eqnss
+      |> map (map (fn x => (#fun_args x, #ctr_no x, #prems x, #auto_gen x))
+        #> fst o (fn xs => fold_map (fn x => fn ys => ((x, ys), ys @ [x])) xs [])
+        #> maps (uncurry (map o pair)
+          #> map (fn ((fun_args, c, x, a), (_, c', y, a')) =>
+              ((c, c', a orelse a'), (x, s_not (s_conjs y)))
+            ||> apfst (map HOLogic.mk_Trueprop) o apsnd HOLogic.mk_Trueprop
+            ||> Logic.list_implies
+            ||> curry Logic.list_all (map dest_Free fun_args))))
+  in
+    map (list_comb o rpair corec_args) corecs
+    |> map2 (fn Ts => fn t => if length Ts = 0 then t $ HOLogic.unit else t) arg_Tss
+    |> map2 currys arg_Tss
+    |> Syntax.check_terms lthy
+    |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.conceal (Thm.def_binding b), []), t)))
+      bs mxs
+    |> rpair excludess'
+  end;
+
+fun mk_actual_disc_eqns fun_binding arg_Ts exhaustive ({ctr_specs, ...} : corec_spec)
+    (sel_eqns : coeqn_data_sel list) (disc_eqns : coeqn_data_disc list) =
+  let val num_disc_eqns = length disc_eqns in
+    if (exhaustive andalso num_disc_eqns <> 0) orelse num_disc_eqns <> length ctr_specs - 1 then
+      disc_eqns
+    else
+      let
+        val n = 0 upto length ctr_specs
+          |> the o find_first (fn idx => not (exists (curry (op =) idx o #ctr_no) disc_eqns));
+        val {ctr, disc, ...} = nth ctr_specs n;
+        val fun_args = (try (#fun_args o hd) disc_eqns, try (#fun_args o hd) sel_eqns)
+          |> the_default (map (curry Free Name.uu) arg_Ts) o merge_options;
+        val sel_eqn_opt = find_first (equal ctr o #ctr) sel_eqns;
+        val extra_disc_eqn = {
+          fun_name = Binding.name_of fun_binding,
+          fun_T = arg_Ts ---> body_type (fastype_of (#ctr (hd ctr_specs))),
+          fun_args = fun_args,
+          ctr = ctr,
+          ctr_no = n,
+          disc = disc,
+          prems = maps (s_not_conj o #prems) disc_eqns,
+          auto_gen = true,
+          ctr_rhs_opt = Option.map #ctr_rhs_opt sel_eqn_opt |> the_default NONE,
+          code_rhs_opt = Option.map #ctr_rhs_opt sel_eqn_opt |> the_default NONE,
+          eqn_pos = Option.map (curry (op +) 1 o #eqn_pos) sel_eqn_opt |> the_default 100000 (* FIXME *),
+          user_eqn = undef_const};
+      in
+        chop n disc_eqns ||> cons extra_disc_eqn |> (op @)
+      end
+  end;
+
+fun find_corec_calls ctxt has_call basic_ctr_specs ({ctr, sel, rhs_term, ...} : coeqn_data_sel) =
+  let
+    val sel_no = find_first (curry (op =) ctr o #ctr) basic_ctr_specs
+      |> find_index (curry (op =) sel) o #sels o the;
+    fun find t = if has_call t then snd (fold_rev_let_if_case ctxt (K cons) [] t []) else [];
+  in
+    find rhs_term
+    |> K |> nth_map sel_no |> AList.map_entry (op =) ctr
+  end;
+
+fun applied_fun_of fun_name fun_T fun_args =
+  list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0));
+
+fun is_trivial_implies thm =
+  uncurry (member (op aconv)) (Logic.strip_horn (Thm.prop_of thm));
+
+fun add_primcorec_ursive auto opts fixes specs of_specs_opt lthy =
+  let
+    val thy = Proof_Context.theory_of lthy;
+
+    val (bs, mxs) = map_split (apfst fst) fixes;
+    val (arg_Ts, res_Ts) = map (strip_type o snd o fst #>> HOLogic.mk_tupleT) fixes |> split_list;
+
+    val _ = (case filter_out (fn (_, T) => Sign.of_sort thy (T, HOLogic.typeS)) (bs ~~ arg_Ts) of
+        [] => ()
+      | (b, _) :: _ => primcorec_error ("type of " ^ Binding.print b ^ " contains top sort"));
+
+    val actual_nn = length bs;
+
+    val sequentials = replicate actual_nn (member (op =) opts Sequential_Option);
+    val exhaustives = replicate actual_nn (member (op =) opts Exhaustive_Option);
+
+    val fun_names = map Binding.name_of bs;
+    val basic_ctr_specss = map (basic_corec_specs_of lthy) res_Ts;
+    val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
+    val eqns_data =
+      fold_map2 (dissect_coeqn lthy has_call fun_names sequentials basic_ctr_specss) (tag_list 0 (map snd specs))
+        of_specs_opt []
+      |> flat o fst;
+
+    val callssss =
+      map_filter (try (fn Sel x => x)) eqns_data
+      |> partition_eq (op = o pairself #fun_name)
+      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
+      |> map (flat o snd)
+      |> map2 (fold o find_corec_calls lthy has_call) basic_ctr_specss
+      |> map2 (curry (op |>)) (map (map (fn {ctr, sels, ...} =>
+        (ctr, map (K []) sels))) basic_ctr_specss);
+
+(*
+val _ = tracing ("callssss = " ^ @{make_string} callssss);
+*)
+
+    val ((n2m, corec_specs', _, coinduct_thm, strong_coinduct_thm, coinduct_thms,
+          strong_coinduct_thms), lthy') =
+      corec_specs_of bs arg_Ts res_Ts (get_indices fixes) callssss lthy;
+    val corec_specs = take actual_nn corec_specs';
+    val ctr_specss = map #ctr_specs corec_specs;
+
+    val disc_eqnss' = map_filter (try (fn Disc x => x)) eqns_data
+      |> partition_eq (op = o pairself #fun_name)
+      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
+      |> map (sort (op < o pairself #ctr_no |> make_ord) o flat o snd);
+    val _ = disc_eqnss' |> map (fn x =>
+      let val d = duplicates (op = o pairself #ctr_no) x in null d orelse
+        primcorec_error_eqns "excess discriminator formula in definition"
+          (maps (fn t => filter (curry (op =) (#ctr_no t) o #ctr_no) x) d |> map #user_eqn) end);
+
+    val sel_eqnss = map_filter (try (fn Sel x => x)) eqns_data
+      |> partition_eq (op = o pairself #fun_name)
+      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
+      |> map (flat o snd);
+
+    val arg_Tss = map (binder_types o snd o fst) fixes;
+    val disc_eqnss = map6 mk_actual_disc_eqns bs arg_Tss exhaustives corec_specs sel_eqnss
+      disc_eqnss';
+    val (defs, excludess') =
+      build_codefs lthy' bs mxs has_call arg_Tss corec_specs disc_eqnss sel_eqnss;
+
+    val tac_opts =
+      map (fn {code_rhs_opt, ...} :: _ =>
+        if auto orelse is_some code_rhs_opt then SOME (auto_tac o #context) else NONE) disc_eqnss;
+
+    fun exclude_tac tac_opt sequential (c, c', a) =
+      if a orelse c = c' orelse sequential then
+        SOME (K (HEADGOAL (mk_primcorec_assumption_tac lthy [])))
+      else
+        tac_opt;
+
+(*
+val _ = tracing ("exclusiveness properties:\n    \<cdot> " ^
+ space_implode "\n    \<cdot> " (maps (map (Syntax.string_of_term lthy o snd)) excludess'));
+*)
+
+    val excludess'' = map3 (fn tac_opt => fn sequential => map (fn (idx, goal) =>
+        (idx, (Option.map (Goal.prove_sorry lthy [] [] goal #> Thm.close_derivation)
+           (exclude_tac tac_opt sequential idx), goal))))
+      tac_opts sequentials excludess';
+    val taut_thmss = map (map (apsnd (the o fst)) o filter (is_some o fst o snd)) excludess'';
+    val (goal_idxss, exclude_goalss) = excludess''
+      |> map (map (apsnd (rpair [] o snd)) o filter (is_none o fst o snd))
+      |> split_list o map split_list;
+
+    fun list_all_fun_args extras =
+      map2 (fn [] => I
+          | {fun_args, ...} :: _ => map (curry Logic.list_all (extras @ map dest_Free fun_args)))
+        disc_eqnss;
+
+    val syntactic_exhaustives =
+      map (fn disc_eqns => forall (null o #prems orf is_some o #code_rhs_opt) disc_eqns
+          orelse exists #auto_gen disc_eqns)
+        disc_eqnss;
+    val de_facto_exhaustives =
+      map2 (fn b => fn b' => b orelse b') exhaustives syntactic_exhaustives;
+
+    val nchotomy_goalss =
+      map2 (fn false => K [] | true => single o HOLogic.mk_Trueprop o mk_dnf o map #prems)
+        de_facto_exhaustives disc_eqnss
+      |> list_all_fun_args []
+    val nchotomy_taut_thmss =
+      map6 (fn tac_opt => fn {disc_exhausts = res_disc_exhausts, ...} => fn arg_Ts =>
+          fn {code_rhs_opt, ...} :: _ => fn [] => K []
+            | [goal] => fn true =>
+              let
+                val (_, _, arg_disc_exhausts, _, _) =
+                  case_thms_of_term lthy arg_Ts (the_default Term.dummy code_rhs_opt);
+              in
+                [Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
+                   mk_primcorec_nchotomy_tac ctxt (res_disc_exhausts @ arg_disc_exhausts))
+                 |> Thm.close_derivation]
+              end
+            | false =>
+              (case tac_opt of
+                SOME tac => [Goal.prove_sorry lthy [] [] goal tac |> Thm.close_derivation]
+              | NONE => []))
+        tac_opts corec_specs arg_Tss disc_eqnss nchotomy_goalss syntactic_exhaustives;
+
+    val syntactic_exhaustives =
+      map (fn disc_eqns => forall (null o #prems orf is_some o #code_rhs_opt) disc_eqns
+          orelse exists #auto_gen disc_eqns)
+        disc_eqnss;
+
+    val nchotomy_goalss =
+      map2 (fn (NONE, false) => map (rpair []) | _ => K []) (tac_opts ~~ syntactic_exhaustives)
+        nchotomy_goalss;
+
+    val goalss = nchotomy_goalss @ exclude_goalss;
+
+    fun prove thmss'' def_thms' lthy =
+      let
+        val def_thms = map (snd o snd) def_thms';
+
+        val (nchotomy_thmss, exclude_thmss) =
+          (map2 append (take actual_nn thmss'') nchotomy_taut_thmss, drop actual_nn thmss'');
+
+        val ps =
+          Variable.variant_frees lthy (maps (maps #fun_args) disc_eqnss) [("P", HOLogic.boolT)];
+
+        val exhaust_thmss =
+          map2 (fn false => K []
+              | true => fn disc_eqns as {fun_args, ...} :: _ =>
+                let
+                  val p = Bound (length fun_args);
+                  fun mk_imp_p Qs = Logic.list_implies (Qs, HOLogic.mk_Trueprop p);
+                in
+                  [mk_imp_p (map (mk_imp_p o map HOLogic.mk_Trueprop o #prems) disc_eqns)]
+                end)
+            de_facto_exhaustives disc_eqnss
+          |> list_all_fun_args ps
+          |> map3 (fn disc_eqns as {fun_args, ...} :: _ => fn [] => K []
+              | [nchotomy_thm] => fn [goal] =>
+                [mk_primcorec_exhaust_tac lthy ("" (* for "P" *) :: map (fst o dest_Free) fun_args)
+                   (length disc_eqns) nchotomy_thm
+                 |> K |> Goal.prove_sorry lthy [] [] goal
+                 |> Thm.close_derivation])
+            disc_eqnss nchotomy_thmss;
+        val nontriv_exhaust_thmss = map (filter_out is_trivial_implies) exhaust_thmss;
+
+        val excludess' = map (op ~~) (goal_idxss ~~ exclude_thmss);
+        fun mk_excludesss excludes n =
+          fold (fn ((c, c', _), thm) => nth_map c (nth_map c' (K [thm])))
+            excludes (map (fn k => replicate k [asm_rl] @ replicate (n - k) []) (0 upto n - 1));
+        val excludessss =
+          map2 (fn excludes => mk_excludesss excludes o length o #ctr_specs)
+            (map2 append excludess' taut_thmss) corec_specs;
+
+        fun prove_disc ({ctr_specs, ...} : corec_spec) excludesss
+            ({fun_name, fun_T, fun_args, ctr_no, prems, eqn_pos, ...} : coeqn_data_disc) =
+          if Term.aconv_untyped (#disc (nth ctr_specs ctr_no), @{term "\<lambda>x. x = x"}) then
+            []
+          else
+            let
+              val {disc, disc_corec, ...} = nth ctr_specs ctr_no;
+              val k = 1 + ctr_no;
+              val m = length prems;
+              val goal =
+                applied_fun_of fun_name fun_T fun_args
+                |> curry betapply disc
+                |> HOLogic.mk_Trueprop
+                |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
+                |> curry Logic.list_all (map dest_Free fun_args);
+            in
+              if prems = [@{term False}] then
+                []
+              else
+                mk_primcorec_disc_tac lthy def_thms disc_corec k m excludesss
+                |> K |> Goal.prove_sorry lthy [] [] goal
+                |> Thm.close_derivation
+                |> pair (#disc (nth ctr_specs ctr_no))
+                |> pair eqn_pos
+                |> single
+            end;
+
+        fun prove_sel ({nested_maps, nested_map_idents, nested_map_comps, ctr_specs, ...}
+              : corec_spec) (disc_eqns : coeqn_data_disc list) excludesss
+            ({fun_name, fun_T, fun_args, ctr, sel, rhs_term, eqn_pos, ...} : coeqn_data_sel) =
+          let
+            val SOME ctr_spec = find_first (curry (op =) ctr o #ctr) ctr_specs;
+            val ctr_no = find_index (curry (op =) ctr o #ctr) ctr_specs;
+            val prems = the_default (maps (s_not_conj o #prems) disc_eqns)
+              (find_first (curry (op =) ctr_no o #ctr_no) disc_eqns |> Option.map #prems);
+            val sel_corec = find_index (curry (op =) sel) (#sels ctr_spec)
+              |> nth (#sel_corecs ctr_spec);
+            val k = 1 + ctr_no;
+            val m = length prems;
+            val goal =
+              applied_fun_of fun_name fun_T fun_args
+              |> curry betapply sel
+              |> rpair (abstract (List.rev fun_args) rhs_term)
+              |> HOLogic.mk_Trueprop o HOLogic.mk_eq
+              |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
+              |> curry Logic.list_all (map dest_Free fun_args);
+            val (distincts, _, _, sel_splits, sel_split_asms) = case_thms_of_term lthy [] rhs_term;
+          in
+            mk_primcorec_sel_tac lthy def_thms distincts sel_splits sel_split_asms nested_maps
+              nested_map_idents nested_map_comps sel_corec k m excludesss
+            |> K |> Goal.prove_sorry lthy [] [] goal
+            |> Thm.close_derivation
+            |> pair sel
+            |> pair eqn_pos
+          end;
+
+        fun prove_ctr disc_alist sel_alist (disc_eqns : coeqn_data_disc list)
+            (sel_eqns : coeqn_data_sel list) ({ctr, disc, sels, collapse, ...} : corec_ctr_spec) =
+          (* don't try to prove theorems when some sel_eqns are missing *)
+          if not (exists (curry (op =) ctr o #ctr) disc_eqns)
+              andalso not (exists (curry (op =) ctr o #ctr) sel_eqns)
+            orelse
+              filter (curry (op =) ctr o #ctr) sel_eqns
+              |> fst o finds (op = o apsnd #sel) sels
+              |> exists (null o snd) then
+            []
+          else
+            let
+              val (fun_name, fun_T, fun_args, prems, rhs_opt, eqn_pos) =
+                (find_first (curry (op =) ctr o #ctr) disc_eqns,
+                 find_first (curry (op =) ctr o #ctr) sel_eqns)
+                |>> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #prems x,
+                  #ctr_rhs_opt x, #eqn_pos x))
+                ||> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, [], #ctr_rhs_opt x,
+                  #eqn_pos x))
+                |> the o merge_options;
+              val m = length prems;
+              val goal =
+                (case rhs_opt of
+                  SOME rhs => rhs
+                | NONE =>
+                  filter (curry (op =) ctr o #ctr) sel_eqns
+                  |> fst o finds (op = o apsnd #sel) sels
+                  |> map (snd #> (fn [x] => (List.rev (#fun_args x), #rhs_term x)) #-> abstract)
+                  |> curry list_comb ctr)
+                |> curry mk_Trueprop_eq (applied_fun_of fun_name fun_T fun_args)
+                |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
+                |> curry Logic.list_all (map dest_Free fun_args);
+              val disc_thm_opt = AList.lookup (op =) disc_alist disc;
+              val sel_thms = map snd (filter (member (op =) sels o fst) sel_alist);
+            in
+              if prems = [@{term False}] then [] else
+                mk_primcorec_ctr_tac lthy m collapse disc_thm_opt sel_thms
+                |> K |> Goal.prove_sorry lthy [] [] goal
+                |> Thm.close_derivation
+                |> pair ctr
+                |> pair eqn_pos
+                |> single
+            end;
+
+        fun prove_code exhaustive disc_eqns sel_eqns nchotomys ctr_alist ctr_specs =
+          let
+            val fun_data_opt =
+              (find_first (member (op =) (map #ctr ctr_specs) o #ctr) disc_eqns,
+               find_first (member (op =) (map #ctr ctr_specs) o #ctr) sel_eqns)
+              |>> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #code_rhs_opt x))
+              ||> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #code_rhs_opt x))
+              |> merge_options;
+          in
+            (case fun_data_opt of
+              NONE => []
+            | SOME (fun_name, fun_T, fun_args, rhs_opt) =>
+              let
+                val bound_Ts = List.rev (map fastype_of fun_args);
+
+                val lhs = applied_fun_of fun_name fun_T fun_args;
+                val rhs_info_opt =
+                  (case rhs_opt of
+                    SOME rhs =>
+                    let
+                      val raw_rhs = expand_corec_code_rhs lthy has_call bound_Ts rhs;
+                      val cond_ctrs =
+                        fold_rev_corec_code_rhs lthy (K oo (cons oo pair)) bound_Ts raw_rhs [];
+                      val ctr_thms =
+                        map (the_default FalseE o AList.lookup (op =) ctr_alist o snd) cond_ctrs;
+                    in SOME (false, rhs, raw_rhs, ctr_thms) end
+                  | NONE =>
+                    let
+                      fun prove_code_ctr {ctr, sels, ...} =
+                        if not (exists (curry (op =) ctr o fst) ctr_alist) then NONE else
+                          let
+                            val prems = find_first (curry (op =) ctr o #ctr) disc_eqns
+                              |> Option.map #prems |> the_default [];
+                            val t =
+                              filter (curry (op =) ctr o #ctr) sel_eqns
+                              |> fst o finds (op = o apsnd #sel) sels
+                              |> map (snd #> (fn [x] => (List.rev (#fun_args x), #rhs_term x))
+                                #-> abstract)
+                              |> curry list_comb ctr;
+                          in
+                            SOME (prems, t)
+                          end;
+                      val ctr_conds_argss_opt = map prove_code_ctr ctr_specs;
+                      val exhaustive_code =
+                        exhaustive
+                        orelse exists (is_some andf (null o fst o the)) ctr_conds_argss_opt
+                        orelse forall is_some ctr_conds_argss_opt
+                          andalso exists #auto_gen disc_eqns;
+                      val rhs =
+                        (if exhaustive_code then
+                           split_last (map_filter I ctr_conds_argss_opt) ||> snd
+                         else
+                           Const (@{const_name Code.abort}, @{typ String.literal} -->
+                               (@{typ unit} --> body_type fun_T) --> body_type fun_T) $
+                             HOLogic.mk_literal fun_name $
+                             absdummy @{typ unit} (incr_boundvars 1 lhs)
+                           |> pair (map_filter I ctr_conds_argss_opt))
+                         |-> fold_rev (fn (prems, u) => mk_If (s_conjs prems) u)
+                    in
+                      SOME (exhaustive_code, rhs, rhs, map snd ctr_alist)
+                    end);
+              in
+                (case rhs_info_opt of
+                  NONE => []
+                | SOME (exhaustive_code, rhs, raw_rhs, ctr_thms) =>
+                  let
+                    val ms = map (Logic.count_prems o prop_of) ctr_thms;
+                    val (raw_goal, goal) = (raw_rhs, rhs)
+                      |> pairself (curry mk_Trueprop_eq (applied_fun_of fun_name fun_T fun_args)
+                        #> curry Logic.list_all (map dest_Free fun_args));
+                    val (distincts, discIs, _, sel_splits, sel_split_asms) =
+                      case_thms_of_term lthy bound_Ts raw_rhs;
+
+                    val raw_code_thm = mk_primcorec_raw_code_tac lthy distincts discIs sel_splits
+                        sel_split_asms ms ctr_thms
+                        (if exhaustive_code then try the_single nchotomys else NONE)
+                      |> K |> Goal.prove_sorry lthy [] [] raw_goal
+                      |> Thm.close_derivation;
+                  in
+                    mk_primcorec_code_tac lthy distincts sel_splits raw_code_thm
+                    |> K |> Goal.prove_sorry lthy [] [] goal
+                    |> Thm.close_derivation
+                    |> single
+                  end)
+              end)
+          end;
+
+        val disc_alistss = map3 (map oo prove_disc) corec_specs excludessss disc_eqnss;
+        val disc_alists = map (map snd o flat) disc_alistss;
+        val sel_alists = map4 (map ooo prove_sel) corec_specs disc_eqnss excludessss sel_eqnss;
+        val disc_thmss = map (map snd o order_list_duplicates o flat) disc_alistss;
+        val disc_thmsss' = map (map (map (snd o snd))) disc_alistss;
+        val sel_thmss = map (map snd o order_list_duplicates) sel_alists;
+
+        fun prove_disc_iff ({ctr_specs, ...} : corec_spec) exhaust_thms disc_thmss'
+            (({fun_args = exhaust_fun_args, ...} : coeqn_data_disc) :: _) disc_thms
+            ({fun_name, fun_T, fun_args, ctr_no, prems, eqn_pos, ...} : coeqn_data_disc) =
+          if null exhaust_thms orelse null (tl ctr_specs) then
+            []
+          else
+            let
+              val {disc, disc_excludess, ...} = nth ctr_specs ctr_no;
+              val goal =
+                mk_Trueprop_eq (applied_fun_of fun_name fun_T fun_args |> curry betapply disc,
+                  mk_conjs prems)
+                |> curry Logic.list_all (map dest_Free fun_args);
+            in
+              mk_primcorec_disc_iff_tac lthy (map (fst o dest_Free) exhaust_fun_args)
+                (the_single exhaust_thms) disc_thms disc_thmss' (flat disc_excludess)
+              |> K |> Goal.prove_sorry lthy [] [] goal
+              |> Thm.close_derivation
+              |> fold (fn rule => perhaps (try (fn thm => thm RS rule)))
+                @{thms eqTrueE eq_False[THEN iffD1] notnotD}
+              |> pair eqn_pos
+              |> single
+            end;
+
+        val disc_iff_thmss = map6 (flat ooo map2 oooo prove_disc_iff) corec_specs exhaust_thmss
+          disc_thmsss' disc_eqnss disc_thmsss' disc_eqnss
+          |> map order_list_duplicates;
+
+        val ctr_alists = map5 (maps oooo prove_ctr) disc_alists (map (map snd) sel_alists) disc_eqnss
+          sel_eqnss ctr_specss;
+        val ctr_thmss' = map (map snd) ctr_alists;
+        val ctr_thmss = map (map snd o order_list) ctr_alists;
+
+        val code_thmss = map6 prove_code exhaustives disc_eqnss sel_eqnss nchotomy_thmss ctr_thmss'
+          ctr_specss;
+
+        val disc_iff_or_disc_thmss =
+          map2 (fn [] => I | disc_iffs => K disc_iffs) disc_iff_thmss disc_thmss;
+        val simp_thmss = map2 append disc_iff_or_disc_thmss sel_thmss;
+
+        val common_name = mk_common_name fun_names;
+
+        val notes =
+          [(coinductN, map (if n2m then single else K []) coinduct_thms, []),
+           (codeN, code_thmss, code_nitpicksimp_attrs),
+           (ctrN, ctr_thmss, []),
+           (discN, disc_thmss, simp_attrs),
+           (disc_iffN, disc_iff_thmss, []),
+           (excludeN, exclude_thmss, []),
+           (exhaustN, nontriv_exhaust_thmss, []),
+           (selN, sel_thmss, simp_attrs),
+           (simpsN, simp_thmss, []),
+           (strong_coinductN, map (if n2m then single else K []) strong_coinduct_thms, [])]
+          |> maps (fn (thmN, thmss, attrs) =>
+            map2 (fn fun_name => fn thms =>
+                ((Binding.qualify true fun_name (Binding.name thmN), attrs), [(thms, [])]))
+              fun_names (take actual_nn thmss))
+          |> filter_out (null o fst o hd o snd);
+
+        val common_notes =
+          [(coinductN, if n2m then [coinduct_thm] else [], []),
+           (strong_coinductN, if n2m then [strong_coinduct_thm] else [], [])]
+          |> filter_out (null o #2)
+          |> map (fn (thmN, thms, attrs) =>
+            ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
+      in
+        lthy |> Local_Theory.notes (notes @ common_notes) |> snd
+      end;
+
+    fun after_qed thmss' = fold_map Local_Theory.define defs #-> prove thmss';
+  in
+    (goalss, after_qed, lthy')
+  end;
+
+fun add_primcorec_ursive_cmd auto opts (raw_fixes, raw_specs') lthy =
+  let
+    val (raw_specs, of_specs_opt) =
+      split_list raw_specs' ||> map (Option.map (Syntax.read_term lthy));
+    val ((fixes, specs), _) = Specification.read_spec raw_fixes raw_specs lthy;
+  in
+    add_primcorec_ursive auto opts fixes specs of_specs_opt lthy
+    handle ERROR str => primcorec_error str
+  end
+  handle Primcorec_Error (str, eqns) =>
+    if null eqns
+    then error ("primcorec error:\n  " ^ str)
+    else error ("primcorec error:\n  " ^ str ^ "\nin\n  " ^
+      space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns));
+
+val add_primcorecursive_cmd = (fn (goalss, after_qed, lthy) =>
+  lthy
+  |> Proof.theorem NONE after_qed goalss
+  |> Proof.refine (Method.primitive_text (K I))
+  |> Seq.hd) ooo add_primcorec_ursive_cmd false;
+
+val add_primcorec_cmd = (fn (goalss, after_qed, lthy) =>
+  lthy
+  |> after_qed (map (fn [] => []
+      | _ => error "\"auto\" failed -- use \"primcorecursive\" instead of \"primcorec\"")
+    goalss)) ooo add_primcorec_ursive_cmd true;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,209 @@
+(*  Title:      HOL/BNF/Tools/bnf_gfp_rec_sugar_tactics.ML
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2013
+
+Tactics for corecursor sugar.
+*)
+
+signature BNF_GFP_REC_SUGAR_TACTICS =
+sig
+  val mk_primcorec_assumption_tac: Proof.context -> thm list -> int -> tactic
+  val mk_primcorec_code_tac: Proof.context -> thm list -> thm list -> thm -> tactic
+  val mk_primcorec_ctr_tac: Proof.context -> int -> thm -> thm option -> thm list -> tactic
+  val mk_primcorec_disc_tac: Proof.context -> thm list -> thm -> int -> int -> thm list list list ->
+    tactic
+  val mk_primcorec_disc_iff_tac: Proof.context -> string list -> thm -> thm list -> thm list list ->
+    thm list -> tactic
+  val mk_primcorec_exhaust_tac: Proof.context -> string list -> int -> thm -> tactic
+  val mk_primcorec_nchotomy_tac: Proof.context -> thm list -> tactic
+  val mk_primcorec_raw_code_tac: Proof.context -> thm list -> thm list -> thm list -> thm list ->
+    int list -> thm list -> thm option -> tactic
+  val mk_primcorec_sel_tac: Proof.context -> thm list -> thm list -> thm list -> thm list ->
+    thm list -> thm list -> thm list -> thm -> int -> int -> thm list list list -> tactic
+end;
+
+structure BNF_GFP_Rec_Sugar_Tactics : BNF_GFP_REC_SUGAR_TACTICS =
+struct
+
+open BNF_Util
+open BNF_Tactics
+open BNF_FP_Util
+
+val atomize_conjL = @{thm atomize_conjL};
+val falseEs = @{thms not_TrueE FalseE};
+val neq_eq_eq_contradict = @{thm neq_eq_eq_contradict};
+val split_if = @{thm split_if};
+val split_if_asm = @{thm split_if_asm};
+val split_connectI = @{thms allI impI conjI};
+val unfold_lets = @{thms Let_def[abs_def] split_beta}
+
+fun exhaust_inst_as_projs ctxt frees thm =
+  let
+    val num_frees = length frees;
+    val fs = Term.add_vars (prop_of thm) [] |> filter (can dest_funT o snd);
+    fun find s = find_index (curry (op =) s) frees;
+    fun mk_cfp (f as ((s, _), T)) =
+      (certify ctxt (Var f), certify ctxt (mk_proj T num_frees (find s)));
+    val cfps = map mk_cfp fs;
+  in
+    Drule.cterm_instantiate cfps thm
+  end;
+
+val exhaust_inst_as_projs_tac = PRIMITIVE oo exhaust_inst_as_projs;
+
+fun distinct_in_prems_tac distincts =
+  eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac;
+
+fun mk_primcorec_nchotomy_tac ctxt disc_exhausts =
+  HEADGOAL (Method.insert_tac disc_exhausts THEN' clean_blast_tac ctxt);
+
+fun mk_primcorec_exhaust_tac ctxt frees n nchotomy =
+  let val ks = 1 upto n in
+    HEADGOAL (atac ORELSE'
+      cut_tac nchotomy THEN'
+      K (exhaust_inst_as_projs_tac ctxt frees) THEN'
+      EVERY' (map (fn k =>
+          (if k < n then etac disjE else K all_tac) THEN'
+          REPEAT o (dtac meta_mp THEN' atac ORELSE'
+            etac conjE THEN' dtac meta_mp THEN' atac ORELSE'
+            atac))
+        ks))
+  end;
+
+fun mk_primcorec_assumption_tac ctxt discIs =
+  SELECT_GOAL (unfold_thms_tac ctxt @{thms fst_conv snd_conv not_not not_False_eq_True
+      not_True_eq_False de_Morgan_conj de_Morgan_disj} THEN
+    SOLVE (HEADGOAL (REPEAT o (rtac refl ORELSE' atac ORELSE' etac conjE ORELSE'
+    eresolve_tac falseEs ORELSE'
+    resolve_tac @{thms TrueI conjI disjI1 disjI2} ORELSE'
+    dresolve_tac discIs THEN' atac ORELSE'
+    etac notE THEN' atac ORELSE'
+    etac disjE))));
+
+val ss_fst_snd_conv = simpset_of (ss_only @{thms fst_conv snd_conv} @{context});
+
+fun case_atac ctxt = simp_tac (put_simpset ss_fst_snd_conv ctxt);
+
+fun same_case_tac ctxt m =
+  HEADGOAL (if m = 0 then rtac TrueI
+    else REPEAT_DETERM_N (m - 1) o (rtac conjI THEN' case_atac ctxt) THEN' case_atac ctxt);
+
+fun different_case_tac ctxt m exclude =
+  HEADGOAL (if m = 0 then
+      mk_primcorec_assumption_tac ctxt []
+    else
+      dtac exclude THEN' (REPEAT_DETERM_N (m - 1) o case_atac ctxt) THEN'
+      mk_primcorec_assumption_tac ctxt []);
+
+fun cases_tac ctxt k m excludesss =
+  let val n = length excludesss in
+    EVERY (map (fn [] => if k = n then all_tac else same_case_tac ctxt m
+        | [exclude] => different_case_tac ctxt m exclude)
+      (take k (nth excludesss (k - 1))))
+  end;
+
+fun prelude_tac ctxt defs thm =
+  unfold_thms_tac ctxt defs THEN HEADGOAL (rtac thm) THEN unfold_thms_tac ctxt unfold_lets;
+
+fun mk_primcorec_disc_tac ctxt defs disc_corec k m excludesss =
+  prelude_tac ctxt defs disc_corec THEN cases_tac ctxt k m excludesss;
+
+fun mk_primcorec_disc_iff_tac ctxt fun_exhaust_frees fun_exhaust fun_discs fun_discss
+    disc_excludes =
+  HEADGOAL (rtac iffI THEN'
+    rtac fun_exhaust THEN'
+    K (exhaust_inst_as_projs_tac ctxt fun_exhaust_frees) THEN'
+    EVERY' (map (fn [] => etac FalseE
+        | fun_discs' as [fun_disc'] =>
+          if eq_list Thm.eq_thm (fun_discs', fun_discs) then
+            REPEAT_DETERM o etac conjI THEN' (atac ORELSE' rtac TrueI)
+          else
+            rtac FalseE THEN'
+            (rotate_tac 1 THEN' dtac fun_disc' THEN' REPEAT_DETERM o atac ORELSE'
+             cut_tac fun_disc') THEN'
+            dresolve_tac disc_excludes THEN' etac notE THEN' atac)
+      fun_discss) THEN'
+    (etac FalseE ORELSE'
+     resolve_tac (map (unfold_thms ctxt [atomize_conjL]) fun_discs) THEN_MAYBE' atac));
+
+fun mk_primcorec_sel_tac ctxt defs distincts splits split_asms mapsx map_idents map_comps fun_sel k
+    m excludesss =
+  prelude_tac ctxt defs (fun_sel RS trans) THEN
+  cases_tac ctxt k m excludesss THEN
+  HEADGOAL (REPEAT_DETERM o (rtac refl ORELSE' rtac ext ORELSE'
+    eresolve_tac falseEs ORELSE'
+    resolve_tac split_connectI ORELSE'
+    Splitter.split_asm_tac (split_if_asm :: split_asms) ORELSE'
+    Splitter.split_tac (split_if :: splits) ORELSE'
+    eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac ORELSE'
+    etac notE THEN' atac ORELSE'
+    (CHANGED o SELECT_GOAL (unfold_thms_tac ctxt (@{thms fst_conv snd_conv id_def o_def split_def
+       sum.cases} @ mapsx @ map_comps @ map_idents)))));
+
+fun mk_primcorec_ctr_tac ctxt m collapse disc_fun_opt sel_funs =
+  HEADGOAL (rtac ((if null sel_funs then collapse else collapse RS sym) RS trans) THEN'
+    (the_default (K all_tac) (Option.map rtac disc_fun_opt)) THEN' REPEAT_DETERM_N m o atac) THEN
+  unfold_thms_tac ctxt (@{thm split_def} :: unfold_lets @ sel_funs) THEN HEADGOAL (rtac refl);
+
+fun inst_split_eq ctxt split =
+  (case prop_of split of
+    @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ (Var (_, Type (_, [T, _])) $ _) $ _) =>
+    let
+      val s = Name.uu;
+      val eq = Abs (Name.uu, T, HOLogic.mk_eq (Free (s, T), Bound 0));
+      val split' = Drule.instantiate' [] [SOME (certify ctxt eq)] split;
+    in
+      Thm.generalize ([], [s]) (Thm.maxidx_of split' + 1) split'
+    end
+  | _ => split);
+
+fun raw_code_single_tac ctxt distincts discIs splits split_asms m fun_ctr =
+  let
+    val splits' =
+      map (fn th => th RS iffD2) (@{thm split_if_eq2} :: map (inst_split_eq ctxt) splits)
+  in
+    HEADGOAL (REPEAT o (resolve_tac (splits' @ split_connectI))) THEN
+    prelude_tac ctxt [] (fun_ctr RS trans) THEN
+    HEADGOAL ((REPEAT_DETERM_N m o mk_primcorec_assumption_tac ctxt discIs) THEN'
+      SELECT_GOAL (SOLVE (HEADGOAL (REPEAT_DETERM o
+      (rtac refl ORELSE' atac ORELSE'
+       resolve_tac (@{thm Code.abort_def} :: split_connectI) ORELSE'
+       Splitter.split_tac (split_if :: splits) ORELSE'
+       Splitter.split_asm_tac (split_if_asm :: split_asms) ORELSE'
+       mk_primcorec_assumption_tac ctxt discIs ORELSE'
+       distinct_in_prems_tac distincts ORELSE'
+       (TRY o dresolve_tac discIs) THEN' etac notE THEN' atac)))))
+  end;
+
+fun rulify_nchotomy n = funpow (n - 1) (fn thm => thm RS @{thm Meson.make_pos_rule'});
+
+fun mk_primcorec_raw_code_tac ctxt distincts discIs splits split_asms ms fun_ctrs nchotomy_opt =
+  let
+    val n = length ms;
+    val (ms', fun_ctrs') =
+      (case nchotomy_opt of
+        NONE => (ms, fun_ctrs)
+      | SOME nchotomy =>
+        (ms |> split_last ||> K [n - 1] |> op @,
+         fun_ctrs
+         |> split_last
+         ||> unfold_thms ctxt [atomize_conjL]
+         ||> (fn thm => [rulify_nchotomy n nchotomy RS thm] handle THM _ => [thm])
+         |> op @));
+  in
+    EVERY (map2 (raw_code_single_tac ctxt distincts discIs splits split_asms) ms' fun_ctrs') THEN
+    IF_UNSOLVED (unfold_thms_tac ctxt @{thms Code.abort_def} THEN
+      HEADGOAL (REPEAT_DETERM o resolve_tac (refl :: split_connectI)))
+  end;
+
+fun mk_primcorec_code_tac ctxt distincts splits raw =
+  HEADGOAL (rtac raw ORELSE' rtac (raw RS trans) THEN'
+    SELECT_GOAL (unfold_thms_tac ctxt unfold_lets) THEN' REPEAT_DETERM o
+    (rtac refl ORELSE' atac ORELSE'
+     resolve_tac split_connectI ORELSE'
+     Splitter.split_tac (split_if :: splits) ORELSE'
+     distinct_in_prems_tac distincts ORELSE'
+     rtac sym THEN' atac ORELSE'
+     etac notE THEN' atac));
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_gfp_tactics.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,1304 @@
+(*  Title:      HOL/BNF/Tools/bnf_gfp_tactics.ML
+    Author:     Dmitriy Traytel, TU Muenchen
+    Author:     Andrei Popescu, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2012
+
+Tactics for the codatatype construction.
+*)
+
+signature BNF_GFP_TACTICS =
+sig
+  val mk_Lev_sbd_tac: Proof.context -> cterm option list -> thm list -> thm list ->
+    thm list list -> tactic
+  val mk_bis_Gr_tac: thm -> thm list -> thm list -> thm list -> thm list ->
+    {prems: 'a, context: Proof.context} -> tactic
+  val mk_bis_O_tac: Proof.context -> int -> thm -> thm list -> thm list -> tactic
+  val mk_bis_Union_tac: thm -> thm list -> {prems: 'a, context: Proof.context} -> tactic
+  val mk_bis_converse_tac: int -> thm -> thm list -> thm list -> tactic
+  val mk_bis_rel_tac: Proof.context -> int -> thm -> thm list -> thm list -> thm list ->
+    thm list list -> tactic
+  val mk_coalgT_tac: int -> thm list -> thm list -> thm list list ->
+    {prems: 'a, context: Proof.context} -> tactic
+  val mk_coalg_final_tac: int -> thm -> thm list -> thm list -> thm list list -> thm list list ->
+    tactic
+  val mk_coalg_set_tac: thm -> tactic
+  val mk_coind_wit_tac: thm -> thm list -> thm list -> thm list ->
+    {prems: 'a, context: Proof.context} -> tactic
+  val mk_col_bd_tac: int -> int -> cterm option list -> thm list -> thm list -> thm -> thm ->
+    thm list list -> tactic
+  val mk_col_natural_tac: cterm option list -> thm list -> thm list -> thm list -> thm list list ->
+    {prems: 'a, context: Proof.context} -> tactic
+  val mk_congruent_str_final_tac: int -> thm -> thm -> thm -> thm list -> tactic
+  val mk_corec_tac: int -> thm list -> thm -> thm -> thm list ->
+    {prems: 'a, context: Proof.context} -> tactic
+  val mk_corec_unique_mor_tac: thm list -> thm list -> thm -> {prems: 'a, context: Proof.context} ->
+    tactic
+  val mk_dtor_map_coinduct_tac: int -> int list -> thm -> thm -> tactic
+  val mk_dtor_coinduct_tac: int -> thm -> thm -> thm list -> tactic
+  val mk_dtor_rel_tac: Proof.context -> thm list -> int -> thm -> thm -> thm -> thm -> thm list ->
+    thm -> thm -> thm list -> thm list -> thm list list -> tactic
+  val mk_dtor_o_ctor_tac: thm -> thm -> thm -> thm -> thm list ->
+    {prems: 'a, context: Proof.context} -> tactic
+  val mk_equiv_lsbis_tac: thm -> thm -> thm -> thm -> thm -> thm -> tactic
+  val mk_hset_minimal_tac: int -> thm list -> thm -> {prems: 'a, context: Proof.context} -> tactic
+  val mk_hset_rec_minimal_tac: int -> cterm option list -> thm list -> thm list ->
+    {prems: 'a, context: Proof.context} -> tactic
+  val mk_incl_lsbis_tac: int -> int -> thm -> tactic
+  val mk_length_Lev'_tac: thm -> tactic
+  val mk_length_Lev_tac: Proof.context -> cterm option list -> thm list -> thm list -> tactic
+  val mk_map_comp0_tac: thm list -> thm list -> thm -> tactic
+  val mk_mcong_tac: Proof.context -> int -> (int -> tactic) -> thm list -> thm list -> thm list ->
+    thm list list -> thm list list -> thm list list list -> tactic
+  val mk_map_id0_tac: thm list -> thm -> thm -> tactic
+  val mk_map_tac: int -> int -> ctyp option -> thm -> thm -> thm -> tactic
+  val mk_dtor_map_unique_tac: thm -> thm list -> Proof.context -> tactic
+  val mk_mor_Abs_tac: thm list -> thm list -> {prems: 'a, context: Proof.context} -> tactic
+  val mk_mor_Rep_tac: int -> thm list -> thm list -> thm list -> thm list list -> thm list ->
+    thm list -> {prems: 'a, context: Proof.context} -> tactic
+  val mk_mor_T_final_tac: thm -> thm list -> thm list -> tactic
+  val mk_mor_UNIV_tac: thm list -> thm -> tactic
+  val mk_mor_beh_tac: int -> thm -> thm -> thm list -> thm list -> thm list -> thm list ->
+    thm list list -> thm list list -> thm list -> thm list -> thm list -> thm list -> thm list ->
+    thm list -> thm list -> thm list -> thm list list -> thm list list list -> thm list list list ->
+    thm list list list -> thm list list -> thm list list -> thm list -> thm list -> thm list ->
+    {prems: 'a, context: Proof.context} -> tactic
+  val mk_mor_comp_tac: thm -> thm list -> thm list -> thm list -> tactic
+  val mk_mor_elim_tac: thm -> tactic
+  val mk_mor_hset_rec_tac: int -> int -> cterm option list -> int -> thm list -> thm list ->
+    thm list -> thm list list -> thm list list -> tactic
+  val mk_mor_incl_tac: thm -> thm list -> tactic
+  val mk_mor_str_tac: 'a list -> thm -> tactic
+  val mk_mor_sum_case_tac: 'a list -> thm -> tactic
+  val mk_mor_unfold_tac: int -> thm -> thm list -> thm list -> thm list -> thm list -> thm list ->
+    thm list -> tactic
+  val mk_prefCl_Lev_tac: Proof.context -> cterm option list -> thm list -> thm list -> tactic
+  val mk_raw_coind_tac: thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm list ->
+    thm list -> thm list -> thm -> thm list -> tactic
+  val mk_rel_coinduct_tac: thm list -> thm list -> thm list list -> thm list -> thm list ->
+    {prems: thm list, context: Proof.context} -> tactic
+  val mk_rel_coinduct_coind_tac: int -> thm -> int list -> thm list -> thm list -> thm list ->
+    thm list list -> thm list -> thm list -> {prems: 'a, context: Proof.context} -> tactic
+  val mk_rel_coinduct_ind_tac: int -> int list -> thm list -> thm list list -> int -> thm ->
+    {prems: 'a, context: Proof.context} -> tactic
+  val mk_rv_last_tac: ctyp option list -> cterm option list -> thm list -> thm list -> tactic
+  val mk_sbis_lsbis_tac: Proof.context -> thm list -> thm -> thm -> tactic
+  val mk_set_Lev_tac: Proof.context -> cterm option list -> thm list -> thm list -> thm list ->
+    thm list -> thm list list -> tactic
+  val mk_set_bd_tac: thm -> thm -> thm -> tactic
+  val mk_set_hset_incl_hset_tac: int -> thm list -> thm -> int -> tactic
+  val mk_set_image_Lev_tac: Proof.context -> cterm option list -> thm list -> thm list ->
+    thm list -> thm list -> thm list list -> thm list list -> tactic
+  val mk_set_incl_hset_tac: thm -> thm -> tactic
+  val mk_set_ge_tac: int  -> thm -> thm list -> tactic
+  val mk_set_le_tac: int -> thm -> thm list -> thm list list -> tactic
+  val mk_set_map0_tac: thm -> thm -> tactic
+  val mk_set_rv_Lev_tac: Proof.context -> int -> cterm option list -> thm list -> thm list ->
+    thm list -> thm list -> thm list list -> thm list list -> tactic
+  val mk_unfold_unique_mor_tac: thm list -> thm -> thm -> thm list -> tactic
+  val mk_unfold_transfer_tac: int -> thm -> thm list -> thm list ->
+    {prems: thm list, context: Proof.context} -> tactic
+  val mk_wit_tac: int -> thm list -> thm list -> thm list -> thm list -> Proof.context -> tactic
+  val mk_le_rel_OO_tac: thm -> thm list -> thm list -> tactic
+end;
+
+structure BNF_GFP_Tactics : BNF_GFP_TACTICS =
+struct
+
+open BNF_Tactics
+open BNF_Util
+open BNF_FP_Util
+open BNF_GFP_Util
+
+val fst_convol_fun_cong_sym = @{thm fst_convol} RS fun_cong RS sym;
+val list_inject_iffD1 = @{thm list.inject[THEN iffD1]};
+val nat_induct = @{thm nat_induct};
+val o_apply_trans_sym = o_apply RS trans RS sym;
+val ord_eq_le_trans = @{thm ord_eq_le_trans};
+val ord_eq_le_trans_trans_fun_cong_image_id_id_apply =
+  @{thm ord_eq_le_trans[OF trans[OF fun_cong[OF image_id] id_apply]]};
+val ordIso_ordLeq_trans = @{thm ordIso_ordLeq_trans};
+val snd_convol_fun_cong_sym = @{thm snd_convol} RS fun_cong RS sym;
+val sum_case_weak_cong = @{thm sum_case_weak_cong};
+val trans_fun_cong_image_id_id_apply = @{thm trans[OF fun_cong[OF image_id] id_apply]};
+val Collect_splitD_set_mp = @{thm Collect_splitD[OF set_mp]};
+val rev_bspec = Drule.rotate_prems 1 bspec;
+val Un_cong = @{thm arg_cong2[of _ _ _ _ "op \<union>"]};
+val conversep_in_rel_Id_on =
+   @{thm trans[OF conversep_in_rel arg_cong[of _ _ in_rel, OF converse_Id_on]]};
+val relcompp_in_rel_Id_on =
+   @{thm   trans[OF relcompp_in_rel arg_cong[of _ _ in_rel, OF Id_on_Comp[symmetric]]]};
+val converse_shift = @{thm converse_subset_swap} RS iffD1;
+
+fun mk_coalg_set_tac coalg_def =
+  dtac (coalg_def RS iffD1) 1 THEN
+  REPEAT_DETERM (etac conjE 1) THEN
+  EVERY' [dtac rev_bspec, atac] 1 THEN
+  REPEAT_DETERM (eresolve_tac [CollectE, conjE] 1) THEN atac 1;
+
+fun mk_mor_elim_tac mor_def =
+  (dtac (subst OF [mor_def]) THEN'
+  REPEAT o etac conjE THEN'
+  TRY o rtac @{thm image_subsetI} THEN'
+  etac bspec THEN'
+  atac) 1;
+
+fun mk_mor_incl_tac mor_def map_ids =
+  (stac mor_def THEN'
+  rtac conjI THEN'
+  CONJ_WRAP' (K (EVERY' [rtac ballI, etac set_mp, stac id_apply, atac])) map_ids THEN'
+  CONJ_WRAP' (fn thm =>
+   (EVERY' [rtac ballI, rtac (thm RS trans), rtac sym, rtac (id_apply RS arg_cong)])) map_ids) 1;
+
+fun mk_mor_comp_tac mor_def mor_images morEs map_comp_ids =
+  let
+    fun fbetw_tac image = EVERY' [rtac ballI, stac o_apply, etac image, etac image, atac];
+    fun mor_tac ((mor_image, morE), map_comp_id) =
+      EVERY' [rtac ballI, stac o_apply, rtac trans, rtac (map_comp_id RS sym), rtac trans,
+        etac (morE RS arg_cong), atac, etac morE, etac mor_image, atac];
+  in
+    (stac mor_def THEN' rtac conjI THEN'
+    CONJ_WRAP' fbetw_tac mor_images THEN'
+    CONJ_WRAP' mor_tac ((mor_images ~~ morEs) ~~ map_comp_ids)) 1
+  end;
+
+fun mk_mor_UNIV_tac morEs mor_def =
+  let
+    val n = length morEs;
+    fun mor_tac morE = EVERY' [rtac ext, rtac trans, rtac o_apply, rtac trans, etac morE,
+      rtac UNIV_I, rtac sym, rtac o_apply];
+  in
+    EVERY' [rtac iffI, CONJ_WRAP' mor_tac morEs,
+    stac mor_def, rtac conjI, CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) morEs,
+    CONJ_WRAP' (fn i =>
+      EVERY' [dtac (mk_conjunctN n i), rtac ballI, etac @{thm comp_eq_dest}]) (1 upto n)] 1
+  end;
+
+fun mk_mor_str_tac ks mor_UNIV =
+  (stac mor_UNIV THEN' CONJ_WRAP' (K (rtac refl)) ks) 1;
+
+fun mk_mor_sum_case_tac ks mor_UNIV =
+  (stac mor_UNIV THEN' CONJ_WRAP' (K (rtac @{thm sum_case_o_inj(1)[symmetric]})) ks) 1;
+
+fun mk_set_incl_hset_tac def rec_Suc =
+  EVERY' (stac def ::
+    map rtac [@{thm SUP_upper2}, UNIV_I, @{thm ord_le_eq_trans}, @{thm Un_upper1},
+      sym, rec_Suc]) 1;
+
+fun mk_set_hset_incl_hset_tac n defs rec_Suc i =
+  EVERY' (map (TRY oo stac) defs @
+    map rtac [@{thm UN_least}, subsetI, @{thm UN_I}, UNIV_I, set_mp, equalityD2, rec_Suc, UnI2,
+      mk_UnIN n i] @
+    [etac @{thm UN_I}, atac]) 1;
+
+fun mk_hset_rec_minimal_tac m cts rec_0s rec_Sucs {context = ctxt, prems = _} =
+  EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
+    REPEAT_DETERM o rtac allI,
+    CONJ_WRAP' (fn thm => EVERY'
+      [rtac ord_eq_le_trans, rtac thm, rtac @{thm empty_subsetI}]) rec_0s,
+    REPEAT_DETERM o rtac allI,
+    CONJ_WRAP' (fn rec_Suc => EVERY'
+      [rtac ord_eq_le_trans, rtac rec_Suc,
+        if m = 0 then K all_tac
+        else (rtac @{thm Un_least} THEN' Goal.assume_rule_tac ctxt),
+        CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least}))
+          (K (EVERY' [rtac @{thm UN_least}, REPEAT_DETERM o eresolve_tac [allE, conjE],
+            rtac subset_trans, atac, Goal.assume_rule_tac ctxt])) rec_0s])
+      rec_Sucs] 1;
+
+fun mk_hset_minimal_tac n hset_defs hset_rec_minimal {context = ctxt, prems = _} =
+  (CONJ_WRAP' (fn def => (EVERY' [rtac ord_eq_le_trans, rtac def,
+    rtac @{thm UN_least}, rtac rev_mp, rtac hset_rec_minimal,
+    EVERY' (replicate ((n + 1) * n) (Goal.assume_rule_tac ctxt)), rtac impI,
+    REPEAT_DETERM o eresolve_tac [allE, conjE], atac])) hset_defs) 1
+
+fun mk_mor_hset_rec_tac m n cts j rec_0s rec_Sucs morEs set_map0ss coalg_setss =
+  EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
+    REPEAT_DETERM o rtac allI,
+    CONJ_WRAP' (fn thm => EVERY' (map rtac [impI, thm RS trans, thm RS sym])) rec_0s,
+    REPEAT_DETERM o rtac allI,
+    CONJ_WRAP'
+      (fn (rec_Suc, (morE, ((passive_set_map0s, active_set_map0s), coalg_sets))) =>
+        EVERY' [rtac impI, rtac (rec_Suc RS trans), rtac (rec_Suc RS trans RS sym),
+          if m = 0 then K all_tac
+          else EVERY' [rtac Un_cong, rtac box_equals,
+            rtac (nth passive_set_map0s (j - 1) RS sym),
+            rtac trans_fun_cong_image_id_id_apply, etac (morE RS arg_cong), atac],
+          CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 Un_cong))
+            (fn (i, (set_map0, coalg_set)) =>
+              EVERY' [rtac sym, rtac trans, rtac (refl RSN (2, @{thm UN_cong})),
+                etac (morE RS sym RS arg_cong RS trans), atac, rtac set_map0,
+                rtac (@{thm UN_simps(10)} RS trans), rtac (refl RS @{thm UN_cong}),
+                ftac coalg_set, atac, dtac set_mp, atac, rtac mp, rtac (mk_conjunctN n i),
+                REPEAT_DETERM o etac allE, atac, atac])
+            (rev ((1 upto n) ~~ (active_set_map0s ~~ coalg_sets)))])
+      (rec_Sucs ~~ (morEs ~~ (map (chop m) set_map0ss ~~ map (drop m) coalg_setss)))] 1;
+
+fun mk_bis_rel_tac ctxt m bis_def rel_OO_Grps map_comp0s map_cong0s set_map0ss =
+  let
+    val n = length rel_OO_Grps;
+    val thms = ((1 upto n) ~~ map_comp0s ~~ map_cong0s ~~ set_map0ss ~~ rel_OO_Grps);
+
+    fun mk_if_tac ((((i, map_comp0), map_cong0), set_map0s), rel_OO_Grp) =
+      EVERY' [rtac allI, rtac allI, rtac impI, dtac (mk_conjunctN n i),
+        etac allE, etac allE, etac impE, atac, etac bexE, etac conjE,
+        rtac (rel_OO_Grp RS sym RS @{thm eq_refl} RS @{thm predicate2D}),
+        rtac @{thm relcomppI}, rtac @{thm conversepI},
+        EVERY' (map (fn thm =>
+          EVERY' [rtac @{thm GrpI},
+            rtac trans, rtac trans, rtac map_comp0, rtac map_cong0, REPEAT_DETERM_N m o rtac thm,
+            REPEAT_DETERM_N n o rtac (@{thm o_id} RS fun_cong), atac,
+            REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac CollectI,
+            CONJ_WRAP' (fn (i, thm) =>
+              if i <= m
+              then EVERY' [rtac ord_eq_le_trans, rtac thm, rtac subset_trans,
+                etac @{thm image_mono}, rtac @{thm image_subsetI},
+                etac @{thm Collect_split_in_relI[OF Id_onI]}]
+              else EVERY' [rtac ord_eq_le_trans, rtac trans, rtac thm,
+                rtac trans_fun_cong_image_id_id_apply, etac @{thm Collect_split_in_rel_leI}])
+            (1 upto (m + n) ~~ set_map0s)])
+          @{thms fst_diag_id snd_diag_id})];
+
+    fun mk_only_if_tac ((((i, map_comp0), map_cong0), set_map0s), rel_OO_Grp) =
+      EVERY' [dtac (mk_conjunctN n i), rtac allI, rtac allI, rtac impI,
+        etac allE, etac allE, etac impE, atac,
+        dtac (rel_OO_Grp RS @{thm eq_refl} RS @{thm predicate2D}),
+        REPEAT_DETERM o eresolve_tac ([CollectE, conjE, exE] @
+          @{thms GrpE relcomppE conversepE CollectE Collect_split_in_rel_leE}),
+        hyp_subst_tac ctxt,
+        rtac bexI, rtac conjI, rtac trans, rtac map_comp0, rtac trans, rtac map_cong0,
+        REPEAT_DETERM_N m o rtac (@{thm id_o} RS fun_cong),
+        REPEAT_DETERM_N n o rtac (@{thm o_id} RS fun_cong),
+        atac, rtac trans, rtac map_comp0, rtac trans, rtac map_cong0,
+        REPEAT_DETERM_N m o rtac (@{thm id_o} RS fun_cong),
+        REPEAT_DETERM_N n o rtac (@{thm o_id} RS fun_cong),
+        rtac trans, rtac map_cong0,
+        REPEAT_DETERM_N m o EVERY' [rtac @{thm Id_onD'}, etac set_mp, atac],
+        REPEAT_DETERM_N n o rtac refl,
+        atac, rtac CollectI,
+        CONJ_WRAP' (fn (i, thm) =>
+          if i <= m
+          then EVERY' [rtac ord_eq_le_trans, rtac thm, rtac @{thm image_subsetI},
+            rtac @{thm Id_on_fst}, etac set_mp, atac]
+          else EVERY' [rtac ord_eq_le_trans, rtac trans, rtac thm,
+            rtac trans_fun_cong_image_id_id_apply, atac])
+        (1 upto (m + n) ~~ set_map0s)];
+  in
+    EVERY' [rtac (bis_def RS trans),
+      rtac iffI, etac conjE, etac conjI, CONJ_WRAP' mk_if_tac thms,
+      etac conjE, etac conjI, CONJ_WRAP' mk_only_if_tac thms] 1
+  end;
+
+fun mk_bis_converse_tac m bis_rel rel_congs rel_converseps =
+  EVERY' [stac bis_rel, dtac (bis_rel RS iffD1),
+    REPEAT_DETERM o etac conjE, rtac conjI,
+    CONJ_WRAP' (K (EVERY' [rtac converse_shift, etac subset_trans,
+      rtac equalityD2, rtac @{thm converse_Times}])) rel_congs,
+    CONJ_WRAP' (fn (rel_cong, rel_conversep) =>
+      EVERY' [rtac allI, rtac allI, rtac impI,
+        rtac (rel_cong RS @{thm eq_refl} RS @{thm predicate2D}),
+        REPEAT_DETERM_N m o rtac conversep_in_rel_Id_on,
+        REPEAT_DETERM_N (length rel_congs) o rtac @{thm conversep_in_rel},
+        rtac (rel_conversep RS sym RS @{thm eq_refl} RS @{thm predicate2D}),
+        REPEAT_DETERM o etac allE,
+        rtac @{thm conversepI}, etac mp, etac @{thm converseD}]) (rel_congs ~~ rel_converseps)] 1;
+
+fun mk_bis_O_tac ctxt m bis_rel rel_congs rel_OOs =
+  EVERY' [stac bis_rel, REPEAT_DETERM o dtac (bis_rel RS iffD1),
+    REPEAT_DETERM o etac conjE, rtac conjI,
+    CONJ_WRAP' (K (EVERY' [etac @{thm relcomp_subset_Sigma}, atac])) rel_congs,
+    CONJ_WRAP' (fn (rel_cong, rel_OO) =>
+      EVERY' [rtac allI, rtac allI, rtac impI,
+        rtac (rel_cong RS @{thm eq_refl} RS @{thm predicate2D}),
+        REPEAT_DETERM_N m o rtac relcompp_in_rel_Id_on,
+        REPEAT_DETERM_N (length rel_congs) o rtac @{thm relcompp_in_rel},
+        rtac (rel_OO RS sym RS @{thm eq_refl} RS @{thm predicate2D}),
+        etac @{thm relcompE},
+        REPEAT_DETERM o dtac Pair_eqD,
+        etac conjE, hyp_subst_tac ctxt,
+        REPEAT_DETERM o etac allE, rtac @{thm relcomppI},
+        etac mp, atac, etac mp, atac]) (rel_congs ~~ rel_OOs)] 1;
+
+fun mk_bis_Gr_tac bis_rel rel_Grps mor_images morEs coalg_ins
+  {context = ctxt, prems = _} =
+  unfold_thms_tac ctxt (bis_rel :: @{thm Id_on_Gr} :: @{thm in_rel_Gr} :: rel_Grps) THEN
+  EVERY' [rtac conjI,
+    CONJ_WRAP' (fn thm => rtac (@{thm Gr_incl} RS ssubst) THEN' etac thm) mor_images,
+    CONJ_WRAP' (fn (coalg_in, morE) =>
+      EVERY' [rtac allI, rtac allI, rtac impI, rtac @{thm GrpI}, etac (morE RS trans),
+        etac @{thm GrD1}, etac (@{thm GrD2} RS arg_cong), etac coalg_in, etac @{thm GrD1}])
+    (coalg_ins ~~ morEs)] 1;
+
+fun mk_bis_Union_tac bis_def in_monos {context = ctxt, prems = _} =
+  let
+    val n = length in_monos;
+    val ks = 1 upto n;
+  in
+    unfold_thms_tac ctxt [bis_def] THEN
+    EVERY' [rtac conjI,
+      CONJ_WRAP' (fn i =>
+        EVERY' [rtac @{thm UN_least}, dtac bspec, atac,
+          dtac conjunct1, etac (mk_conjunctN n i)]) ks,
+      CONJ_WRAP' (fn (i, in_mono) =>
+        EVERY' [rtac allI, rtac allI, rtac impI, etac @{thm UN_E}, dtac bspec, atac,
+          dtac conjunct2, dtac (mk_conjunctN n i), etac allE, etac allE, dtac mp,
+          atac, etac bexE, rtac bexI, atac, rtac in_mono,
+          REPEAT_DETERM_N n o etac @{thm SUP_upper2[OF _ subset_refl]},
+          atac]) (ks ~~ in_monos)] 1
+  end;
+
+fun mk_sbis_lsbis_tac ctxt lsbis_defs bis_Union bis_cong =
+  let
+    val n = length lsbis_defs;
+  in
+    EVERY' [rtac (Thm.permute_prems 0 1 bis_cong), EVERY' (map rtac lsbis_defs),
+      rtac bis_Union, rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE, exE],
+      hyp_subst_tac ctxt, etac bis_cong, EVERY' (map (rtac o mk_nth_conv n) (1 upto n))] 1
+  end;
+
+fun mk_incl_lsbis_tac n i lsbis_def =
+  EVERY' [rtac @{thm xt1(3)}, rtac lsbis_def, rtac @{thm SUP_upper2}, rtac CollectI,
+    REPEAT_DETERM_N n o rtac exI, rtac conjI, rtac refl, atac, rtac equalityD2,
+    rtac (mk_nth_conv n i)] 1;
+
+fun mk_equiv_lsbis_tac sbis_lsbis lsbis_incl incl_lsbis bis_Id_on bis_converse bis_O =
+  EVERY' [rtac (@{thm equiv_def} RS iffD2),
+
+    rtac conjI, rtac (@{thm refl_on_def} RS iffD2),
+    rtac conjI, rtac lsbis_incl, rtac ballI, rtac set_mp,
+    rtac incl_lsbis, rtac bis_Id_on, atac, etac @{thm Id_onI},
+
+    rtac conjI, rtac (@{thm sym_def} RS iffD2),
+    rtac allI, rtac allI, rtac impI, rtac set_mp,
+    rtac incl_lsbis, rtac bis_converse, rtac sbis_lsbis, etac @{thm converseI},
+
+    rtac (@{thm trans_def} RS iffD2),
+    rtac allI, rtac allI, rtac allI, rtac impI, rtac impI, rtac set_mp,
+    rtac incl_lsbis, rtac bis_O, rtac sbis_lsbis, rtac sbis_lsbis,
+    etac @{thm relcompI}, atac] 1;
+
+fun mk_coalgT_tac m defs strT_defs set_map0ss {context = ctxt, prems = _} =
+  let
+    val n = length strT_defs;
+    val ks = 1 upto n;
+    fun coalg_tac (i, ((passive_sets, active_sets), def)) =
+      EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE],
+        hyp_subst_tac ctxt, rtac (def RS trans RS @{thm ssubst_mem}), etac (arg_cong RS trans),
+        rtac (mk_sum_casesN n i), rtac CollectI,
+        EVERY' (map (fn thm => EVERY' [rtac conjI, rtac (thm RS ord_eq_le_trans),
+          etac ((trans OF [@{thm image_id} RS fun_cong, id_apply]) RS ord_eq_le_trans)])
+          passive_sets),
+        CONJ_WRAP' (fn (i, thm) => EVERY' [rtac (thm RS ord_eq_le_trans),
+          rtac @{thm image_subsetI}, rtac CollectI, rtac exI, rtac exI, rtac conjI, rtac refl,
+          rtac conjI,
+          rtac conjI, etac @{thm empty_Shift}, dtac set_rev_mp,
+            etac equalityD1, etac CollectD,
+          rtac conjI, etac @{thm Shift_clists},
+          rtac conjI, etac @{thm Shift_prefCl},
+          rtac conjI, rtac ballI,
+            rtac conjI, dtac @{thm iffD1[OF ball_conj_distrib]}, dtac conjunct1,
+            SELECT_GOAL (unfold_thms_tac ctxt @{thms Succ_Shift shift_def}),
+            etac bspec, etac @{thm ShiftD},
+            CONJ_WRAP' (fn i => EVERY' [rtac ballI, etac CollectE, dtac @{thm ShiftD},
+              dtac bspec, etac thin_rl, atac, dtac conjunct2, dtac (mk_conjunctN n i),
+              dtac bspec, rtac CollectI, etac @{thm set_mp[OF equalityD1[OF Succ_Shift]]},
+              REPEAT_DETERM o eresolve_tac [exE, conjE], rtac exI,
+              rtac conjI, rtac (@{thm shift_def} RS fun_cong RS trans),
+              rtac (@{thm append_Cons} RS sym RS arg_cong RS trans), atac,
+              REPEAT_DETERM_N m o (rtac conjI THEN' atac),
+              CONJ_WRAP' (K (EVERY' [etac trans, rtac @{thm Collect_cong},
+                rtac @{thm eqset_imp_iff}, rtac sym, rtac trans, rtac @{thm Succ_Shift},
+                rtac (@{thm append_Cons} RS sym RS arg_cong)])) ks]) ks,
+          rtac allI, rtac impI, REPEAT_DETERM o eresolve_tac [allE, impE],
+          etac @{thm not_in_Shift}, rtac trans, rtac (@{thm shift_def} RS fun_cong), atac,
+          dtac bspec, atac, dtac conjunct2, dtac (mk_conjunctN n i), dtac bspec,
+          etac @{thm set_mp[OF equalityD1]}, atac,
+          REPEAT_DETERM o eresolve_tac [exE, conjE], rtac exI,
+          rtac conjI, rtac (@{thm shift_def} RS fun_cong RS trans),
+          etac (@{thm append_Nil} RS sym RS arg_cong RS trans),
+          REPEAT_DETERM_N m o (rtac conjI THEN' atac),
+          CONJ_WRAP' (K (EVERY' [etac trans, rtac @{thm Collect_cong},
+            rtac @{thm eqset_imp_iff}, rtac sym, rtac trans, rtac @{thm Succ_Shift},
+            rtac (@{thm append_Nil} RS sym RS arg_cong)])) ks]) (ks ~~ active_sets)];
+  in
+    unfold_thms_tac ctxt defs THEN
+    CONJ_WRAP' coalg_tac (ks ~~ (map (chop m) set_map0ss ~~ strT_defs)) 1
+  end;
+
+fun mk_Lev_sbd_tac ctxt cts Lev_0s Lev_Sucs to_sbdss =
+  let
+    val n = length Lev_0s;
+    val ks = 1 upto n;
+  in
+    EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
+      REPEAT_DETERM o rtac allI,
+      CONJ_WRAP' (fn Lev_0 =>
+        EVERY' (map rtac [ord_eq_le_trans, Lev_0, @{thm Nil_clists}])) Lev_0s,
+      REPEAT_DETERM o rtac allI,
+      CONJ_WRAP' (fn (Lev_Suc, to_sbds) =>
+        EVERY' [rtac ord_eq_le_trans, rtac Lev_Suc,
+          CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least}))
+            (fn (i, to_sbd) => EVERY' [rtac subsetI,
+              REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac ctxt,
+              rtac @{thm Cons_clists}, rtac (mk_InN_Field n i), etac to_sbd,
+              etac set_rev_mp, REPEAT_DETERM o etac allE,
+              etac (mk_conjunctN n i)])
+          (rev (ks ~~ to_sbds))])
+      (Lev_Sucs ~~ to_sbdss)] 1
+  end;
+
+fun mk_length_Lev_tac ctxt cts Lev_0s Lev_Sucs =
+  let
+    val n = length Lev_0s;
+    val ks = n downto 1;
+  in
+    EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
+      REPEAT_DETERM o rtac allI,
+      CONJ_WRAP' (fn Lev_0 =>
+        EVERY' [rtac impI, dtac (Lev_0 RS equalityD1 RS set_mp),
+          etac @{thm singletonE}, etac ssubst, rtac @{thm list.size(3)}]) Lev_0s,
+      REPEAT_DETERM o rtac allI,
+      CONJ_WRAP' (fn Lev_Suc =>
+        EVERY' [rtac impI, dtac (Lev_Suc RS equalityD1 RS set_mp),
+          CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
+            (fn i =>
+              EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac ctxt,
+                rtac trans, rtac @{thm length_Cons}, rtac @{thm arg_cong[of _ _ Suc]},
+                REPEAT_DETERM o etac allE, dtac (mk_conjunctN n i), etac mp, atac]) ks])
+      Lev_Sucs] 1
+  end;
+
+fun mk_length_Lev'_tac length_Lev =
+  EVERY' [ftac length_Lev, etac ssubst, atac] 1;
+
+fun mk_prefCl_Lev_tac ctxt cts Lev_0s Lev_Sucs =
+  let
+    val n = length Lev_0s;
+    val ks = n downto 1;
+  in
+    EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
+      REPEAT_DETERM o rtac allI,
+      CONJ_WRAP' (fn Lev_0 =>
+        EVERY' [rtac impI, etac conjE, dtac (Lev_0 RS equalityD1 RS set_mp),
+          etac @{thm singletonE}, hyp_subst_tac ctxt,
+          dtac @{thm prefixeq_Nil[THEN subst, of "%x. x"]},
+          hyp_subst_tac ctxt,
+          rtac @{thm set_mp[OF equalityD2[OF trans[OF arg_cong[OF list.size(3)]]]]},
+          rtac Lev_0, rtac @{thm singletonI}]) Lev_0s,
+      REPEAT_DETERM o rtac allI,
+      CONJ_WRAP' (fn (Lev_0, Lev_Suc) =>
+        EVERY' [rtac impI, etac conjE, dtac (Lev_Suc RS equalityD1 RS set_mp),
+          CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
+            (fn i =>
+              EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac ctxt,
+              dtac @{thm prefixeq_Cons[THEN subst, of "%x. x"]}, etac disjE, hyp_subst_tac ctxt,
+              rtac @{thm set_mp[OF equalityD2[OF trans[OF arg_cong[OF list.size(3)]]]]},
+              rtac Lev_0, rtac @{thm singletonI},
+              REPEAT_DETERM o eresolve_tac [exE, conjE], hyp_subst_tac ctxt,
+              rtac @{thm set_mp[OF equalityD2[OF trans[OF arg_cong[OF length_Cons]]]]},
+              rtac Lev_Suc, rtac (mk_UnIN n i), rtac CollectI, REPEAT_DETERM o rtac exI, rtac conjI,
+              rtac refl, etac conjI, REPEAT_DETERM o etac allE, dtac (mk_conjunctN n i),
+              etac mp, etac conjI, atac]) ks])
+      (Lev_0s ~~ Lev_Sucs)] 1
+  end;
+
+fun mk_rv_last_tac cTs cts rv_Nils rv_Conss =
+  let
+    val n = length rv_Nils;
+    val ks = 1 upto n;
+  in
+    EVERY' [rtac (Drule.instantiate' cTs cts @{thm list.induct}),
+      REPEAT_DETERM o rtac allI,
+      CONJ_WRAP' (fn rv_Cons =>
+        CONJ_WRAP' (fn (i, rv_Nil) => (EVERY' [rtac exI,
+          rtac (@{thm append_Nil} RS arg_cong RS trans),
+          rtac (rv_Cons RS trans), rtac (mk_sum_casesN n i RS arg_cong RS trans), rtac rv_Nil]))
+        (ks ~~ rv_Nils))
+      rv_Conss,
+      REPEAT_DETERM o rtac allI, rtac (mk_sumEN n),
+      EVERY' (map (fn i =>
+        CONJ_WRAP' (fn rv_Cons => EVERY' [REPEAT_DETERM o etac allE, dtac (mk_conjunctN n i),
+          CONJ_WRAP' (fn i' => EVERY' [dtac (mk_conjunctN n i'), etac exE, rtac exI,
+            rtac (@{thm append_Cons} RS arg_cong RS trans),
+            rtac (rv_Cons RS trans), etac (sum_case_weak_cong RS arg_cong RS trans),
+            rtac (mk_sum_casesN n i RS arg_cong RS trans), atac])
+          ks])
+        rv_Conss)
+      ks)] 1
+  end;
+
+fun mk_set_rv_Lev_tac ctxt m cts Lev_0s Lev_Sucs rv_Nils rv_Conss coalg_setss from_to_sbdss =
+  let
+    val n = length Lev_0s;
+    val ks = 1 upto n;
+  in
+    EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
+      REPEAT_DETERM o rtac allI,
+      CONJ_WRAP' (fn (i, ((Lev_0, rv_Nil), coalg_sets)) =>
+        EVERY' [rtac impI, REPEAT_DETERM o etac conjE,
+          dtac (Lev_0 RS equalityD1 RS set_mp), etac @{thm singletonE}, etac ssubst,
+          rtac (rv_Nil RS arg_cong RS iffD2),
+          rtac (mk_sum_casesN n i RS iffD2),
+          CONJ_WRAP' (fn thm => etac thm THEN' atac) (take m coalg_sets)])
+      (ks ~~ ((Lev_0s ~~ rv_Nils) ~~ coalg_setss)),
+      REPEAT_DETERM o rtac allI,
+      CONJ_WRAP' (fn ((Lev_Suc, rv_Cons), (from_to_sbds, coalg_sets)) =>
+        EVERY' [rtac impI, etac conjE, dtac (Lev_Suc RS equalityD1 RS set_mp),
+          CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
+            (fn (i, (from_to_sbd, coalg_set)) =>
+              EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac ctxt,
+              rtac (rv_Cons RS arg_cong RS iffD2),
+              rtac (mk_sum_casesN n i RS arg_cong RS trans RS iffD2),
+              etac (from_to_sbd RS arg_cong), REPEAT_DETERM o etac allE,
+              dtac (mk_conjunctN n i), etac mp, etac conjI, etac set_rev_mp,
+              etac coalg_set, atac])
+          (rev (ks ~~ (from_to_sbds ~~ drop m coalg_sets)))])
+      ((Lev_Sucs ~~ rv_Conss) ~~ (from_to_sbdss ~~ coalg_setss))] 1
+  end;
+
+fun mk_set_Lev_tac ctxt cts Lev_0s Lev_Sucs rv_Nils rv_Conss from_to_sbdss =
+  let
+    val n = length Lev_0s;
+    val ks = 1 upto n;
+  in
+    EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
+      REPEAT_DETERM o rtac allI,
+      CONJ_WRAP' (fn ((i, (Lev_0, Lev_Suc)), rv_Nil) =>
+        EVERY' [rtac impI, dtac (Lev_0 RS equalityD1 RS set_mp),
+          etac @{thm singletonE}, hyp_subst_tac ctxt,
+          CONJ_WRAP' (fn i' => rtac impI THEN' dtac (sym RS trans) THEN' rtac rv_Nil THEN'
+            (if i = i'
+            then EVERY' [dtac (mk_InN_inject n i), hyp_subst_tac ctxt,
+              CONJ_WRAP' (fn (i'', Lev_0'') =>
+                EVERY' [rtac impI, rtac @{thm ssubst_mem[OF append_Nil]},
+                  rtac (Lev_Suc RS equalityD2 RS set_mp), rtac (mk_UnIN n i''),
+                  rtac CollectI, REPEAT_DETERM o rtac exI, rtac conjI, rtac refl,
+                  etac conjI, rtac (Lev_0'' RS equalityD2 RS set_mp),
+                  rtac @{thm singletonI}])
+              (ks ~~ Lev_0s)]
+            else etac (mk_InN_not_InM i' i RS notE)))
+          ks])
+      ((ks ~~ (Lev_0s ~~ Lev_Sucs)) ~~ rv_Nils),
+      REPEAT_DETERM o rtac allI,
+      CONJ_WRAP' (fn ((Lev_Suc, rv_Cons), from_to_sbds) =>
+        EVERY' [rtac impI, dtac (Lev_Suc RS equalityD1 RS set_mp),
+          CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
+            (fn (i, from_to_sbd) =>
+              EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac ctxt,
+                CONJ_WRAP' (fn i' => rtac impI THEN'
+                  CONJ_WRAP' (fn i'' =>
+                    EVERY' [rtac impI, rtac (Lev_Suc RS equalityD2 RS set_mp),
+                      rtac @{thm ssubst_mem[OF append_Cons]}, rtac (mk_UnIN n i),
+                      rtac CollectI, REPEAT_DETERM o rtac exI, rtac conjI, rtac refl,
+                      rtac conjI, atac, dtac (sym RS trans RS sym),
+                      rtac (rv_Cons RS trans), rtac (mk_sum_casesN n i RS trans),
+                      etac (from_to_sbd RS arg_cong), REPEAT_DETERM o etac allE,
+                      dtac (mk_conjunctN n i), dtac mp, atac,
+                      dtac (mk_conjunctN n i'), dtac mp, atac,
+                      dtac (mk_conjunctN n i''), etac mp, atac])
+                  ks)
+                ks])
+          (rev (ks ~~ from_to_sbds))])
+      ((Lev_Sucs ~~ rv_Conss) ~~ from_to_sbdss)] 1
+  end;
+
+fun mk_set_image_Lev_tac ctxt cts Lev_0s Lev_Sucs rv_Nils rv_Conss from_to_sbdss to_sbd_injss =
+  let
+    val n = length Lev_0s;
+    val ks = 1 upto n;
+  in
+    EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
+      REPEAT_DETERM o rtac allI,
+      CONJ_WRAP' (fn ((i, (Lev_0, Lev_Suc)), rv_Nil) =>
+        EVERY' [rtac impI, dtac (Lev_0 RS equalityD1 RS set_mp),
+          etac @{thm singletonE}, hyp_subst_tac ctxt,
+          CONJ_WRAP' (fn i' => rtac impI THEN'
+            CONJ_WRAP' (fn i'' => rtac impI  THEN' dtac (sym RS trans) THEN' rtac rv_Nil THEN'
+              (if i = i''
+              then EVERY' [dtac @{thm ssubst_mem[OF sym[OF append_Nil]]},
+                dtac (Lev_Suc RS equalityD1 RS set_mp), dtac (mk_InN_inject n i),
+                hyp_subst_tac ctxt,
+                CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
+                  (fn k => REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE] THEN'
+                    dtac list_inject_iffD1 THEN' etac conjE THEN'
+                    (if k = i'
+                    then EVERY' [dtac (mk_InN_inject n k), hyp_subst_tac ctxt, etac imageI]
+                    else etac (mk_InN_not_InM i' k RS notE)))
+                (rev ks)]
+              else etac (mk_InN_not_InM i'' i RS notE)))
+            ks)
+          ks])
+      ((ks ~~ (Lev_0s ~~ Lev_Sucs)) ~~ rv_Nils),
+      REPEAT_DETERM o rtac allI,
+      CONJ_WRAP' (fn ((Lev_Suc, rv_Cons), (from_to_sbds, to_sbd_injs)) =>
+        EVERY' [rtac impI, dtac (Lev_Suc RS equalityD1 RS set_mp),
+          CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
+            (fn (i, (from_to_sbd, to_sbd_inj)) =>
+              REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE] THEN' hyp_subst_tac ctxt THEN'
+              CONJ_WRAP' (fn i' => rtac impI THEN'
+                dtac @{thm ssubst_mem[OF sym[OF append_Cons]]} THEN'
+                dtac (Lev_Suc RS equalityD1 RS set_mp) THEN'
+                CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE)) (fn k =>
+                  REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE] THEN'
+                  dtac list_inject_iffD1 THEN' etac conjE THEN'
+                  (if k = i
+                  then EVERY' [dtac (mk_InN_inject n i),
+                    dtac (Thm.permute_prems 0 2 (to_sbd_inj RS iffD1)),
+                    atac, atac, hyp_subst_tac ctxt] THEN'
+                    CONJ_WRAP' (fn i'' =>
+                      EVERY' [rtac impI, dtac (sym RS trans),
+                        rtac (rv_Cons RS trans), rtac (mk_sum_casesN n i RS arg_cong RS trans),
+                        etac (from_to_sbd RS arg_cong),
+                        REPEAT_DETERM o etac allE,
+                        dtac (mk_conjunctN n i), dtac mp, atac,
+                        dtac (mk_conjunctN n i'), dtac mp, atac,
+                        dtac (mk_conjunctN n i''), etac mp, etac sym])
+                    ks
+                  else etac (mk_InN_not_InM i k RS notE)))
+                (rev ks))
+              ks)
+          (rev (ks ~~ (from_to_sbds ~~ to_sbd_injs)))])
+      ((Lev_Sucs ~~ rv_Conss) ~~ (from_to_sbdss ~~ to_sbd_injss))] 1
+  end;
+
+fun mk_mor_beh_tac m mor_def mor_cong beh_defs carT_defs strT_defs isNode_defs
+  to_sbd_injss from_to_sbdss Lev_0s Lev_Sucs rv_Nils rv_Conss Lev_sbds length_Levs length_Lev's
+  prefCl_Levs rv_lastss set_rv_Levsss set_Levsss set_image_Levsss set_map0ss coalg_setss
+  map_comp_ids map_cong0s map_arg_congs {context = ctxt, prems = _} =
+  let
+    val n = length beh_defs;
+    val ks = 1 upto n;
+
+    fun fbetw_tac (i, (carT_def, (isNode_def, (Lev_0, (rv_Nil, (Lev_sbd,
+      ((length_Lev, length_Lev'), (prefCl_Lev, (rv_lasts, (set_map0s,
+        (coalg_sets, (set_rv_Levss, (set_Levss, set_image_Levss))))))))))))) =
+      EVERY' [rtac ballI, rtac (carT_def RS equalityD2 RS set_mp),
+        rtac CollectI, REPEAT_DETERM o rtac exI, rtac conjI, rtac refl, rtac conjI,
+        rtac conjI,
+          rtac @{thm UN_I}, rtac UNIV_I, rtac (Lev_0 RS equalityD2 RS set_mp),
+          rtac @{thm singletonI},
+        rtac conjI,
+          rtac @{thm UN_least}, rtac Lev_sbd,
+        rtac conjI,
+          rtac @{thm prefCl_UN}, rtac ssubst, rtac @{thm PrefCl_def}, REPEAT_DETERM o rtac allI,
+          rtac impI, etac conjE, rtac exI, rtac conjI, rtac @{thm ord_le_eq_trans},
+          etac @{thm prefixeq_length_le}, etac length_Lev, rtac prefCl_Lev, etac conjI, atac,
+        rtac conjI,
+          rtac ballI, etac @{thm UN_E}, rtac conjI,
+          if n = 1 then K all_tac else rtac (mk_sumEN n),
+          EVERY' (map6 (fn i => fn isNode_def => fn set_map0s =>
+            fn set_rv_Levs => fn set_Levs => fn set_image_Levs =>
+            EVERY' [rtac (mk_disjIN n i), rtac (isNode_def RS ssubst),
+              rtac exI, rtac conjI,
+              (if n = 1 then rtac @{thm if_P} THEN' etac length_Lev'
+              else rtac (@{thm if_P} RS arg_cong RS trans) THEN' etac length_Lev' THEN'
+                etac (sum_case_weak_cong RS trans) THEN' rtac (mk_sum_casesN n i)),
+              EVERY' (map2 (fn set_map0 => fn set_rv_Lev =>
+                EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac (set_map0 RS trans),
+                  rtac trans_fun_cong_image_id_id_apply,
+                  etac set_rv_Lev, TRY o atac, etac conjI, atac])
+              (take m set_map0s) set_rv_Levs),
+              CONJ_WRAP' (fn (set_map0, (set_Lev, set_image_Lev)) =>
+                EVERY' [rtac (set_map0 RS trans), rtac equalityI, rtac @{thm image_subsetI},
+                  rtac CollectI, rtac @{thm SuccI}, rtac @{thm UN_I}, rtac UNIV_I, etac set_Lev,
+                  if n = 1 then rtac refl else atac, atac, rtac subsetI,
+                  REPEAT_DETERM o eresolve_tac [CollectE, @{thm SuccE}, @{thm UN_E}],
+                  rtac set_image_Lev, atac, dtac length_Lev, hyp_subst_tac ctxt, dtac length_Lev',
+                  etac @{thm set_mp[OF equalityD1[OF arg_cong[OF length_append_singleton]]]},
+                  if n = 1 then rtac refl else atac])
+              (drop m set_map0s ~~ (set_Levs ~~ set_image_Levs))])
+          ks isNode_defs set_map0ss set_rv_Levss set_Levss set_image_Levss),
+          CONJ_WRAP' (fn (i, (rv_last, (isNode_def, (set_map0s,
+            (set_rv_Levs, (set_Levs, set_image_Levs)))))) =>
+            EVERY' [rtac ballI,
+              REPEAT_DETERM o eresolve_tac [CollectE, @{thm SuccE}, @{thm UN_E}],
+              rtac (rev_mp OF [rv_last, impI]), etac exE, rtac (isNode_def RS ssubst),
+              rtac exI, rtac conjI,
+              (if n = 1 then rtac @{thm if_P} THEN' etac length_Lev'
+              else rtac (@{thm if_P} RS trans) THEN' etac length_Lev' THEN'
+                etac (sum_case_weak_cong RS trans) THEN' rtac (mk_sum_casesN n i)),
+              EVERY' (map2 (fn set_map0 => fn set_rv_Lev =>
+                EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac (set_map0 RS trans),
+                  rtac trans_fun_cong_image_id_id_apply,
+                  etac set_rv_Lev, TRY o atac, etac conjI, atac])
+              (take m set_map0s) set_rv_Levs),
+              CONJ_WRAP' (fn (set_map0, (set_Lev, set_image_Lev)) =>
+                EVERY' [rtac (set_map0 RS trans), rtac equalityI, rtac @{thm image_subsetI},
+                  rtac CollectI, rtac @{thm SuccI}, rtac @{thm UN_I}, rtac UNIV_I, etac set_Lev,
+                  if n = 1 then rtac refl else atac, atac, rtac subsetI,
+                  REPEAT_DETERM o eresolve_tac [CollectE, @{thm SuccE}, @{thm UN_E}],
+                  REPEAT_DETERM_N 4 o etac thin_rl,
+                  rtac set_image_Lev,
+                  atac, dtac length_Lev, hyp_subst_tac ctxt, dtac length_Lev',
+                  etac @{thm set_mp[OF equalityD1[OF arg_cong[OF length_append_singleton]]]},
+                  if n = 1 then rtac refl else atac])
+              (drop m set_map0s ~~ (set_Levs ~~ set_image_Levs))])
+          (ks ~~ (rv_lasts ~~ (isNode_defs ~~ (set_map0ss ~~
+            (set_rv_Levss ~~ (set_Levss ~~ set_image_Levss)))))),
+        (**)
+          rtac allI, rtac impI, rtac @{thm if_not_P}, rtac notI,
+          etac notE, etac @{thm UN_I[OF UNIV_I]},
+        (*root isNode*)
+          rtac (isNode_def RS ssubst), rtac exI, rtac conjI, rtac (@{thm if_P} RS trans),
+          rtac length_Lev', rtac (Lev_0 RS equalityD2 RS set_mp), rtac @{thm singletonI},
+          CONVERSION (Conv.top_conv
+            (K (Conv.try_conv (Conv.rewr_conv (rv_Nil RS eq_reflection)))) ctxt),
+          if n = 1 then rtac refl else rtac (mk_sum_casesN n i),
+          EVERY' (map2 (fn set_map0 => fn coalg_set =>
+            EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac (set_map0 RS trans),
+              rtac trans_fun_cong_image_id_id_apply, etac coalg_set, atac])
+            (take m set_map0s) (take m coalg_sets)),
+          CONJ_WRAP' (fn (set_map0, (set_Lev, set_image_Lev)) =>
+            EVERY' [rtac (set_map0 RS trans), rtac equalityI, rtac @{thm image_subsetI},
+              rtac CollectI, rtac @{thm SuccI}, rtac @{thm UN_I}, rtac UNIV_I, rtac set_Lev,
+              rtac (Lev_0 RS equalityD2 RS set_mp), rtac @{thm singletonI}, rtac rv_Nil,
+              atac, rtac subsetI,
+              REPEAT_DETERM o eresolve_tac [CollectE, @{thm SuccE}, @{thm UN_E}],
+              rtac set_image_Lev, rtac (Lev_0 RS equalityD2 RS set_mp),
+              rtac @{thm singletonI}, dtac length_Lev',
+              etac @{thm set_mp[OF equalityD1[OF arg_cong[OF
+                trans[OF length_append_singleton arg_cong[of _ _ Suc, OF list.size(3)]]]]]},
+              rtac rv_Nil])
+          (drop m set_map0s ~~ (nth set_Levss (i - 1) ~~ nth set_image_Levss (i - 1)))];
+
+    fun mor_tac (i, (strT_def, (((Lev_0, Lev_Suc), (rv_Nil, rv_Cons)),
+      ((map_comp_id, (map_cong0, map_arg_cong)), (length_Lev', (from_to_sbds, to_sbd_injs)))))) =
+      EVERY' [rtac ballI, rtac sym, rtac trans, rtac strT_def,
+        rtac (@{thm if_P} RS
+          (if n = 1 then map_arg_cong else sum_case_weak_cong) RS trans),
+        rtac (@{thm list.size(3)} RS arg_cong RS trans RS equalityD2 RS set_mp),
+        rtac Lev_0, rtac @{thm singletonI},
+        CONVERSION (Conv.top_conv
+          (K (Conv.try_conv (Conv.rewr_conv (rv_Nil RS eq_reflection)))) ctxt),
+        if n = 1 then K all_tac
+        else (rtac (sum_case_weak_cong RS trans) THEN'
+          rtac (mk_sum_casesN n i) THEN' rtac (mk_sum_casesN n i RS trans)),
+        rtac (map_comp_id RS trans), rtac (map_cong0 OF replicate m refl),
+        EVERY' (map3 (fn i' => fn to_sbd_inj => fn from_to_sbd =>
+          DETERM o EVERY' [rtac trans, rtac o_apply, rtac Pair_eqI, rtac conjI,
+            rtac trans, rtac @{thm Shift_def},
+            rtac equalityI, rtac subsetI, etac thin_rl, etac thin_rl,
+            REPEAT_DETERM o eresolve_tac [CollectE, @{thm UN_E}], dtac length_Lev', dtac asm_rl,
+            etac thin_rl, dtac @{thm set_rev_mp[OF _ equalityD1]},
+            rtac (@{thm length_Cons} RS arg_cong RS trans), rtac Lev_Suc,
+            CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE)) (fn i'' =>
+              EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE],
+                dtac list_inject_iffD1, etac conjE,
+                if i' = i'' then EVERY' [dtac (mk_InN_inject n i'),
+                  dtac (Thm.permute_prems 0 2 (to_sbd_inj RS iffD1)),
+                  atac, atac, hyp_subst_tac ctxt, etac @{thm UN_I[OF UNIV_I]}]
+                else etac (mk_InN_not_InM i' i'' RS notE)])
+            (rev ks),
+            rtac @{thm UN_least}, rtac subsetI, rtac CollectI, rtac @{thm UN_I[OF UNIV_I]},
+            rtac (Lev_Suc RS equalityD2 RS set_mp), rtac (mk_UnIN n i'), rtac CollectI,
+            REPEAT_DETERM o rtac exI, rtac conjI, rtac refl, etac conjI, atac,
+            rtac trans, rtac @{thm shift_def}, rtac ssubst, rtac @{thm fun_eq_iff}, rtac allI,
+            rtac @{thm if_cong}, rtac (@{thm length_Cons} RS arg_cong RS trans), rtac iffI,
+            dtac asm_rl, dtac asm_rl, dtac asm_rl,
+            dtac (Lev_Suc RS equalityD1 RS set_mp),
+            CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE)) (fn i'' =>
+              EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE],
+                dtac list_inject_iffD1, etac conjE,
+                if i' = i'' then EVERY' [dtac (mk_InN_inject n i'),
+                  dtac (Thm.permute_prems 0 2 (to_sbd_inj RS iffD1)),
+                  atac, atac, hyp_subst_tac ctxt, atac]
+                else etac (mk_InN_not_InM i' i'' RS notE)])
+            (rev ks),
+            rtac (Lev_Suc RS equalityD2 RS set_mp), rtac (mk_UnIN n i'), rtac CollectI,
+            REPEAT_DETERM o rtac exI, rtac conjI, rtac refl, etac conjI, atac,
+            CONVERSION (Conv.top_conv
+              (K (Conv.try_conv (Conv.rewr_conv (rv_Cons RS eq_reflection)))) ctxt),
+            if n = 1 then K all_tac
+            else rtac sum_case_weak_cong THEN' rtac (mk_sum_casesN n i' RS trans),
+            SELECT_GOAL (unfold_thms_tac ctxt [from_to_sbd]), rtac refl,
+            rtac refl])
+        ks to_sbd_injs from_to_sbds)];
+  in
+    (rtac mor_cong THEN'
+    EVERY' (map (fn thm => rtac (thm RS ext)) beh_defs) THEN'
+    stac mor_def THEN' rtac conjI THEN'
+    CONJ_WRAP' fbetw_tac
+      (ks ~~ (carT_defs ~~ (isNode_defs ~~ (Lev_0s ~~ (rv_Nils ~~ (Lev_sbds ~~
+        ((length_Levs ~~ length_Lev's) ~~ (prefCl_Levs ~~ (rv_lastss ~~
+          (set_map0ss ~~ (coalg_setss ~~
+            (set_rv_Levsss ~~ (set_Levsss ~~ set_image_Levsss))))))))))))) THEN'
+    CONJ_WRAP' mor_tac
+      (ks ~~ (strT_defs ~~ (((Lev_0s ~~ Lev_Sucs) ~~ (rv_Nils ~~ rv_Conss)) ~~
+        ((map_comp_ids ~~ (map_cong0s ~~ map_arg_congs)) ~~
+          (length_Lev's ~~ (from_to_sbdss ~~ to_sbd_injss))))))) 1
+  end;
+
+fun mk_congruent_str_final_tac m lsbisE map_comp_id map_cong0 equiv_LSBISs =
+  EVERY' [rtac @{thm congruentI}, dtac lsbisE,
+    REPEAT_DETERM o eresolve_tac [CollectE, conjE, bexE], rtac (o_apply RS trans),
+    etac (sym RS arg_cong RS trans), rtac (map_comp_id RS trans),
+    rtac (map_cong0 RS trans), REPEAT_DETERM_N m o rtac refl,
+    EVERY' (map (fn equiv_LSBIS =>
+      EVERY' [rtac @{thm equiv_proj}, rtac equiv_LSBIS, etac set_mp, atac])
+    equiv_LSBISs), rtac sym, rtac (o_apply RS trans),
+    etac (sym RS arg_cong RS trans), rtac map_comp_id] 1;
+
+fun mk_coalg_final_tac m coalg_def congruent_str_finals equiv_LSBISs set_map0ss coalgT_setss =
+  EVERY' [stac coalg_def,
+    CONJ_WRAP' (fn ((set_map0s, coalgT_sets), (equiv_LSBIS, congruent_str_final)) =>
+      EVERY' [rtac @{thm univ_preserves}, rtac equiv_LSBIS, rtac congruent_str_final,
+        rtac ballI, rtac @{thm ssubst_mem}, rtac o_apply, rtac CollectI,
+        EVERY' (map2 (fn set_map0 => fn coalgT_set =>
+          EVERY' [rtac conjI, rtac (set_map0 RS ord_eq_le_trans),
+            rtac ord_eq_le_trans_trans_fun_cong_image_id_id_apply,
+            etac coalgT_set])
+        (take m set_map0s) (take m coalgT_sets)),
+        CONJ_WRAP' (fn (equiv_LSBIS, (set_map0, coalgT_set)) =>
+          EVERY' [rtac (set_map0 RS ord_eq_le_trans),
+            rtac @{thm image_subsetI}, rtac ssubst, rtac @{thm proj_in_iff},
+            rtac equiv_LSBIS, etac set_rev_mp, etac coalgT_set])
+        (equiv_LSBISs ~~ drop m (set_map0s ~~ coalgT_sets))])
+    ((set_map0ss ~~ coalgT_setss) ~~ (equiv_LSBISs ~~ congruent_str_finals))] 1;
+
+fun mk_mor_T_final_tac mor_def congruent_str_finals equiv_LSBISs =
+  EVERY' [stac mor_def, rtac conjI,
+    CONJ_WRAP' (fn equiv_LSBIS =>
+      EVERY' [rtac ballI, rtac ssubst, rtac @{thm proj_in_iff}, rtac equiv_LSBIS, atac])
+    equiv_LSBISs,
+    CONJ_WRAP' (fn (equiv_LSBIS, congruent_str_final) =>
+      EVERY' [rtac ballI, rtac sym, rtac trans, rtac @{thm univ_commute}, rtac equiv_LSBIS,
+        rtac congruent_str_final, atac, rtac o_apply])
+    (equiv_LSBISs ~~ congruent_str_finals)] 1;
+
+fun mk_mor_Rep_tac m defs Reps Abs_inverses coalg_final_setss map_comp_ids map_cong0Ls
+  {context = ctxt, prems = _} =
+  unfold_thms_tac ctxt defs THEN
+  EVERY' [rtac conjI,
+    CONJ_WRAP' (fn thm => rtac ballI THEN' rtac thm) Reps,
+    CONJ_WRAP' (fn (Rep, ((map_comp_id, map_cong0L), coalg_final_sets)) =>
+      EVERY' [rtac ballI, rtac (map_comp_id RS trans), rtac map_cong0L,
+        EVERY' (map2 (fn Abs_inverse => fn coalg_final_set =>
+          EVERY' [rtac ballI, rtac (o_apply RS trans), rtac Abs_inverse,
+            etac set_rev_mp, rtac coalg_final_set, rtac Rep])
+        Abs_inverses (drop m coalg_final_sets))])
+    (Reps ~~ ((map_comp_ids ~~ map_cong0Ls) ~~ coalg_final_setss))] 1;
+
+fun mk_mor_Abs_tac defs Abs_inverses {context = ctxt, prems = _} =
+  unfold_thms_tac ctxt defs THEN
+  EVERY' [rtac conjI,
+    CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) Abs_inverses,
+    CONJ_WRAP' (fn thm => rtac ballI THEN' etac (thm RS arg_cong RS sym)) Abs_inverses] 1;
+
+fun mk_mor_unfold_tac m mor_UNIV dtor_defs unfold_defs Abs_inverses morEs map_comp_ids map_cong0s =
+  EVERY' [rtac iffD2, rtac mor_UNIV,
+    CONJ_WRAP' (fn ((Abs_inverse, morE), ((dtor_def, unfold_def), (map_comp_id, map_cong0))) =>
+      EVERY' [rtac ext, rtac (o_apply RS trans RS sym), rtac (dtor_def RS trans),
+        rtac (unfold_def RS arg_cong RS trans), rtac (Abs_inverse RS arg_cong RS trans),
+        rtac (morE RS arg_cong RS trans), rtac (map_comp_id RS trans),
+        rtac (o_apply RS trans RS sym), rtac map_cong0,
+        REPEAT_DETERM_N m o rtac refl,
+        EVERY' (map (fn thm => rtac (thm RS trans) THEN' rtac (o_apply RS sym)) unfold_defs)])
+    ((Abs_inverses ~~ morEs) ~~ ((dtor_defs ~~ unfold_defs) ~~ (map_comp_ids ~~ map_cong0s)))] 1;
+
+fun mk_raw_coind_tac bis_def bis_cong bis_O bis_converse bis_Gr tcoalg coalgT mor_T_final
+  sbis_lsbis lsbis_incls incl_lsbiss equiv_LSBISs mor_Rep Rep_injects =
+  let
+    val n = length Rep_injects;
+  in
+    EVERY' [rtac rev_mp, ftac (bis_def RS iffD1),
+      REPEAT_DETERM o etac conjE, rtac bis_cong, rtac bis_O, rtac bis_converse,
+      rtac bis_Gr, rtac tcoalg, rtac mor_Rep, rtac bis_O, atac, rtac bis_Gr, rtac tcoalg,
+      rtac mor_Rep, REPEAT_DETERM_N n o etac @{thm relImage_Gr},
+      rtac impI, rtac rev_mp, rtac bis_cong, rtac bis_O, rtac bis_Gr, rtac coalgT,
+      rtac mor_T_final, rtac bis_O, rtac sbis_lsbis, rtac bis_converse, rtac bis_Gr, rtac coalgT,
+      rtac mor_T_final, EVERY' (map (fn thm => rtac (thm RS @{thm relInvImage_Gr})) lsbis_incls),
+      rtac impI,
+      CONJ_WRAP' (fn (Rep_inject, (equiv_LSBIS , (incl_lsbis, lsbis_incl))) =>
+        EVERY' [rtac subset_trans, rtac @{thm relInvImage_UNIV_relImage}, rtac subset_trans,
+          rtac @{thm relInvImage_mono}, rtac subset_trans, etac incl_lsbis,
+          rtac ord_eq_le_trans, rtac @{thm sym[OF relImage_relInvImage]},
+          rtac @{thm xt1(3)}, rtac @{thm Sigma_cong},
+          rtac @{thm proj_image}, rtac @{thm proj_image}, rtac lsbis_incl,
+          rtac subset_trans, rtac @{thm relImage_mono}, rtac incl_lsbis, atac,
+          rtac @{thm relImage_proj}, rtac equiv_LSBIS, rtac @{thm relInvImage_Id_on},
+          rtac Rep_inject])
+      (Rep_injects ~~ (equiv_LSBISs ~~ (incl_lsbiss ~~ lsbis_incls)))] 1
+  end;
+
+fun mk_unfold_unique_mor_tac raw_coinds bis mor unfold_defs =
+  CONJ_WRAP' (fn (raw_coind, unfold_def) =>
+    EVERY' [rtac ext, etac (bis RS raw_coind RS set_mp RS @{thm IdD}), rtac mor,
+      rtac @{thm image2_eqI}, rtac refl, rtac (unfold_def RS arg_cong RS trans),
+      rtac (o_apply RS sym), rtac UNIV_I]) (raw_coinds ~~ unfold_defs) 1;
+
+fun mk_dtor_o_ctor_tac ctor_def unfold map_comp_id map_cong0L unfold_o_dtors
+  {context = ctxt, prems = _} =
+  unfold_thms_tac ctxt [ctor_def] THEN EVERY' [rtac ext, rtac trans, rtac o_apply,
+    rtac trans, rtac unfold, rtac trans, rtac map_comp_id, rtac trans, rtac map_cong0L,
+    EVERY' (map (fn thm =>
+      rtac ballI THEN' rtac (trans OF [thm RS fun_cong, id_apply])) unfold_o_dtors),
+    rtac sym, rtac id_apply] 1;
+
+fun mk_corec_tac m corec_defs unfold map_cong0 corec_Inls {context = ctxt, prems = _} =
+  unfold_thms_tac ctxt corec_defs THEN EVERY' [rtac trans, rtac (o_apply RS arg_cong),
+    rtac trans, rtac unfold, fo_rtac (@{thm sum.cases(2)} RS arg_cong RS trans) ctxt, rtac map_cong0,
+    REPEAT_DETERM_N m o rtac refl,
+    EVERY' (map (fn thm => rtac @{thm sum_case_expand_Inr} THEN' rtac thm) corec_Inls)] 1;
+
+fun mk_corec_unique_mor_tac corec_defs corec_Inls unfold_unique_mor {context = ctxt, prems = _} =
+  unfold_thms_tac ctxt
+    (corec_defs @ map (fn thm => thm RS @{thm sum_case_expand_Inr'}) corec_Inls) THEN
+  etac unfold_unique_mor 1;
+
+fun mk_dtor_coinduct_tac m raw_coind bis_rel rel_congs =
+  EVERY' [rtac rev_mp, rtac raw_coind, rtac ssubst, rtac bis_rel, rtac conjI,
+    CONJ_WRAP' (K (rtac @{thm ord_le_eq_trans[OF subset_UNIV UNIV_Times_UNIV[THEN sym]]}))
+    rel_congs,
+    CONJ_WRAP' (fn rel_cong => EVERY' [rtac allI, rtac allI, rtac impI,
+      REPEAT_DETERM o etac allE, rtac (rel_cong RS @{thm eq_refl} RS @{thm predicate2D}),
+      REPEAT_DETERM_N m o rtac @{thm in_rel_Id_on_UNIV[symmetric]},
+      REPEAT_DETERM_N (length rel_congs) o rtac @{thm in_rel_Collect_split_eq[symmetric]},
+      etac mp, etac CollectE, etac @{thm splitD}])
+    rel_congs,
+    rtac impI, REPEAT_DETERM o etac conjE,
+    CONJ_WRAP' (K (EVERY' [rtac impI, rtac @{thm IdD}, etac set_mp,
+      rtac CollectI, etac @{thm prod_caseI}])) rel_congs] 1;
+
+fun mk_dtor_map_coinduct_tac m ks raw_coind bis_def =
+  let
+    val n = length ks;
+  in
+    EVERY' [rtac rev_mp, rtac raw_coind, rtac ssubst, rtac bis_def, rtac conjI,
+      CONJ_WRAP' (K (rtac @{thm ord_le_eq_trans[OF subset_UNIV UNIV_Times_UNIV[THEN sym]]})) ks,
+      CONJ_WRAP' (fn i => EVERY' [select_prem_tac n (dtac asm_rl) i, REPEAT_DETERM o rtac allI,
+        rtac impI, REPEAT_DETERM o dtac @{thm meta_spec}, etac CollectE, etac @{thm meta_impE},
+        atac, etac exE, etac conjE, etac conjE, rtac bexI, rtac conjI,
+        etac @{thm fst_conv[THEN subst]}, etac @{thm snd_conv[THEN subst]},
+        rtac CollectI, REPEAT_DETERM_N m o (rtac conjI THEN' rtac subset_UNIV),
+        CONJ_WRAP' (fn i' => EVERY' [rtac subsetI, rtac CollectI, dtac (mk_conjunctN n i'),
+          REPEAT_DETERM o etac allE, etac mp, rtac @{thm ssubst_mem[OF pair_collapse]}, atac])
+        ks])
+      ks,
+      rtac impI,
+      CONJ_WRAP' (fn i => EVERY' [rtac impI, dtac (mk_conjunctN n i),
+        rtac @{thm subst[OF pair_in_Id_conv]}, etac set_mp,
+        rtac CollectI, etac (refl RSN (2, @{thm subst_Pair}))]) ks] 1
+  end;
+
+fun mk_map_tac m n cT unfold map_comp map_cong0 =
+  EVERY' [rtac ext, rtac (o_apply RS trans RS sym), rtac (o_apply RS trans RS sym),
+    rtac (unfold RS trans), rtac (Thm.permute_prems 0 1 (map_comp RS box_equals)), rtac map_cong0,
+    REPEAT_DETERM_N m o rtac (@{thm id_o} RS fun_cong),
+    REPEAT_DETERM_N n o rtac (@{thm o_id} RS fun_cong),
+    rtac (o_apply RS (Drule.instantiate' [cT] [] arg_cong) RS sym)] 1;
+
+fun mk_set_le_tac n hset_minimal set_hsets set_hset_hsetss =
+  EVERY' [rtac hset_minimal,
+    REPEAT_DETERM_N n o rtac @{thm Un_upper1},
+    REPEAT_DETERM_N n o
+      EVERY' (map3 (fn i => fn set_hset => fn set_hset_hsets =>
+        EVERY' [rtac subsetI, rtac @{thm UnI2}, rtac (mk_UnIN n i), etac @{thm UN_I},
+          etac UnE, etac set_hset, REPEAT_DETERM_N (n - 1) o etac UnE,
+          EVERY' (map (fn thm => EVERY' [etac @{thm UN_E}, etac thm, atac]) set_hset_hsets)])
+      (1 upto n) set_hsets set_hset_hsetss)] 1;
+
+fun mk_set_ge_tac n set_incl_hset set_hset_incl_hsets =
+  EVERY' [rtac @{thm Un_least}, rtac set_incl_hset,
+    REPEAT_DETERM_N (n - 1) o rtac @{thm Un_least},
+    EVERY' (map (fn thm => rtac @{thm UN_least} THEN' etac thm) set_hset_incl_hsets)] 1;
+
+fun mk_map_id0_tac maps unfold_unique unfold_dtor =
+  EVERY' [rtac (unfold_unique RS trans), EVERY' (map (rtac o mk_sym) maps),
+    rtac unfold_dtor] 1;
+
+fun mk_map_comp0_tac maps map_comp0s map_unique =
+  EVERY' [rtac map_unique,
+    EVERY' (map2 (fn map_thm => fn map_comp0 =>
+      EVERY' (map rtac
+        [@{thm o_assoc} RS trans,
+        @{thm arg_cong2[of _ _ _ _ "op o"]} OF [map_thm, refl] RS trans,
+        @{thm o_assoc} RS sym RS trans, map_thm RS arg_cong RS trans, @{thm o_assoc} RS trans,
+        @{thm arg_cong2[of _ _ _ _ "op o"]} OF [map_comp0 RS sym, refl]]))
+    maps map_comp0s)] 1;
+
+fun mk_mcong_tac ctxt m coinduct_tac map_comps dtor_maps map_cong0s set_map0ss set_hsetss
+  set_hset_hsetsss =
+  let
+    val n = length map_comps;
+    val ks = 1 upto n;
+  in
+    EVERY' ([rtac rev_mp,
+      coinduct_tac] @
+      maps (fn (((((map_comp_trans, dtor_maps_trans), map_cong0), set_map0s), set_hsets),
+        set_hset_hsetss) =>
+        [REPEAT_DETERM o eresolve_tac [exE, conjE], hyp_subst_tac ctxt, rtac exI, rtac conjI,
+         rtac conjI, rtac map_comp_trans, rtac sym, rtac dtor_maps_trans, rtac map_cong0,
+         REPEAT_DETERM_N m o (rtac o_apply_trans_sym THEN' rtac id_apply),
+         REPEAT_DETERM_N n o rtac fst_convol_fun_cong_sym,
+         rtac map_comp_trans, rtac sym, rtac dtor_maps_trans, rtac map_cong0,
+         EVERY' (maps (fn set_hset =>
+           [rtac o_apply_trans_sym, rtac (id_apply RS trans), etac CollectE,
+           REPEAT_DETERM o etac conjE, etac bspec, etac set_hset]) set_hsets),
+         REPEAT_DETERM_N n o rtac snd_convol_fun_cong_sym,
+         CONJ_WRAP' (fn (set_map0, set_hset_hsets) =>
+           EVERY' [REPEAT_DETERM o rtac allI, rtac impI, rtac @{thm image_convolD},
+             etac set_rev_mp, rtac ord_eq_le_trans, rtac set_map0,
+             rtac @{thm image_mono}, rtac subsetI, rtac CollectI, etac CollectE,
+             REPEAT_DETERM o etac conjE,
+             CONJ_WRAP' (fn set_hset_hset =>
+               EVERY' [rtac ballI, etac bspec, etac set_hset_hset, atac]) set_hset_hsets])
+           (drop m set_map0s ~~ set_hset_hsetss)])
+        (map (fn th => th RS trans) map_comps ~~ map (fn th => th RS trans) dtor_maps ~~
+          map_cong0s ~~ set_map0ss ~~ set_hsetss ~~ set_hset_hsetsss) @
+      [rtac impI,
+       CONJ_WRAP' (fn k =>
+         EVERY' [rtac impI, dtac (mk_conjunctN n k), etac mp, rtac exI, rtac conjI, etac CollectI,
+           rtac conjI, rtac refl, rtac refl]) ks]) 1
+  end
+
+fun mk_dtor_map_unique_tac unfold_unique sym_map_comps ctxt =
+  rtac unfold_unique 1 THEN
+  unfold_thms_tac ctxt (sym_map_comps @ @{thms o_assoc id_o o_id}) THEN
+  ALLGOALS (etac sym);
+
+fun mk_col_natural_tac cts rec_0s rec_Sucs dtor_maps set_map0ss
+  {context = ctxt, prems = _} =
+  let
+    val n = length dtor_maps;
+  in
+    EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
+      REPEAT_DETERM o rtac allI, SELECT_GOAL (unfold_thms_tac ctxt rec_0s),
+      CONJ_WRAP' (K (rtac @{thm image_empty})) rec_0s,
+      REPEAT_DETERM o rtac allI,
+      CONJ_WRAP' (fn (rec_Suc, (dtor_map, set_nats)) => EVERY'
+        [SELECT_GOAL (unfold_thms_tac ctxt
+          (rec_Suc :: dtor_map :: set_nats @ @{thms image_Un image_UN UN_simps(10)})),
+        rtac Un_cong, rtac refl,
+        CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 Un_cong))
+          (fn i => EVERY' [rtac @{thm UN_cong[OF refl]},
+            REPEAT_DETERM o etac allE, etac (mk_conjunctN n i)]) (n downto 1)])
+      (rec_Sucs ~~ (dtor_maps ~~ set_map0ss))] 1
+  end;
+
+fun mk_set_map0_tac hset_def col_natural =
+  EVERY' (map rtac [ext, (o_apply RS trans), (hset_def RS trans), sym,
+    (o_apply RS trans), (@{thm image_cong} OF [hset_def, refl] RS trans),
+    (@{thm image_UN} RS trans), (refl RS @{thm UN_cong}), col_natural]) 1;
+
+fun mk_col_bd_tac m j cts rec_0s rec_Sucs sbd_Card_order sbd_Cinfinite set_sbdss =
+  let
+    val n = length rec_0s;
+  in
+    EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
+      REPEAT_DETERM o rtac allI,
+      CONJ_WRAP' (fn rec_0 => EVERY' (map rtac [ordIso_ordLeq_trans,
+        @{thm card_of_ordIso_subst}, rec_0, @{thm Card_order_empty}, sbd_Card_order])) rec_0s,
+      REPEAT_DETERM o rtac allI,
+      CONJ_WRAP' (fn (rec_Suc, set_sbds) => EVERY'
+        [rtac ordIso_ordLeq_trans, rtac @{thm card_of_ordIso_subst}, rtac rec_Suc,
+        rtac (sbd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})), rtac (nth set_sbds (j - 1)),
+        REPEAT_DETERM_N (n - 1) o rtac (sbd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})),
+        EVERY' (map2 (fn i => fn set_sbd => EVERY' [rtac @{thm UNION_Cinfinite_bound},
+          rtac set_sbd, rtac ballI, REPEAT_DETERM o etac allE,
+          etac (mk_conjunctN n i), rtac sbd_Cinfinite]) (1 upto n) (drop m set_sbds))])
+      (rec_Sucs ~~ set_sbdss)] 1
+  end;
+
+fun mk_set_bd_tac sbd_Cinfinite hset_def col_bd =
+  EVERY' (map rtac [ordIso_ordLeq_trans, @{thm card_of_ordIso_subst}, hset_def,
+    @{thm UNION_Cinfinite_bound}, ordIso_ordLeq_trans, @{thm card_of_nat},
+    @{thm natLeq_ordLeq_cinfinite}, sbd_Cinfinite, ballI, col_bd, sbd_Cinfinite]) 1;
+
+fun mk_le_rel_OO_tac coinduct rel_Jrels rel_OOs =
+  EVERY' (rtac coinduct :: map2 (fn rel_Jrel => fn rel_OO =>
+    let val Jrel_imp_rel = rel_Jrel RS iffD1;
+    in
+      EVERY' [rtac (rel_OO RS sym RS @{thm eq_refl} RS @{thm predicate2D}), etac @{thm relcomppE},
+      rtac @{thm relcomppI}, etac Jrel_imp_rel, etac Jrel_imp_rel]
+    end)
+  rel_Jrels rel_OOs) 1;
+
+fun mk_wit_tac n dtor_ctors dtor_set wit coind_wits ctxt =
+  ALLGOALS (TRY o (eresolve_tac coind_wits THEN' rtac refl)) THEN
+  REPEAT_DETERM (atac 1 ORELSE
+    EVERY' [dtac set_rev_mp, rtac equalityD1, resolve_tac dtor_set,
+    K (unfold_thms_tac ctxt dtor_ctors),
+    REPEAT_DETERM_N n o etac UnE,
+    REPEAT_DETERM o
+      (TRY o REPEAT_DETERM o etac UnE THEN' TRY o etac @{thm UN_E} THEN'
+        (eresolve_tac wit ORELSE'
+        (dresolve_tac wit THEN'
+          (etac FalseE ORELSE'
+          EVERY' [hyp_subst_tac ctxt, dtac set_rev_mp, rtac equalityD1, resolve_tac dtor_set,
+            K (unfold_thms_tac ctxt dtor_ctors), REPEAT_DETERM_N n o etac UnE]))))] 1);
+
+fun mk_coind_wit_tac induct unfolds set_nats wits {context = ctxt, prems = _} =
+  rtac induct 1 THEN ALLGOALS (TRY o rtac impI THEN' TRY o hyp_subst_tac ctxt) THEN
+  unfold_thms_tac ctxt (unfolds @ set_nats @ @{thms image_id id_apply}) THEN
+  ALLGOALS (REPEAT_DETERM o etac imageE THEN' TRY o hyp_subst_tac ctxt) THEN
+  ALLGOALS (TRY o
+    FIRST' [rtac TrueI, rtac refl, etac (refl RSN (2, mp)), dresolve_tac wits THEN' etac FalseE])
+
+fun mk_dtor_rel_tac ctxt in_Jrels i in_rel map_comp0 map_cong0 dtor_map dtor_sets dtor_inject
+    dtor_ctor set_map0s dtor_set_incls dtor_set_set_inclss =
+  let
+    val m = length dtor_set_incls;
+    val n = length dtor_set_set_inclss;
+    val (passive_set_map0s, active_set_map0s) = chop m set_map0s;
+    val in_Jrel = nth in_Jrels (i - 1);
+    val if_tac =
+      EVERY' [dtac (in_Jrel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
+        rtac (in_rel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
+        EVERY' (map2 (fn set_map0 => fn set_incl =>
+          EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac set_map0,
+            rtac ord_eq_le_trans, rtac trans_fun_cong_image_id_id_apply,
+            etac (set_incl RS @{thm subset_trans})])
+        passive_set_map0s dtor_set_incls),
+        CONJ_WRAP' (fn (in_Jrel, (set_map0, dtor_set_set_incls)) =>
+          EVERY' [rtac ord_eq_le_trans, rtac set_map0, rtac @{thm image_subsetI}, rtac CollectI,
+            rtac @{thm prod_caseI}, rtac (in_Jrel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
+            CONJ_WRAP' (fn thm => etac (thm RS @{thm subset_trans}) THEN' atac) dtor_set_set_incls,
+            rtac conjI, rtac refl, rtac refl])
+        (in_Jrels ~~ (active_set_map0s ~~ dtor_set_set_inclss)),
+        CONJ_WRAP' (fn conv =>
+          EVERY' [rtac trans, rtac map_comp0, rtac trans, rtac map_cong0,
+          REPEAT_DETERM_N m o rtac @{thm fun_cong[OF o_id]},
+          REPEAT_DETERM_N n o EVERY' (map rtac [trans, o_apply, conv]),
+          rtac trans, rtac sym, rtac dtor_map, rtac (dtor_inject RS iffD2), atac])
+        @{thms fst_conv snd_conv}];
+    val only_if_tac =
+      EVERY' [dtac (in_rel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
+        rtac (in_Jrel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
+        CONJ_WRAP' (fn (dtor_set, passive_set_map0) =>
+          EVERY' [rtac ord_eq_le_trans, rtac dtor_set, rtac @{thm Un_least},
+            rtac ord_eq_le_trans, rtac box_equals, rtac passive_set_map0,
+            rtac (dtor_ctor RS sym RS arg_cong), rtac trans_fun_cong_image_id_id_apply, atac,
+            CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least}))
+              (fn (active_set_map0, in_Jrel) => EVERY' [rtac ord_eq_le_trans,
+                rtac @{thm UN_cong[OF _ refl]}, rtac @{thm box_equals[OF _ _ refl]},
+                rtac active_set_map0, rtac (dtor_ctor RS sym RS arg_cong), rtac @{thm UN_least},
+                dtac set_rev_mp, etac @{thm image_mono}, etac imageE,
+                dtac @{thm ssubst_mem[OF pair_collapse]},
+                REPEAT_DETERM o eresolve_tac (CollectE :: conjE ::
+                  @{thms prod_caseE iffD1[OF Pair_eq, elim_format]}),
+                hyp_subst_tac ctxt,
+                dtac (in_Jrel RS iffD1),
+                dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE,
+                TRY o
+                  EVERY' [dtac (Thm.permute_prems 0 1 @{thm ssubst_mem}), atac, hyp_subst_tac ctxt],
+                REPEAT_DETERM o eresolve_tac [CollectE, conjE], atac])
+            (rev (active_set_map0s ~~ in_Jrels))])
+        (dtor_sets ~~ passive_set_map0s),
+        rtac conjI,
+        REPEAT_DETERM_N 2 o EVERY'[rtac (dtor_inject RS iffD1), rtac trans, rtac dtor_map,
+          rtac box_equals, rtac map_comp0, rtac (dtor_ctor RS sym RS arg_cong), rtac trans,
+          rtac map_cong0, REPEAT_DETERM_N m o rtac @{thm fun_cong[OF o_id]},
+          EVERY' (map (fn in_Jrel => EVERY' [rtac trans, rtac o_apply, dtac set_rev_mp, atac,
+            dtac @{thm ssubst_mem[OF pair_collapse]},
+            REPEAT_DETERM o
+              eresolve_tac (CollectE :: conjE :: @{thms prod_caseE iffD1[OF Pair_eq, elim_format]}),
+            hyp_subst_tac ctxt, dtac (in_Jrel RS iffD1),
+            dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE, atac]) in_Jrels),
+          atac]]
+  in
+    EVERY' [rtac iffI, if_tac, only_if_tac] 1
+  end;
+
+fun mk_rel_coinduct_coind_tac m coinduct ks map_comp0s map_congs map_arg_congs set_map0ss 
+  dtor_unfolds dtor_maps {context = ctxt, prems = _} =
+  let val n = length ks;
+  in
+    EVERY' [DETERM o rtac coinduct,
+      EVERY' (map7 (fn i => fn map_comp0 => fn map_cong => fn map_arg_cong => fn set_map0s =>
+          fn dtor_unfold => fn dtor_map =>
+        EVERY' [REPEAT_DETERM o eresolve_tac [exE, conjE],
+          select_prem_tac (length ks) (dtac @{thm spec2}) i, dtac mp, atac,
+          REPEAT_DETERM o eresolve_tac [CollectE, conjE], hyp_subst_tac ctxt,
+          rtac exI, rtac conjI, rtac conjI,
+          rtac (map_comp0 RS trans), rtac (dtor_map RS trans RS sym),
+          rtac (dtor_unfold RS map_arg_cong RS trans), rtac (trans OF [map_comp0, map_cong]),
+          REPEAT_DETERM_N m o rtac @{thm trans[OF fun_cong[OF o_id] sym[OF fun_cong[OF id_o]]]},
+          REPEAT_DETERM_N n o (rtac @{thm sym[OF trans[OF o_apply]]} THEN' rtac @{thm fst_conv}),
+          rtac (map_comp0 RS trans), rtac (map_cong RS trans),
+          REPEAT_DETERM_N m o rtac @{thm fun_cong[OF id_o]},
+          REPEAT_DETERM_N n o (rtac @{thm trans[OF o_apply]} THEN' rtac @{thm snd_conv}),
+          etac (@{thm prod.cases} RS map_arg_cong RS trans),
+          SELECT_GOAL (unfold_thms_tac ctxt @{thms prod.cases}), 
+          CONJ_WRAP' (fn set_map0 =>
+            EVERY' [REPEAT_DETERM o resolve_tac [allI, impI],
+              dtac (set_map0 RS equalityD1 RS set_mp),
+              REPEAT_DETERM o
+                eresolve_tac (imageE :: conjE :: @{thms iffD1[OF Pair_eq, elim_format]}),
+              hyp_subst_tac ctxt, rtac exI, rtac conjI, etac Collect_splitD_set_mp, atac,
+              rtac (o_apply RS trans), rtac (@{thm surjective_pairing} RS arg_cong)])
+          (drop m set_map0s)])
+      ks map_comp0s map_congs map_arg_congs set_map0ss dtor_unfolds dtor_maps)] 1
+  end;
+
+val split_id_unfolds = @{thms prod.cases image_id id_apply};
+
+fun mk_rel_coinduct_ind_tac m ks unfolds set_map0ss j set_induct {context = ctxt, prems = _} =
+  let val n = length ks;
+  in
+    rtac set_induct 1 THEN
+    EVERY' (map3 (fn unfold => fn set_map0s => fn i =>
+      EVERY' [REPEAT_DETERM o resolve_tac [allI, impI], etac conjE,
+        select_prem_tac n (dtac @{thm spec2}) i, dtac mp, atac,
+        REPEAT_DETERM o eresolve_tac [CollectE, conjE, Collect_splitD_set_mp, set_rev_mp],
+        hyp_subst_tac ctxt,
+        SELECT_GOAL (unfold_thms_tac ctxt ([unfold, nth set_map0s (j - 1)] @ split_id_unfolds)),
+        rtac subset_refl])
+    unfolds set_map0ss ks) 1 THEN
+    EVERY' (map3 (fn unfold => fn set_map0s => fn i =>
+      EVERY' (map (fn set_map0 =>
+        EVERY' [REPEAT_DETERM o resolve_tac [allI, impI], etac conjE,
+        select_prem_tac n (dtac @{thm spec2}) i, dtac mp, atac,
+        REPEAT_DETERM o eresolve_tac [CollectE, conjE], hyp_subst_tac ctxt,
+        SELECT_GOAL (unfold_thms_tac ctxt ([unfold, set_map0] @ split_id_unfolds)),
+        etac imageE, hyp_subst_tac ctxt, REPEAT_DETERM o eresolve_tac [allE, mp],
+        rtac conjI, etac Collect_splitD_set_mp, atac, rtac (@{thm surjective_pairing} RS arg_cong)])
+      (drop m set_map0s)))
+    unfolds set_map0ss ks) 1
+  end;
+
+fun mk_rel_coinduct_tac in_rels in_Jrels helper_indss helper_coind1s helper_coind2s
+  {context = ctxt, prems = CIHs} =
+  let val n = length in_rels;
+  in
+    Method.insert_tac CIHs 1 THEN
+    unfold_thms_tac ctxt (@{thm choice_iff} :: @{thm ex_simps(6)[symmetric]} :: in_rels) THEN
+    REPEAT_DETERM (etac exE 1) THEN
+    CONJ_WRAP' (fn (in_Jrel, (helper_inds, (helper_coind1, helper_coind2))) =>
+      EVERY' [rtac @{thm predicate2I}, rtac (in_Jrel RS iffD2), rtac exI, rtac conjI,
+        if null helper_inds then rtac UNIV_I
+        else rtac CollectI THEN'
+          CONJ_WRAP' (fn helper_ind =>
+            EVERY' [rtac (helper_ind RS rev_mp), REPEAT_DETERM_N n o atac,
+              REPEAT_DETERM_N n o etac thin_rl, rtac impI,
+              REPEAT_DETERM o resolve_tac [subsetI, CollectI, @{thm iffD2[OF split_beta]}],
+              dtac bspec, atac, REPEAT_DETERM o eresolve_tac [allE, mp, conjE],
+              etac (refl RSN (2, conjI))])
+          helper_inds,
+        rtac conjI,
+        rtac (helper_coind1 RS rev_mp), REPEAT_DETERM_N n o atac, REPEAT_DETERM_N n o etac thin_rl,
+        rtac impI, etac mp, rtac exI, etac (refl RSN (2, conjI)),
+        rtac (helper_coind2 RS rev_mp), REPEAT_DETERM_N n o atac, REPEAT_DETERM_N n o etac thin_rl,
+        rtac impI, etac mp, rtac exI, etac (refl RSN (2, conjI))])
+    (in_Jrels ~~ (helper_indss ~~ (helper_coind1s ~~ helper_coind2s))) 1
+  end;
+
+fun mk_unfold_transfer_tac m rel_coinduct map_transfers unfolds {context = ctxt, prems = _} =
+  let
+    val n = length map_transfers;
+  in
+    unfold_thms_tac ctxt
+      @{thms fun_rel_def_butlast all_conj_distrib[symmetric] imp_conjR[symmetric]} THEN
+    unfold_thms_tac ctxt @{thms fun_rel_iff_geq_image2p} THEN
+    HEADGOAL (EVERY'
+      [REPEAT_DETERM o resolve_tac [allI, impI], rtac rel_coinduct,
+      EVERY' (map (fn map_transfer => EVERY'
+        [REPEAT_DETERM o resolve_tac [allI, impI], etac @{thm image2pE}, hyp_subst_tac ctxt,
+        SELECT_GOAL (unfold_thms_tac ctxt unfolds),
+        rtac (funpow (m + n + 1) (fn thm => thm RS @{thm fun_relD}) map_transfer),
+        REPEAT_DETERM_N m o rtac @{thm id_transfer},
+        REPEAT_DETERM_N n o rtac @{thm fun_rel_image2p},
+        etac @{thm predicate2D}, etac @{thm image2pI}])
+      map_transfers)])
+  end;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_gfp_util.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,182 @@
+(*  Title:      HOL/BNF/Tools/bnf_gfp_util.ML
+    Author:     Dmitriy Traytel, TU Muenchen
+    Copyright   2012
+
+Library for the codatatype construction.
+*)
+
+signature BNF_GFP_UTIL =
+sig
+  val mk_rec_simps: int -> thm -> thm list -> thm list list
+
+  val dest_listT: typ -> typ
+
+  val mk_Cons: term -> term -> term
+  val mk_Shift: term -> term -> term
+  val mk_Succ: term -> term -> term
+  val mk_Times: term * term -> term
+  val mk_append: term * term -> term
+  val mk_congruent: term -> term -> term
+  val mk_clists: term -> term
+  val mk_Id_on: term -> term
+  val mk_in_rel: term -> term
+  val mk_equiv: term -> term -> term
+  val mk_fromCard: term -> term -> term
+  val mk_list_rec: term -> term -> term
+  val mk_nat_rec: term -> term -> term
+  val mk_prefCl: term -> term
+  val mk_prefixeq: term -> term -> term
+  val mk_proj: term -> term
+  val mk_quotient: term -> term -> term
+  val mk_shift: term -> term -> term
+  val mk_size: term -> term
+  val mk_toCard: term -> term -> term
+  val mk_undefined: typ -> term
+  val mk_univ: term -> term
+
+  val mk_specN: int -> thm -> thm
+
+  val mk_InN_Field: int -> int -> thm
+  val mk_InN_inject: int -> int -> thm
+  val mk_InN_not_InM: int -> int -> thm
+end;
+
+structure BNF_GFP_Util : BNF_GFP_UTIL =
+struct
+
+open BNF_Util
+
+val mk_append = HOLogic.mk_binop @{const_name append};
+
+fun mk_equiv B R =
+  Const (@{const_name equiv}, fastype_of B --> fastype_of R --> HOLogic.boolT) $ B $ R;
+
+fun mk_Sigma (A, B) =
+  let
+    val AT = fastype_of A;
+    val BT = fastype_of B;
+    val ABT = mk_relT (HOLogic.dest_setT AT, HOLogic.dest_setT (range_type BT));
+  in Const (@{const_name Sigma}, AT --> BT --> ABT) $ A $ B end;
+
+fun mk_Id_on A =
+  let
+    val AT = fastype_of A;
+    val AAT = mk_relT (HOLogic.dest_setT AT, HOLogic.dest_setT AT);
+  in Const (@{const_name Id_on}, AT --> AAT) $ A end;
+
+fun mk_in_rel R =
+  let
+    val ((A, B), RT) = `dest_relT (fastype_of R);
+  in Const (@{const_name in_rel}, RT --> mk_pred2T A B) $ R end;
+
+fun mk_Times (A, B) =
+  let val AT = HOLogic.dest_setT (fastype_of A);
+  in mk_Sigma (A, Term.absdummy AT B) end;
+
+fun dest_listT (Type (@{type_name list}, [T])) = T
+  | dest_listT T = raise TYPE ("dest_setT: set type expected", [T], []);
+
+fun mk_Succ Kl kl =
+  let val T = fastype_of kl;
+  in
+    Const (@{const_name Succ},
+      HOLogic.mk_setT T --> T --> HOLogic.mk_setT (dest_listT T)) $ Kl $ kl
+  end;
+
+fun mk_Shift Kl k =
+  let val T = fastype_of Kl;
+  in
+    Const (@{const_name Shift}, T --> dest_listT (HOLogic.dest_setT T) --> T) $ Kl $ k
+  end;
+
+fun mk_shift lab k =
+  let val T = fastype_of lab;
+  in
+    Const (@{const_name shift}, T --> dest_listT (Term.domain_type T) --> T) $ lab $ k
+  end;
+
+fun mk_prefCl A =
+  Const (@{const_name prefCl}, fastype_of A --> HOLogic.boolT) $ A;
+
+fun mk_prefixeq xs ys =
+  Const (@{const_name prefixeq}, fastype_of xs --> fastype_of ys --> HOLogic.boolT) $ xs $ ys;
+
+fun mk_clists r =
+  let val T = fastype_of r;
+  in Const (@{const_name clists}, T --> mk_relT (`I (HOLogic.listT (fst (dest_relT T))))) $ r end;
+
+fun mk_toCard A r =
+  let
+    val AT = fastype_of A;
+    val rT = fastype_of r;
+  in
+    Const (@{const_name toCard},
+      AT --> rT --> HOLogic.dest_setT AT --> fst (dest_relT rT)) $ A $ r
+  end;
+
+fun mk_fromCard A r =
+  let
+    val AT = fastype_of A;
+    val rT = fastype_of r;
+  in
+    Const (@{const_name fromCard},
+      AT --> rT --> fst (dest_relT rT) --> HOLogic.dest_setT AT) $ A $ r
+  end;
+
+fun mk_Cons x xs =
+  let val T = fastype_of xs;
+  in Const (@{const_name Cons}, dest_listT T --> T --> T) $ x $ xs end;
+
+fun mk_size t = HOLogic.size_const (fastype_of t) $ t;
+
+fun mk_quotient A R =
+  let val T = fastype_of A;
+  in Const (@{const_name quotient}, T --> fastype_of R --> HOLogic.mk_setT T) $ A $ R end;
+
+fun mk_proj R =
+  let val ((AT, BT), T) = `dest_relT (fastype_of R);
+  in Const (@{const_name proj}, T --> AT --> HOLogic.mk_setT BT) $ R end;
+
+fun mk_univ f =
+  let val ((AT, BT), T) = `dest_funT (fastype_of f);
+  in Const (@{const_name univ}, T --> HOLogic.mk_setT AT --> BT) $ f end;
+
+fun mk_congruent R f =
+  Const (@{const_name congruent}, fastype_of R --> fastype_of f --> HOLogic.boolT) $ R $ f;
+
+fun mk_undefined T = Const (@{const_name undefined}, T);
+
+fun mk_nat_rec Zero Suc =
+  let val T = fastype_of Zero;
+  in Const (@{const_name nat_rec}, T --> fastype_of Suc --> HOLogic.natT --> T) $ Zero $ Suc end;
+
+fun mk_list_rec Nil Cons =
+  let
+    val T = fastype_of Nil;
+    val (U, consT) = `(Term.domain_type) (fastype_of Cons);
+  in
+    Const (@{const_name list_rec}, T --> consT --> HOLogic.listT U --> T) $ Nil $ Cons
+  end;
+
+fun mk_InN_not_InM 1 _ = @{thm Inl_not_Inr}
+  | mk_InN_not_InM n m =
+    if n > m then mk_InN_not_InM m n RS @{thm not_sym}
+    else mk_InN_not_InM (n - 1) (m - 1) RS @{thm not_arg_cong_Inr};
+
+fun mk_InN_Field 1 1 = @{thm TrueE[OF TrueI]}
+  | mk_InN_Field _ 1 = @{thm Inl_Field_csum}
+  | mk_InN_Field 2 2 = @{thm Inr_Field_csum}
+  | mk_InN_Field n m = mk_InN_Field (n - 1) (m - 1) RS @{thm Inr_Field_csum};
+
+fun mk_InN_inject 1 _ = @{thm TrueE[OF TrueI]}
+  | mk_InN_inject _ 1 = @{thm Inl_inject}
+  | mk_InN_inject 2 2 = @{thm Inr_inject}
+  | mk_InN_inject n m = @{thm Inr_inject} RS mk_InN_inject (n - 1) (m - 1);
+
+fun mk_specN 0 thm = thm
+  | mk_specN n thm = mk_specN (n - 1) (thm RS spec);
+
+fun mk_rec_simps n rec_thm defs = map (fn i =>
+  map (fn def => def RS rec_thm RS mk_nthI n i RS fun_cong) defs) (1 upto n);
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_lfp.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,1881 @@
+(*  Title:      HOL/BNF/Tools/bnf_lfp.ML
+    Author:     Dmitriy Traytel, TU Muenchen
+    Author:     Andrei Popescu, TU Muenchen
+    Copyright   2012
+
+Datatype construction.
+*)
+
+signature BNF_LFP =
+sig
+  val construct_lfp: mixfix list -> binding list -> binding list -> binding list list ->
+    binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list ->
+    local_theory -> BNF_FP_Util.fp_result * local_theory
+end;
+
+structure BNF_LFP : BNF_LFP =
+struct
+
+open BNF_Def
+open BNF_Util
+open BNF_Tactics
+open BNF_Comp
+open BNF_FP_Util
+open BNF_FP_Def_Sugar
+open BNF_LFP_Rec_Sugar
+open BNF_LFP_Util
+open BNF_LFP_Tactics
+
+(*all BNFs have the same lives*)
+fun construct_lfp mixfixes map_bs rel_bs set_bss0 bs resBs (resDs, Dss) bnfs lthy =
+  let
+    val time = time lthy;
+    val timer = time (Timer.startRealTimer ());
+
+    val live = live_of_bnf (hd bnfs);
+    val n = length bnfs; (*active*)
+    val ks = 1 upto n;
+    val m = live - n; (*passive, if 0 don't generate a new BNF*)
+
+    val note_all = Config.get lthy bnf_note_all;
+    val b_names = map Binding.name_of bs;
+    val b_name = mk_common_name b_names;
+    val b = Binding.name b_name;
+    val mk_internal_b = Binding.name #> Binding.prefix true b_name #> Binding.conceal;
+    fun mk_internal_bs name =
+      map (fn b =>
+        Binding.prefix true b_name (Binding.prefix_name (name ^ "_") b) |> Binding.conceal) bs;
+    val external_bs = map2 (Binding.prefix false) b_names bs
+      |> note_all = false ? map Binding.conceal;
+
+    (* TODO: check if m, n, etc., are sane *)
+
+    val deads = fold (union (op =)) Dss resDs;
+    val names_lthy = fold Variable.declare_typ deads lthy;
+    val passives = map fst (subtract (op = o apsnd TFree) deads resBs);
+
+    (* tvars *)
+    val (((((passiveAs, activeAs), passiveBs), activeBs), passiveCs), activeCs) =
+      names_lthy
+      |> variant_tfrees passives
+      ||>> mk_TFrees n
+      ||>> variant_tfrees passives
+      ||>> mk_TFrees n
+      ||>> variant_tfrees passives
+      ||>> mk_TFrees n
+      |> fst;
+
+    val allAs = passiveAs @ activeAs;
+    val allBs' = passiveBs @ activeBs;
+    val Ass = replicate n allAs;
+    val allBs = passiveAs @ activeBs;
+    val Bss = replicate n allBs;
+    val allCs = passiveAs @ activeCs;
+    val allCs' = passiveBs @ activeCs;
+    val Css' = replicate n allCs';
+
+    (* types *)
+    val dead_poss =
+      map (fn x => if member (op =) deads (TFree x) then SOME (TFree x) else NONE) resBs;
+    fun mk_param NONE passive = (hd passive, tl passive)
+      | mk_param (SOME a) passive = (a, passive);
+    val mk_params = fold_map mk_param dead_poss #> fst;
+
+    fun mk_FTs Ts = map2 (fn Ds => mk_T_of_bnf Ds Ts) Dss bnfs;
+    val (params, params') = `(map Term.dest_TFree) (mk_params passiveAs);
+    val FTsAs = mk_FTs allAs;
+    val FTsBs = mk_FTs allBs;
+    val FTsCs = mk_FTs allCs;
+    val ATs = map HOLogic.mk_setT passiveAs;
+    val BTs = map HOLogic.mk_setT activeAs;
+    val B'Ts = map HOLogic.mk_setT activeBs;
+    val B''Ts = map HOLogic.mk_setT activeCs;
+    val sTs = map2 (curry op -->) FTsAs activeAs;
+    val s'Ts = map2 (curry op -->) FTsBs activeBs;
+    val s''Ts = map2 (curry op -->) FTsCs activeCs;
+    val fTs = map2 (curry op -->) activeAs activeBs;
+    val inv_fTs = map2 (curry op -->) activeBs activeAs;
+    val self_fTs = map2 (curry op -->) activeAs activeAs;
+    val gTs = map2 (curry op -->) activeBs activeCs;
+    val all_gTs = map2 (curry op -->) allBs allCs';
+    val prodBsAs = map2 (curry HOLogic.mk_prodT) activeBs activeAs;
+    val prodFTs = mk_FTs (passiveAs @ prodBsAs);
+    val prod_sTs = map2 (curry op -->) prodFTs activeAs;
+
+    (* terms *)
+    val mapsAsAs = map4 mk_map_of_bnf Dss Ass Ass bnfs;
+    val mapsAsBs = map4 mk_map_of_bnf Dss Ass Bss bnfs;
+    val mapsBsAs = map4 mk_map_of_bnf Dss Bss Ass bnfs;
+    val mapsBsCs' = map4 mk_map_of_bnf Dss Bss Css' bnfs;
+    val mapsAsCs' = map4 mk_map_of_bnf Dss Ass Css' bnfs;
+    val map_fsts = map4 mk_map_of_bnf Dss (replicate n (passiveAs @ prodBsAs)) Bss bnfs;
+    val map_fsts_rev = map4 mk_map_of_bnf Dss Bss (replicate n (passiveAs @ prodBsAs)) bnfs;
+    fun mk_setss Ts = map3 mk_sets_of_bnf (map (replicate live) Dss)
+      (map (replicate live) (replicate n Ts)) bnfs;
+    val setssAs = mk_setss allAs;
+    val bd0s = map3 mk_bd_of_bnf Dss Ass bnfs;
+    val bds =
+      map3 (fn bd0 => fn Ds => fn bnf => mk_csum bd0
+        (mk_card_of (HOLogic.mk_UNIV
+          (mk_T_of_bnf Ds (replicate live (fst (dest_relT (fastype_of bd0)))) bnf))))
+      bd0s Dss bnfs;
+    val witss = map wits_of_bnf bnfs;
+
+    val (((((((((((((((((((zs, zs'), As), Bs), Bs_copy), B's), B''s), ss), prod_ss), s's), s''s),
+      fs), fs_copy), inv_fs), self_fs), gs), all_gs), (xFs, xFs')), (yFs, yFs')),
+      names_lthy) = lthy
+      |> mk_Frees' "z" activeAs
+      ||>> mk_Frees "A" ATs
+      ||>> mk_Frees "B" BTs
+      ||>> mk_Frees "B" BTs
+      ||>> mk_Frees "B'" B'Ts
+      ||>> mk_Frees "B''" B''Ts
+      ||>> mk_Frees "s" sTs
+      ||>> mk_Frees "prods" prod_sTs
+      ||>> mk_Frees "s'" s'Ts
+      ||>> mk_Frees "s''" s''Ts
+      ||>> mk_Frees "f" fTs
+      ||>> mk_Frees "f" fTs
+      ||>> mk_Frees "f" inv_fTs
+      ||>> mk_Frees "f" self_fTs
+      ||>> mk_Frees "g" gTs
+      ||>> mk_Frees "g" all_gTs
+      ||>> mk_Frees' "x" FTsAs
+      ||>> mk_Frees' "y" FTsBs;
+
+    val passive_UNIVs = map HOLogic.mk_UNIV passiveAs;
+    val active_UNIVs = map HOLogic.mk_UNIV activeAs;
+    val prod_UNIVs = map HOLogic.mk_UNIV prodBsAs;
+    val passive_ids = map HOLogic.id_const passiveAs;
+    val active_ids = map HOLogic.id_const activeAs;
+    val fsts = map fst_const prodBsAs;
+
+    (* thms *)
+    val bd0_card_orders = map bd_card_order_of_bnf bnfs;
+    val bd0_Card_orders = map bd_Card_order_of_bnf bnfs;
+    val bd0_Cinfinites = map bd_Cinfinite_of_bnf bnfs;
+    val set_bd0ss = map set_bd_of_bnf bnfs;
+
+    val bd_card_orders =
+      map (fn thm => @{thm card_order_csum} OF [thm, @{thm card_of_card_order_on}]) bd0_card_orders;
+    val bd_Card_order = @{thm Card_order_csum};
+    val bd_Card_orders = replicate n bd_Card_order;
+    val bd_Cinfinites = map (fn thm => thm RS @{thm Cinfinite_csum1}) bd0_Cinfinites;
+    val bd_Cnotzeros = map (fn thm => thm RS @{thm Cinfinite_Cnotzero}) bd_Cinfinites;
+    val bd_Cinfinite = hd bd_Cinfinites;
+    val bd_Cnotzero = hd bd_Cnotzeros;
+    val set_bdss =
+      map2 (fn set_bd0s => fn bd0_Card_order =>
+        map (fn thm => ctrans OF [thm, bd0_Card_order RS @{thm ordLeq_csum1}]) set_bd0s)
+      set_bd0ss bd0_Card_orders;
+    val in_bds = map in_bd_of_bnf bnfs;
+    val sym_map_comps = map (fn bnf => map_comp0_of_bnf bnf RS sym) bnfs;
+    val map_comps = map map_comp_of_bnf bnfs;
+    val map_cong0s = map map_cong0_of_bnf bnfs;
+    val map_id0s = map map_id0_of_bnf bnfs;
+    val map_ids = map map_id_of_bnf bnfs;
+    val set_mapss = map set_map_of_bnf bnfs;
+    val rel_mono_strongs = map rel_mono_strong_of_bnf bnfs;
+    val rel_OOs = map rel_OO_of_bnf bnfs;
+
+    val timer = time (timer "Extracted terms & thms");
+
+    (* nonemptiness check *)
+    fun new_wit X (wit: nonemptiness_witness) = subset (op =) (#I wit, (0 upto m - 1) @ map snd X);
+
+    val all = m upto m + n - 1;
+
+    fun enrich X = map_filter (fn i =>
+      (case find_first (fn (_, i') => i = i') X of
+        NONE =>
+          (case find_index (new_wit X) (nth witss (i - m)) of
+            ~1 => NONE
+          | j => SOME (j, i))
+      | SOME ji => SOME ji)) all;
+    val reachable = fixpoint (op =) enrich [];
+    val _ = (case subtract (op =) (map snd reachable) all of
+        [] => ()
+      | i :: _ => error ("Cannot define empty datatype " ^ quote (Binding.name_of (nth bs (i - m)))));
+
+    val wit_thms = flat (map2 (fn bnf => fn (j, _) => nth (wit_thmss_of_bnf bnf) j) bnfs reachable);
+
+    val timer = time (timer "Checked nonemptiness");
+
+    (* derived thms *)
+
+    (*map g1 ... gm g(m+1) ... g(m+n) (map id ... id f(m+1) ... f(m+n) x) =
+      map g1 ... gm (g(m+1) o f(m+1)) ... (g(m+n) o f(m+n)) x*)
+    fun mk_map_comp_id x mapAsBs mapBsCs mapAsCs map_comp0 =
+      let
+        val lhs = Term.list_comb (mapBsCs, all_gs) $
+          (Term.list_comb (mapAsBs, passive_ids @ fs) $ x);
+        val rhs = Term.list_comb (mapAsCs,
+          take m all_gs @ map HOLogic.mk_comp (drop m all_gs ~~ fs)) $ x;
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (x :: fs @ all_gs) (mk_Trueprop_eq (lhs, rhs)))
+          (K (mk_map_comp_id_tac map_comp0))
+        |> Thm.close_derivation
+      end;
+
+    val map_comp_id_thms = map5 mk_map_comp_id xFs mapsAsBs mapsBsCs' mapsAsCs' map_comps;
+
+    (*forall a : set(m+1) x. f(m+1) a = a; ...; forall a : set(m+n) x. f(m+n) a = a ==>
+      map id ... id f(m+1) ... f(m+n) x = x*)
+    fun mk_map_cong0L x mapAsAs sets map_cong0 map_id =
+      let
+        fun mk_prem set f z z' = HOLogic.mk_Trueprop
+          (mk_Ball (set $ x) (Term.absfree z' (HOLogic.mk_eq (f $ z, z))));
+        val prems = map4 mk_prem (drop m sets) self_fs zs zs';
+        val goal = mk_Trueprop_eq (Term.list_comb (mapAsAs, passive_ids @ self_fs) $ x, x);
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (x :: self_fs) (Logic.list_implies (prems, goal)))
+          (K (mk_map_cong0L_tac m map_cong0 map_id))
+        |> Thm.close_derivation
+      end;
+
+    val map_cong0L_thms = map5 mk_map_cong0L xFs mapsAsAs setssAs map_cong0s map_ids;
+    val in_mono'_thms = map (fn bnf => in_mono_of_bnf bnf OF (replicate m subset_refl)) bnfs;
+    val in_cong'_thms = map (fn bnf => in_cong_of_bnf bnf OF (replicate m refl)) bnfs;
+
+    val timer = time (timer "Derived simple theorems");
+
+    (* algebra *)
+
+    val alg_bind = mk_internal_b algN;
+    val alg_name = Binding.name_of alg_bind;
+    val alg_def_bind = (Thm.def_binding alg_bind, []);
+
+    (*forall i = 1 ... n: (\<forall>x \<in> Fi_in A1 .. Am B1 ... Bn. si x \<in> Bi)*)
+    val alg_spec =
+      let
+        val algT = Library.foldr (op -->) (ATs @ BTs @ sTs, HOLogic.boolT);
+
+        val ins = map3 mk_in (replicate n (As @ Bs)) setssAs FTsAs;
+        fun mk_alg_conjunct B s X x x' =
+          mk_Ball X (Term.absfree x' (HOLogic.mk_mem (s $ x, B)));
+
+        val lhs = Term.list_comb (Free (alg_name, algT), As @ Bs @ ss);
+        val rhs = Library.foldr1 HOLogic.mk_conj (map5 mk_alg_conjunct Bs ss ins xFs xFs')
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((alg_free, (_, alg_def_free)), (lthy, lthy_old)) =
+        lthy
+        |> Specification.definition (SOME (alg_bind, NONE, NoSyn), (alg_def_bind, alg_spec))
+        ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+    val alg = fst (Term.dest_Const (Morphism.term phi alg_free));
+    val alg_def = Morphism.thm phi alg_def_free;
+
+    fun mk_alg As Bs ss =
+      let
+        val args = As @ Bs @ ss;
+        val Ts = map fastype_of args;
+        val algT = Library.foldr (op -->) (Ts, HOLogic.boolT);
+      in
+        Term.list_comb (Const (alg, algT), args)
+      end;
+
+    val alg_set_thms =
+      let
+        val alg_prem = HOLogic.mk_Trueprop (mk_alg As Bs ss);
+        fun mk_prem x set B = HOLogic.mk_Trueprop (mk_leq (set $ x) B);
+        fun mk_concl s x B = HOLogic.mk_Trueprop (HOLogic.mk_mem (s $ x, B));
+        val premss = map2 ((fn x => fn sets =>  map2 (mk_prem x) sets (As @ Bs))) xFs setssAs;
+        val concls = map3 mk_concl ss xFs Bs;
+        val goals = map3 (fn x => fn prems => fn concl =>
+          fold_rev Logic.all (x :: As @ Bs @ ss)
+            (Logic.list_implies (alg_prem :: prems, concl))) xFs premss concls;
+      in
+        map (fn goal =>
+          Goal.prove_sorry lthy [] [] goal (K (mk_alg_set_tac alg_def)) |> Thm.close_derivation)
+        goals
+      end;
+
+    fun mk_talg ATs BTs = mk_alg (map HOLogic.mk_UNIV ATs) (map HOLogic.mk_UNIV BTs);
+
+    val talg_thm =
+      let
+        val goal = fold_rev Logic.all ss
+          (HOLogic.mk_Trueprop (mk_talg passiveAs activeAs ss))
+      in
+        Goal.prove_sorry lthy [] [] goal
+          (K (stac alg_def 1 THEN CONJ_WRAP (K (EVERY' [rtac ballI, rtac UNIV_I] 1)) ss))
+        |> Thm.close_derivation
+      end;
+
+    val timer = time (timer "Algebra definition & thms");
+
+    val alg_not_empty_thms =
+      let
+        val alg_prem =
+          HOLogic.mk_Trueprop (mk_alg passive_UNIVs Bs ss);
+        val concls = map (HOLogic.mk_Trueprop o mk_not_empty) Bs;
+        val goals =
+          map (fn concl =>
+            fold_rev Logic.all (Bs @ ss) (Logic.mk_implies (alg_prem, concl))) concls;
+      in
+        map2 (fn goal => fn alg_set =>
+          Goal.prove_sorry lthy [] []
+            goal (K (mk_alg_not_empty_tac lthy alg_set alg_set_thms wit_thms))
+          |> Thm.close_derivation)
+        goals alg_set_thms
+      end;
+
+    val timer = time (timer "Proved nonemptiness");
+
+    (* morphism *)
+
+    val mor_bind = mk_internal_b morN;
+    val mor_name = Binding.name_of mor_bind;
+    val mor_def_bind = (Thm.def_binding mor_bind, []);
+
+    (*fbetw) forall i = 1 ... n: (\<forall>x \<in> Bi. f x \<in> B'i)*)
+    (*mor) forall i = 1 ... n: (\<forall>x \<in> Fi_in UNIV ... UNIV B1 ... Bn.
+       f (s1 x) = s1' (Fi_map id ... id f1 ... fn x))*)
+    val mor_spec =
+      let
+        val morT = Library.foldr (op -->) (BTs @ sTs @ B'Ts @ s'Ts @ fTs, HOLogic.boolT);
+
+        fun mk_fbetw f B1 B2 z z' =
+          mk_Ball B1 (Term.absfree z' (HOLogic.mk_mem (f $ z, B2)));
+        fun mk_mor sets mapAsBs f s s' T x x' =
+          mk_Ball (mk_in (passive_UNIVs @ Bs) sets T)
+            (Term.absfree x' (HOLogic.mk_eq (f $ (s $ x), s' $
+              (Term.list_comb (mapAsBs, passive_ids @ fs) $ x))));
+        val lhs = Term.list_comb (Free (mor_name, morT), Bs @ ss @ B's @ s's @ fs);
+        val rhs = HOLogic.mk_conj
+          (Library.foldr1 HOLogic.mk_conj (map5 mk_fbetw fs Bs B's zs zs'),
+          Library.foldr1 HOLogic.mk_conj
+            (map8 mk_mor setssAs mapsAsBs fs ss s's FTsAs xFs xFs'))
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((mor_free, (_, mor_def_free)), (lthy, lthy_old)) =
+        lthy
+        |> Specification.definition (SOME (mor_bind, NONE, NoSyn), (mor_def_bind, mor_spec))
+        ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+    val mor = fst (Term.dest_Const (Morphism.term phi mor_free));
+    val mor_def = Morphism.thm phi mor_def_free;
+
+    fun mk_mor Bs1 ss1 Bs2 ss2 fs =
+      let
+        val args = Bs1 @ ss1 @ Bs2 @ ss2 @ fs;
+        val Ts = map fastype_of (Bs1 @ ss1 @ Bs2 @ ss2 @ fs);
+        val morT = Library.foldr (op -->) (Ts, HOLogic.boolT);
+      in
+        Term.list_comb (Const (mor, morT), args)
+      end;
+
+    val (mor_image_thms, morE_thms) =
+      let
+        val prem = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs);
+        fun mk_image_goal f B1 B2 = fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs)
+          (Logic.mk_implies (prem, HOLogic.mk_Trueprop (mk_leq (mk_image f $ B1) B2)));
+        val image_goals = map3 mk_image_goal fs Bs B's;
+        fun mk_elim_prem sets x T = HOLogic.mk_Trueprop
+          (HOLogic.mk_mem (x, mk_in (passive_UNIVs @ Bs) sets T));
+        fun mk_elim_goal sets mapAsBs f s s' x T =
+          fold_rev Logic.all (x :: Bs @ ss @ B's @ s's @ fs)
+            (Logic.list_implies ([prem, mk_elim_prem sets x T],
+              mk_Trueprop_eq (f $ (s $ x), s' $ Term.list_comb (mapAsBs, passive_ids @ fs @ [x]))));
+        val elim_goals = map7 mk_elim_goal setssAs mapsAsBs fs ss s's xFs FTsAs;
+        fun prove goal =
+          Goal.prove_sorry lthy [] [] goal (K (mk_mor_elim_tac mor_def)) |> Thm.close_derivation;
+      in
+        (map prove image_goals, map prove elim_goals)
+      end;
+
+    val mor_incl_thm =
+      let
+        val prems = map2 (HOLogic.mk_Trueprop oo mk_leq) Bs Bs_copy;
+        val concl = HOLogic.mk_Trueprop (mk_mor Bs ss Bs_copy ss active_ids);
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (Bs @ ss @ Bs_copy) (Logic.list_implies (prems, concl)))
+          (K (mk_mor_incl_tac mor_def map_ids))
+        |> Thm.close_derivation
+      end;
+
+    val mor_comp_thm =
+      let
+        val prems =
+          [HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs),
+           HOLogic.mk_Trueprop (mk_mor B's s's B''s s''s gs)];
+        val concl =
+          HOLogic.mk_Trueprop (mk_mor Bs ss B''s s''s (map2 (curry HOLogic.mk_comp) gs fs));
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (Bs @ ss @ B's @ s's @ B''s @ s''s @ fs @ gs)
+             (Logic.list_implies (prems, concl)))
+          (K (mk_mor_comp_tac mor_def set_mapss map_comp_id_thms))
+        |> Thm.close_derivation
+      end;
+
+    val mor_inv_thm =
+      let
+        fun mk_inv_prem f inv_f B B' = HOLogic.mk_conj (mk_leq (mk_image inv_f $ B') B,
+          HOLogic.mk_conj (mk_inver inv_f f B, mk_inver f inv_f B'));
+        val prems = map HOLogic.mk_Trueprop
+          ([mk_mor Bs ss B's s's fs,
+          mk_alg passive_UNIVs Bs ss,
+          mk_alg passive_UNIVs B's s's] @
+          map4 mk_inv_prem fs inv_fs Bs B's);
+        val concl = HOLogic.mk_Trueprop (mk_mor B's s's Bs ss inv_fs);
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs @ inv_fs)
+            (Logic.list_implies (prems, concl)))
+          (K (mk_mor_inv_tac alg_def mor_def set_mapss morE_thms map_comp_id_thms map_cong0L_thms))
+        |> Thm.close_derivation
+      end;
+
+    val mor_cong_thm =
+      let
+        val prems = map HOLogic.mk_Trueprop
+         (map2 (curry HOLogic.mk_eq) fs_copy fs @ [mk_mor Bs ss B's s's fs])
+        val concl = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs_copy);
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs @ fs_copy)
+             (Logic.list_implies (prems, concl)))
+          (K ((hyp_subst_tac lthy THEN' atac) 1))
+        |> Thm.close_derivation
+      end;
+
+    val mor_str_thm =
+      let
+        val maps = map2 (fn Ds => fn bnf => Term.list_comb
+          (mk_map_of_bnf Ds (passiveAs @ FTsAs) allAs bnf, passive_ids @ ss)) Dss bnfs;
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all ss (HOLogic.mk_Trueprop
+            (mk_mor (map HOLogic.mk_UNIV FTsAs) maps active_UNIVs ss ss)))
+          (K (mk_mor_str_tac ks mor_def))
+        |> Thm.close_derivation
+      end;
+
+    val mor_convol_thm =
+      let
+        val maps = map3 (fn s => fn prod_s => fn mapx =>
+          mk_convol (HOLogic.mk_comp (s, Term.list_comb (mapx, passive_ids @ fsts)), prod_s))
+          s's prod_ss map_fsts;
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (s's @ prod_ss) (HOLogic.mk_Trueprop
+            (mk_mor prod_UNIVs maps (map HOLogic.mk_UNIV activeBs) s's fsts)))
+          (K (mk_mor_convol_tac ks mor_def))
+        |> Thm.close_derivation
+      end;
+
+    val mor_UNIV_thm =
+      let
+        fun mk_conjunct mapAsBs f s s' = HOLogic.mk_eq
+            (HOLogic.mk_comp (f, s),
+            HOLogic.mk_comp (s', Term.list_comb (mapAsBs, passive_ids @ fs)));
+        val lhs = mk_mor active_UNIVs ss (map HOLogic.mk_UNIV activeBs) s's fs;
+        val rhs = Library.foldr1 HOLogic.mk_conj (map4 mk_conjunct mapsAsBs fs ss s's);
+      in
+        Goal.prove_sorry lthy [] [] (fold_rev Logic.all (ss @ s's @ fs) (mk_Trueprop_eq (lhs, rhs)))
+          (K (mk_mor_UNIV_tac m morE_thms mor_def))
+        |> Thm.close_derivation
+      end;
+
+    val timer = time (timer "Morphism definition & thms");
+
+    (* isomorphism *)
+
+    (*mor Bs1 ss1 Bs2 ss2 fs \<and> (\<exists>gs. mor Bs2 ss2 Bs1 ss1 fs \<and>
+       forall i = 1 ... n. (inver gs[i] fs[i] Bs1[i] \<and> inver fs[i] gs[i] Bs2[i]))*)
+    fun mk_iso Bs1 ss1 Bs2 ss2 fs gs =
+      let
+        val ex_inv_mor = list_exists_free gs
+          (HOLogic.mk_conj (mk_mor Bs2 ss2 Bs1 ss1 gs,
+            Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_conj)
+              (map3 mk_inver gs fs Bs1) (map3 mk_inver fs gs Bs2))));
+      in
+        HOLogic.mk_conj (mk_mor Bs1 ss1 Bs2 ss2 fs, ex_inv_mor)
+      end;
+
+    val iso_alt_thm =
+      let
+        val prems = map HOLogic.mk_Trueprop
+         [mk_alg passive_UNIVs Bs ss,
+         mk_alg passive_UNIVs B's s's]
+        val concl = mk_Trueprop_eq (mk_iso Bs ss B's s's fs inv_fs,
+          HOLogic.mk_conj (mk_mor Bs ss B's s's fs,
+            Library.foldr1 HOLogic.mk_conj (map3 mk_bij_betw fs Bs B's)));
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs) (Logic.list_implies (prems, concl)))
+          (K (mk_iso_alt_tac mor_image_thms mor_inv_thm))
+        |> Thm.close_derivation
+      end;
+
+    val timer = time (timer "Isomorphism definition & thms");
+
+    (* algebra copies *)
+
+    val (copy_alg_thm, ex_copy_alg_thm) =
+      let
+        val prems = map HOLogic.mk_Trueprop
+         (mk_alg passive_UNIVs Bs ss :: map3 mk_bij_betw inv_fs B's Bs);
+        val inver_prems = map HOLogic.mk_Trueprop
+          (map3 mk_inver inv_fs fs Bs @ map3 mk_inver fs inv_fs B's);
+        val all_prems = prems @ inver_prems;
+        fun mk_s f s mapT y y' = Term.absfree y' (f $ (s $
+          (Term.list_comb (mapT, passive_ids @ inv_fs) $ y)));
+
+        val alg = HOLogic.mk_Trueprop
+          (mk_alg passive_UNIVs B's (map5 mk_s fs ss mapsBsAs yFs yFs'));
+        val copy_str_thm = Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (Bs @ ss @ B's @ inv_fs @ fs)
+            (Logic.list_implies (all_prems, alg)))
+          (K (mk_copy_str_tac set_mapss alg_def alg_set_thms))
+          |> Thm.close_derivation;
+
+        val iso = HOLogic.mk_Trueprop
+          (mk_iso B's (map5 mk_s fs ss mapsBsAs yFs yFs') Bs ss inv_fs fs_copy);
+        val copy_alg_thm = Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (Bs @ ss @ B's @ inv_fs @ fs)
+            (Logic.list_implies (all_prems, iso)))
+          (K (mk_copy_alg_tac set_mapss alg_set_thms mor_def iso_alt_thm copy_str_thm))
+          |> Thm.close_derivation;
+
+        val ex = HOLogic.mk_Trueprop
+          (list_exists_free s's
+            (HOLogic.mk_conj (mk_alg passive_UNIVs B's s's,
+              mk_iso B's s's Bs ss inv_fs fs_copy)));
+        val ex_copy_alg_thm = Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (Bs @ ss @ B's @ inv_fs @ fs)
+             (Logic.list_implies (prems, ex)))
+          (K (mk_ex_copy_alg_tac n copy_str_thm copy_alg_thm))
+          |> Thm.close_derivation;
+      in
+        (copy_alg_thm, ex_copy_alg_thm)
+      end;
+
+    val timer = time (timer "Copy thms");
+
+
+    (* bounds *)
+
+    val sum_Card_order = if n = 1 then bd_Card_order else @{thm Card_order_csum};
+    val sum_Cnotzero = if n = 1 then bd_Cnotzero else bd_Cnotzero RS @{thm csum_Cnotzero1};
+    val sum_Cinfinite = if n = 1 then bd_Cinfinite else bd_Cinfinite RS @{thm Cinfinite_csum1};
+    fun mk_set_bd_sums i bd_Card_order bds =
+      if n = 1 then bds
+      else map (fn thm => bd_Card_order RS mk_ordLeq_csum n i thm) bds;
+    val set_bd_sumss = map3 mk_set_bd_sums ks bd_Card_orders set_bdss;
+
+    fun mk_in_bd_sum i Co Cnz bd =
+      if n = 1 then bd
+      else Cnz RS ((Co RS mk_ordLeq_csum n i (Co RS @{thm ordLeq_refl})) RS
+        (bd RS @{thm ordLeq_transitive[OF _ cexp_mono2_Cnotzero[OF _ Card_order_csum]]}));
+    val in_bd_sums = map4 mk_in_bd_sum ks bd_Card_orders bd_Cnotzeros in_bds;
+
+    val sum_bd = Library.foldr1 (uncurry mk_csum) bds;
+    val suc_bd = mk_cardSuc sum_bd;
+    val field_suc_bd = mk_Field suc_bd;
+    val suc_bdT = fst (dest_relT (fastype_of suc_bd));
+    fun mk_Asuc_bd [] = mk_cexp ctwo suc_bd
+      | mk_Asuc_bd As =
+        mk_cexp (mk_csum (Library.foldr1 (uncurry mk_csum) (map mk_card_of As)) ctwo) suc_bd;
+
+    val suc_bd_Card_order = if n = 1 then bd_Card_order RS @{thm cardSuc_Card_order}
+      else @{thm cardSuc_Card_order[OF Card_order_csum]};
+    val suc_bd_Cinfinite = if n = 1 then bd_Cinfinite RS @{thm Cinfinite_cardSuc}
+      else bd_Cinfinite RS @{thm Cinfinite_cardSuc[OF Cinfinite_csum1]};
+    val suc_bd_Cnotzero = suc_bd_Cinfinite RS @{thm Cinfinite_Cnotzero};
+    val suc_bd_worel = suc_bd_Card_order RS @{thm Card_order_wo_rel}
+    val basis_Asuc = if m = 0 then @{thm ordLeq_refl[OF Card_order_ctwo]}
+        else @{thm ordLeq_csum2[OF Card_order_ctwo]};
+    val Asuc_bd_Cinfinite = suc_bd_Cinfinite RS (basis_Asuc RS @{thm Cinfinite_cexp});
+
+    val suc_bd_Asuc_bd = @{thm ordLess_ordLeq_trans[OF ordLess_ctwo_cexp cexp_mono1]} OF
+      [suc_bd_Card_order, basis_Asuc, suc_bd_Card_order];
+
+    val Asuc_bdT = fst (dest_relT (fastype_of (mk_Asuc_bd As)));
+    val II_BTs = replicate n (HOLogic.mk_setT Asuc_bdT);
+    val II_sTs = map2 (fn Ds => fn bnf =>
+      mk_T_of_bnf Ds (passiveAs @ replicate n Asuc_bdT) bnf --> Asuc_bdT) Dss bnfs;
+
+    val (((((((idxs, Asi_name), (idx, idx')), (jdx, jdx')), II_Bs), II_ss), Asuc_fs),
+      names_lthy) = names_lthy
+      |> mk_Frees "i" (replicate n suc_bdT)
+      ||>> (fn ctxt => apfst the_single (mk_fresh_names ctxt 1 "Asi"))
+      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "i") suc_bdT
+      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "j") suc_bdT
+      ||>> mk_Frees "IIB" II_BTs
+      ||>> mk_Frees "IIs" II_sTs
+      ||>> mk_Frees "f" (map (fn T => Asuc_bdT --> T) activeAs);
+
+    val suc_bd_limit_thm =
+      let
+        val prem = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+          (map (fn idx => HOLogic.mk_mem (idx, field_suc_bd)) idxs));
+        fun mk_conjunct idx = HOLogic.mk_conj (mk_not_eq idx jdx,
+          HOLogic.mk_mem (HOLogic.mk_prod (idx, jdx), suc_bd));
+        val concl = HOLogic.mk_Trueprop (mk_Bex field_suc_bd
+          (Term.absfree jdx' (Library.foldr1 HOLogic.mk_conj (map mk_conjunct idxs))));
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all idxs (Logic.list_implies ([prem], concl)))
+          (K (mk_bd_limit_tac n suc_bd_Cinfinite))
+        |> Thm.close_derivation
+      end;
+
+    val timer = time (timer "Bounds");
+
+
+    (* minimal algebra *)
+
+    fun mk_minG Asi i k = mk_UNION (mk_underS suc_bd $ i)
+      (Term.absfree jdx' (mk_nthN n (Asi $ jdx) k));
+
+    fun mk_minH_component As Asi i sets Ts s k =
+      HOLogic.mk_binop @{const_name "sup"}
+      (mk_minG Asi i k, mk_image s $ mk_in (As @ map (mk_minG Asi i) ks) sets Ts);
+
+    fun mk_min_algs As ss =
+      let
+        val BTs = map (range_type o fastype_of) ss;
+        val Ts = map (HOLogic.dest_setT o fastype_of) As @ BTs;
+        val (Asi, Asi') = `Free (Asi_name, suc_bdT -->
+          Library.foldr1 HOLogic.mk_prodT (map HOLogic.mk_setT BTs));
+      in
+         mk_worec suc_bd (Term.absfree Asi' (Term.absfree idx' (HOLogic.mk_tuple
+           (map4 (mk_minH_component As Asi idx) (mk_setss Ts) (mk_FTs Ts) ss ks))))
+      end;
+
+    val (min_algs_thms, min_algs_mono_thms, card_of_min_algs_thm, least_min_algs_thm) =
+      let
+        val i_field = HOLogic.mk_mem (idx, field_suc_bd);
+        val min_algs = mk_min_algs As ss;
+        val min_algss = map (fn k => mk_nthN n (min_algs $ idx) k) ks;
+
+        val concl = HOLogic.mk_Trueprop
+          (HOLogic.mk_eq (min_algs $ idx, HOLogic.mk_tuple
+            (map4 (mk_minH_component As min_algs idx) setssAs FTsAs ss ks)));
+        val goal = fold_rev Logic.all (idx :: As @ ss)
+          (Logic.mk_implies (HOLogic.mk_Trueprop i_field, concl));
+
+        val min_algs_thm = Goal.prove_sorry lthy [] [] goal
+          (K (mk_min_algs_tac suc_bd_worel in_cong'_thms))
+          |> Thm.close_derivation;
+
+        val min_algs_thms = map (fn k => min_algs_thm RS mk_nthI n k) ks;
+
+        fun mk_mono_goal min_alg =
+          fold_rev Logic.all (As @ ss) (HOLogic.mk_Trueprop (mk_relChain suc_bd
+            (Term.absfree idx' min_alg)));
+
+        val monos =
+          map2 (fn goal => fn min_algs =>
+            Goal.prove_sorry lthy [] [] goal (K (mk_min_algs_mono_tac lthy min_algs))
+            |> Thm.close_derivation)
+          (map mk_mono_goal min_algss) min_algs_thms;
+
+        val Asuc_bd = mk_Asuc_bd As;
+
+        fun mk_card_conjunct min_alg = mk_ordLeq (mk_card_of min_alg) Asuc_bd;
+        val card_conjunction = Library.foldr1 HOLogic.mk_conj (map mk_card_conjunct min_algss);
+        val card_cT = certifyT lthy suc_bdT;
+        val card_ct = certify lthy (Term.absfree idx' card_conjunction);
+
+        val card_of = singleton (Proof_Context.export names_lthy lthy)
+          (Goal.prove_sorry lthy [] []
+            (HOLogic.mk_Trueprop (HOLogic.mk_imp (i_field, card_conjunction)))
+            (K (mk_min_algs_card_of_tac card_cT card_ct
+              m suc_bd_worel min_algs_thms in_bd_sums
+              sum_Card_order sum_Cnotzero suc_bd_Card_order suc_bd_Cinfinite suc_bd_Cnotzero
+              suc_bd_Asuc_bd Asuc_bd_Cinfinite)))
+          |> Thm.close_derivation;
+
+        val least_prem = HOLogic.mk_Trueprop (mk_alg As Bs ss);
+        val least_conjunction = Library.foldr1 HOLogic.mk_conj (map2 mk_leq min_algss Bs);
+        val least_cT = certifyT lthy suc_bdT;
+        val least_ct = certify lthy (Term.absfree idx' least_conjunction);
+
+        val least = singleton (Proof_Context.export names_lthy lthy)
+          (Goal.prove_sorry lthy [] []
+            (Logic.mk_implies (least_prem,
+              HOLogic.mk_Trueprop (HOLogic.mk_imp (i_field, least_conjunction))))
+            (K (mk_min_algs_least_tac least_cT least_ct
+              suc_bd_worel min_algs_thms alg_set_thms)))
+          |> Thm.close_derivation;
+      in
+        (min_algs_thms, monos, card_of, least)
+      end;
+
+    val timer = time (timer "min_algs definition & thms");
+
+    val min_alg_binds = mk_internal_bs min_algN;
+    fun min_alg_bind i = nth min_alg_binds (i - 1);
+    fun min_alg_name i = Binding.name_of (min_alg_bind i);
+    val min_alg_def_bind = rpair [] o Thm.def_binding o min_alg_bind;
+
+    fun min_alg_spec i =
+      let
+        val min_algT =
+          Library.foldr (op -->) (ATs @ sTs, HOLogic.mk_setT (nth activeAs (i - 1)));
+
+        val lhs = Term.list_comb (Free (min_alg_name i, min_algT), As @ ss);
+        val rhs = mk_UNION (field_suc_bd)
+          (Term.absfree idx' (mk_nthN n (mk_min_algs As ss $ idx) i));
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((min_alg_frees, (_, min_alg_def_frees)), (lthy, lthy_old)) =
+        lthy
+        |> fold_map (fn i => Specification.definition
+          (SOME (min_alg_bind i, NONE, NoSyn), (min_alg_def_bind i, min_alg_spec i))) ks
+        |>> apsnd split_list o split_list
+        ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+    val min_algs = map (fst o Term.dest_Const o Morphism.term phi) min_alg_frees;
+    val min_alg_defs = map (Morphism.thm phi) min_alg_def_frees;
+
+    fun mk_min_alg As ss i =
+      let
+        val T = HOLogic.mk_setT (range_type (fastype_of (nth ss (i - 1))))
+        val args = As @ ss;
+        val Ts = map fastype_of args;
+        val min_algT = Library.foldr (op -->) (Ts, T);
+      in
+        Term.list_comb (Const (nth min_algs (i - 1), min_algT), args)
+      end;
+
+    val (alg_min_alg_thm, card_of_min_alg_thms, least_min_alg_thms, mor_incl_min_alg_thm) =
+      let
+        val min_algs = map (mk_min_alg As ss) ks;
+
+        val goal = fold_rev Logic.all (As @ ss) (HOLogic.mk_Trueprop (mk_alg As min_algs ss));
+        val alg_min_alg = Goal.prove_sorry lthy [] [] goal
+          (K (mk_alg_min_alg_tac m alg_def min_alg_defs suc_bd_limit_thm sum_Cinfinite
+            set_bd_sumss min_algs_thms min_algs_mono_thms))
+          |> Thm.close_derivation;
+
+        val Asuc_bd = mk_Asuc_bd As;
+        fun mk_card_of_thm min_alg def = Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (As @ ss)
+            (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of min_alg) Asuc_bd)))
+          (K (mk_card_of_min_alg_tac def card_of_min_algs_thm
+            suc_bd_Card_order suc_bd_Asuc_bd Asuc_bd_Cinfinite))
+          |> Thm.close_derivation;
+
+        val least_prem = HOLogic.mk_Trueprop (mk_alg As Bs ss);
+        fun mk_least_thm min_alg B def = Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (As @ Bs @ ss)
+            (Logic.mk_implies (least_prem, HOLogic.mk_Trueprop (mk_leq min_alg B))))
+          (K (mk_least_min_alg_tac def least_min_algs_thm))
+          |> Thm.close_derivation;
+
+        val leasts = map3 mk_least_thm min_algs Bs min_alg_defs;
+
+        val incl_prem = HOLogic.mk_Trueprop (mk_alg passive_UNIVs Bs ss);
+        val incl_min_algs = map (mk_min_alg passive_UNIVs ss) ks;
+        val incl = Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (Bs @ ss)
+            (Logic.mk_implies (incl_prem,
+              HOLogic.mk_Trueprop (mk_mor incl_min_algs ss Bs ss active_ids))))
+          (K (EVERY' (rtac mor_incl_thm :: map etac leasts) 1))
+          |> Thm.close_derivation;
+      in
+        (alg_min_alg, map2 mk_card_of_thm min_algs min_alg_defs, leasts, incl)
+      end;
+
+    val timer = time (timer "Minimal algebra definition & thms");
+
+    val II_repT = HOLogic.mk_prodT (HOLogic.mk_tupleT II_BTs, HOLogic.mk_tupleT II_sTs);
+    val IIT_bind = mk_internal_b IITN;
+
+    val ((IIT_name, (IIT_glob_info, IIT_loc_info)), lthy) =
+      typedef (IIT_bind, params, NoSyn)
+        (HOLogic.mk_UNIV II_repT) NONE (EVERY' [rtac exI, rtac UNIV_I] 1) lthy;
+
+    val IIT = Type (IIT_name, params');
+    val Abs_IIT = Const (#Abs_name IIT_glob_info, II_repT --> IIT);
+    val Rep_IIT = Const (#Rep_name IIT_glob_info, IIT --> II_repT);
+    val Abs_IIT_inverse_thm = UNIV_I RS #Abs_inverse IIT_loc_info;
+
+    val initT = IIT --> Asuc_bdT;
+    val active_initTs = replicate n initT;
+    val init_FTs = map2 (fn Ds => mk_T_of_bnf Ds (passiveAs @ active_initTs)) Dss bnfs;
+    val init_fTs = map (fn T => initT --> T) activeAs;
+
+    val (((((((iidx, iidx'), init_xs), (init_xFs, init_xFs')),
+      init_fs), init_fs_copy), init_phis), names_lthy) = names_lthy
+      |> yield_singleton (apfst (op ~~) oo mk_Frees' "i") IIT
+      ||>> mk_Frees "ix" active_initTs
+      ||>> mk_Frees' "x" init_FTs
+      ||>> mk_Frees "f" init_fTs
+      ||>> mk_Frees "f" init_fTs
+      ||>> mk_Frees "P" (replicate n (mk_pred1T initT));
+
+    val II = HOLogic.mk_Collect (fst iidx', IIT, list_exists_free (II_Bs @ II_ss)
+      (HOLogic.mk_conj (HOLogic.mk_eq (iidx,
+        Abs_IIT $ (HOLogic.mk_prod (HOLogic.mk_tuple II_Bs, HOLogic.mk_tuple II_ss))),
+        mk_alg passive_UNIVs II_Bs II_ss)));
+
+    val select_Bs = map (mk_nthN n (HOLogic.mk_fst (Rep_IIT $ iidx))) ks;
+    val select_ss = map (mk_nthN n (HOLogic.mk_snd (Rep_IIT $ iidx))) ks;
+
+    val str_init_binds = mk_internal_bs str_initN;
+    fun str_init_bind i = nth str_init_binds (i - 1);
+    val str_init_name = Binding.name_of o str_init_bind;
+    val str_init_def_bind = rpair [] o Thm.def_binding o str_init_bind;
+
+    fun str_init_spec i =
+      let
+        val T = nth init_FTs (i - 1);
+        val init_xF = nth init_xFs (i - 1)
+        val select_s = nth select_ss (i - 1);
+        val map = mk_map_of_bnf (nth Dss (i - 1))
+          (passiveAs @ active_initTs) (passiveAs @ replicate n Asuc_bdT)
+          (nth bnfs (i - 1));
+        val map_args = passive_ids @ replicate n (mk_rapp iidx Asuc_bdT);
+        val str_initT = T --> IIT --> Asuc_bdT;
+
+        val lhs = Term.list_comb (Free (str_init_name i, str_initT), [init_xF, iidx]);
+        val rhs = select_s $ (Term.list_comb (map, map_args) $ init_xF);
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((str_init_frees, (_, str_init_def_frees)), (lthy, lthy_old)) =
+      lthy
+      |> fold_map (fn i => Specification.definition
+        (SOME (str_init_bind i, NONE, NoSyn), (str_init_def_bind i, str_init_spec i))) ks
+      |>> apsnd split_list o split_list
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+    val str_inits =
+      map (Term.subst_atomic_types (map (`(Morphism.typ phi)) params') o Morphism.term phi)
+        str_init_frees;
+
+    val str_init_defs = map (Morphism.thm phi) str_init_def_frees;
+
+    val car_inits = map (mk_min_alg passive_UNIVs str_inits) ks;
+
+    (*TODO: replace with instantiate? (problem: figure out right type instantiation)*)
+    val alg_init_thm = Goal.prove_sorry lthy [] []
+      (HOLogic.mk_Trueprop (mk_alg passive_UNIVs car_inits str_inits))
+      (K (rtac alg_min_alg_thm 1))
+      |> Thm.close_derivation;
+
+    val alg_select_thm = Goal.prove_sorry lthy [] []
+      (HOLogic.mk_Trueprop (mk_Ball II
+        (Term.absfree iidx' (mk_alg passive_UNIVs select_Bs select_ss))))
+      (mk_alg_select_tac Abs_IIT_inverse_thm)
+      |> Thm.close_derivation;
+
+    val mor_select_thm =
+      let
+        val alg_prem = HOLogic.mk_Trueprop (mk_alg passive_UNIVs Bs ss);
+        val i_prem = HOLogic.mk_Trueprop (HOLogic.mk_mem (iidx, II));
+        val mor_prem = HOLogic.mk_Trueprop (mk_mor select_Bs select_ss Bs ss Asuc_fs);
+        val prems = [alg_prem, i_prem, mor_prem];
+        val concl = HOLogic.mk_Trueprop
+          (mk_mor car_inits str_inits Bs ss
+            (map (fn f => HOLogic.mk_comp (f, mk_rapp iidx Asuc_bdT)) Asuc_fs));
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (iidx :: Bs @ ss @ Asuc_fs) (Logic.list_implies (prems, concl)))
+          (K (mk_mor_select_tac mor_def mor_cong_thm mor_comp_thm mor_incl_min_alg_thm alg_def
+            alg_select_thm alg_set_thms set_mapss str_init_defs))
+        |> Thm.close_derivation
+      end;
+
+    val (init_ex_mor_thm, init_unique_mor_thms) =
+      let
+        val prem = HOLogic.mk_Trueprop (mk_alg passive_UNIVs Bs ss);
+        val concl = HOLogic.mk_Trueprop
+          (list_exists_free init_fs (mk_mor car_inits str_inits Bs ss init_fs));
+        val ex_mor = Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (Bs @ ss) (Logic.mk_implies (prem, concl)))
+          (mk_init_ex_mor_tac Abs_IIT_inverse_thm ex_copy_alg_thm alg_min_alg_thm
+            card_of_min_alg_thms mor_comp_thm mor_select_thm mor_incl_min_alg_thm)
+          |> Thm.close_derivation;
+
+        val prems = map2 (HOLogic.mk_Trueprop oo curry HOLogic.mk_mem) init_xs car_inits
+        val mor_prems = map HOLogic.mk_Trueprop
+          [mk_mor car_inits str_inits Bs ss init_fs,
+          mk_mor car_inits str_inits Bs ss init_fs_copy];
+        fun mk_fun_eq f g x = HOLogic.mk_eq (f $ x, g $ x);
+        val unique = HOLogic.mk_Trueprop
+          (Library.foldr1 HOLogic.mk_conj (map3 mk_fun_eq init_fs init_fs_copy init_xs));
+        val unique_mor = Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (init_xs @ Bs @ ss @ init_fs @ init_fs_copy)
+            (Logic.list_implies (prems @ mor_prems, unique)))
+          (K (mk_init_unique_mor_tac m alg_def alg_init_thm least_min_alg_thms
+            in_mono'_thms alg_set_thms morE_thms map_cong0s))
+          |> Thm.close_derivation;
+      in
+        (ex_mor, split_conj_thm unique_mor)
+      end;
+
+    val init_setss = mk_setss (passiveAs @ active_initTs);
+    val active_init_setss = map (drop m) init_setss;
+    val init_ins = map2 (fn sets => mk_in (passive_UNIVs @ car_inits) sets) init_setss init_FTs;
+
+    fun mk_closed phis =
+      let
+        fun mk_conjunct phi str_init init_sets init_in x x' =
+          let
+            val prem = Library.foldr1 HOLogic.mk_conj
+              (map2 (fn set => mk_Ball (set $ x)) init_sets phis);
+            val concl = phi $ (str_init $ x);
+          in
+            mk_Ball init_in (Term.absfree x' (HOLogic.mk_imp (prem, concl)))
+          end;
+      in
+        Library.foldr1 HOLogic.mk_conj
+          (map6 mk_conjunct phis str_inits active_init_setss init_ins init_xFs init_xFs')
+      end;
+
+    val init_induct_thm =
+      let
+        val prem = HOLogic.mk_Trueprop (mk_closed init_phis);
+        val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+          (map2 mk_Ball car_inits init_phis));
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all init_phis (Logic.mk_implies (prem, concl)))
+          (K (mk_init_induct_tac m alg_def alg_init_thm least_min_alg_thms alg_set_thms))
+        |> Thm.close_derivation
+      end;
+
+    val timer = time (timer "Initiality definition & thms");
+
+    val ((T_names, (T_glob_infos, T_loc_infos)), lthy) =
+      lthy
+      |> fold_map3 (fn b => fn mx => fn car_init =>
+        typedef (Binding.conceal b, params, mx) car_init NONE
+          (EVERY' [rtac ssubst, rtac @{thm ex_in_conv}, resolve_tac alg_not_empty_thms,
+            rtac alg_init_thm] 1)) bs mixfixes car_inits
+      |>> apsnd split_list o split_list;
+
+    val Ts = map (fn name => Type (name, params')) T_names;
+    fun mk_Ts passive = map (Term.typ_subst_atomic (passiveAs ~~ passive)) Ts;
+    val Ts' = mk_Ts passiveBs;
+    val Rep_Ts = map2 (fn info => fn T => Const (#Rep_name info, T --> initT)) T_glob_infos Ts;
+    val Abs_Ts = map2 (fn info => fn T => Const (#Abs_name info, initT --> T)) T_glob_infos Ts;
+
+    val type_defs = map #type_definition T_loc_infos;
+    val Reps = map #Rep T_loc_infos;
+    val Rep_casess = map #Rep_cases T_loc_infos;
+    val Rep_injects = map #Rep_inject T_loc_infos;
+    val Rep_inverses = map #Rep_inverse T_loc_infos;
+    val Abs_inverses = map #Abs_inverse T_loc_infos;
+
+    fun mk_inver_thm mk_tac rep abs X thm =
+      Goal.prove_sorry lthy [] []
+        (HOLogic.mk_Trueprop (mk_inver rep abs X))
+        (K (EVERY' [rtac ssubst, rtac @{thm inver_def}, rtac ballI, mk_tac thm] 1))
+      |> Thm.close_derivation;
+
+    val inver_Reps = map4 (mk_inver_thm rtac) Abs_Ts Rep_Ts (map HOLogic.mk_UNIV Ts) Rep_inverses;
+    val inver_Abss = map4 (mk_inver_thm etac) Rep_Ts Abs_Ts car_inits Abs_inverses;
+
+    val timer = time (timer "THE TYPEDEFs & Rep/Abs thms");
+
+    val UNIVs = map HOLogic.mk_UNIV Ts;
+    val FTs = mk_FTs (passiveAs @ Ts);
+    val FTs' = mk_FTs (passiveBs @ Ts');
+    fun mk_set_Ts T = passiveAs @ replicate n (HOLogic.mk_setT T);
+    val setFTss = map (mk_FTs o mk_set_Ts) passiveAs;
+    val FTs_setss = mk_setss (passiveAs @ Ts);
+    val FTs'_setss = mk_setss (passiveBs @ Ts');
+    val map_FT_inits = map2 (fn Ds =>
+      mk_map_of_bnf Ds (passiveAs @ Ts) (passiveAs @ active_initTs)) Dss bnfs;
+    val fTs = map2 (curry op -->) Ts activeAs;
+    val foldT = Library.foldr1 HOLogic.mk_prodT (map2 (curry op -->) Ts activeAs);
+    val rec_sTs = map (Term.typ_subst_atomic (activeBs ~~ Ts)) prod_sTs;
+    val rec_maps = map (Term.subst_atomic_types (activeBs ~~ Ts)) map_fsts;
+    val rec_maps_rev = map (Term.subst_atomic_types (activeBs ~~ Ts)) map_fsts_rev;
+    val rec_fsts = map (Term.subst_atomic_types (activeBs ~~ Ts)) fsts;
+    val rec_UNIVs = map2 (HOLogic.mk_UNIV oo curry HOLogic.mk_prodT) Ts activeAs;
+
+    val (((((((((Izs1, Izs1'), (Izs2, Izs2')), (xFs, xFs')), yFs), (AFss, AFss')),
+      (fold_f, fold_f')), fs), rec_ss), names_lthy) = names_lthy
+      |> mk_Frees' "z1" Ts
+      ||>> mk_Frees' "z2" Ts'
+      ||>> mk_Frees' "x" FTs
+      ||>> mk_Frees "y" FTs'
+      ||>> mk_Freess' "z" setFTss
+      ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "f") foldT
+      ||>> mk_Frees "f" fTs
+      ||>> mk_Frees "s" rec_sTs;
+
+    val Izs = map2 retype_free Ts zs;
+    val phis = map2 retype_free (map mk_pred1T Ts) init_phis;
+    val phi2s = map2 retype_free (map2 mk_pred2T Ts Ts') init_phis;
+
+    fun ctor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctorN ^ "_");
+    val ctor_name = Binding.name_of o ctor_bind;
+    val ctor_def_bind = rpair [] o Binding.conceal o Thm.def_binding o ctor_bind;
+
+    fun ctor_spec i abs str map_FT_init x x' =
+      let
+        val ctorT = nth FTs (i - 1) --> nth Ts (i - 1);
+
+        val lhs = Free (ctor_name i, ctorT);
+        val rhs = Term.absfree x' (abs $ (str $
+          (Term.list_comb (map_FT_init, map HOLogic.id_const passiveAs @ Rep_Ts) $ x)));
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((ctor_frees, (_, ctor_def_frees)), (lthy, lthy_old)) =
+      lthy
+      |> fold_map6 (fn i => fn abs => fn str => fn mapx => fn x => fn x' =>
+        Specification.definition
+          (SOME (ctor_bind i, NONE, NoSyn), (ctor_def_bind i, ctor_spec i abs str mapx x x')))
+          ks Abs_Ts str_inits map_FT_inits xFs xFs'
+      |>> apsnd split_list o split_list
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+    fun mk_ctors passive =
+      map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ (mk_params passive)) o
+        Morphism.term phi) ctor_frees;
+    val ctors = mk_ctors passiveAs;
+    val ctor's = mk_ctors passiveBs;
+    val ctor_defs = map (Morphism.thm phi) ctor_def_frees;
+
+    val (mor_Rep_thm, mor_Abs_thm) =
+      let
+        val copy = alg_init_thm RS copy_alg_thm;
+        fun mk_bij inj Rep cases = @{thm bij_betwI'} OF [inj, Rep, cases];
+        val bijs = map3 mk_bij Rep_injects Reps Rep_casess;
+        val mor_Rep =
+          Goal.prove_sorry lthy [] []
+            (HOLogic.mk_Trueprop (mk_mor UNIVs ctors car_inits str_inits Rep_Ts))
+            (mk_mor_Rep_tac ctor_defs copy bijs inver_Abss inver_Reps)
+          |> Thm.close_derivation;
+
+        val inv = mor_inv_thm OF [mor_Rep, talg_thm, alg_init_thm];
+        val mor_Abs =
+          Goal.prove_sorry lthy [] []
+            (HOLogic.mk_Trueprop (mk_mor car_inits str_inits UNIVs ctors Abs_Ts))
+            (K (mk_mor_Abs_tac inv inver_Abss inver_Reps))
+          |> Thm.close_derivation;
+      in
+        (mor_Rep, mor_Abs)
+      end;
+
+    val timer = time (timer "ctor definitions & thms");
+
+    val fold_fun = Term.absfree fold_f'
+      (mk_mor UNIVs ctors active_UNIVs ss (map (mk_nthN n fold_f) ks));
+    val foldx = HOLogic.choice_const foldT $ fold_fun;
+
+    fun fold_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctor_foldN ^ "_");
+    val fold_name = Binding.name_of o fold_bind;
+    val fold_def_bind = rpair [] o Binding.conceal o Thm.def_binding o fold_bind;
+
+    fun fold_spec i T AT =
+      let
+        val foldT = Library.foldr (op -->) (sTs, T --> AT);
+
+        val lhs = Term.list_comb (Free (fold_name i, foldT), ss);
+        val rhs = mk_nthN n foldx i;
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((fold_frees, (_, fold_def_frees)), (lthy, lthy_old)) =
+      lthy
+      |> fold_map3 (fn i => fn T => fn AT =>
+        Specification.definition
+          (SOME (fold_bind i, NONE, NoSyn), (fold_def_bind i, fold_spec i T AT)))
+          ks Ts activeAs
+      |>> apsnd split_list o split_list
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+    val folds = map (Morphism.term phi) fold_frees;
+    val fold_names = map (fst o dest_Const) folds;
+    fun mk_folds passives actives =
+      map3 (fn name => fn T => fn active =>
+        Const (name, Library.foldr (op -->)
+          (map2 (curry op -->) (mk_FTs (passives @ actives)) actives, T --> active)))
+      fold_names (mk_Ts passives) actives;
+    fun mk_fold Ts ss i = Term.list_comb (Const (nth fold_names (i - 1), Library.foldr (op -->)
+      (map fastype_of ss, nth Ts (i - 1) --> range_type (fastype_of (nth ss (i - 1))))), ss);
+    val fold_defs = map (Morphism.thm phi) fold_def_frees;
+
+    val mor_fold_thm =
+      let
+        val ex_mor = talg_thm RS init_ex_mor_thm;
+        val mor_cong = mor_cong_thm OF (map (mk_nth_conv n) ks);
+        val mor_comp = mor_Rep_thm RS mor_comp_thm;
+        val cT = certifyT lthy foldT;
+        val ct = certify lthy fold_fun
+      in
+        singleton (Proof_Context.export names_lthy lthy)
+          (Goal.prove_sorry lthy [] []
+            (HOLogic.mk_Trueprop (mk_mor UNIVs ctors active_UNIVs ss (map (mk_fold Ts ss) ks)))
+            (K (mk_mor_fold_tac cT ct fold_defs ex_mor (mor_comp RS mor_cong))))
+        |> Thm.close_derivation
+      end;
+
+    val ctor_fold_thms = map (fn morE => rule_by_tactic lthy
+      ((rtac CollectI THEN' CONJ_WRAP' (K (rtac @{thm subset_UNIV})) (1 upto m + n)) 1)
+      (mor_fold_thm RS morE)) morE_thms;
+
+    val (fold_unique_mor_thms, fold_unique_mor_thm) =
+      let
+        val prem = HOLogic.mk_Trueprop (mk_mor UNIVs ctors active_UNIVs ss fs);
+        fun mk_fun_eq f i = HOLogic.mk_eq (f, mk_fold Ts ss i);
+        val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_fun_eq fs ks));
+        val unique_mor = Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (ss @ fs) (Logic.mk_implies (prem, unique)))
+          (K (mk_fold_unique_mor_tac type_defs init_unique_mor_thms Reps
+            mor_comp_thm mor_Abs_thm mor_fold_thm))
+          |> Thm.close_derivation;
+      in
+        `split_conj_thm unique_mor
+      end;
+
+    val (ctor_fold_unique_thms, ctor_fold_unique_thm) =
+      `split_conj_thm (mk_conjIN n RS
+        (mor_UNIV_thm RS iffD2 RS fold_unique_mor_thm))
+
+    val fold_ctor_thms =
+      map (fn thm => (mor_incl_thm OF replicate n @{thm subset_UNIV}) RS thm RS sym)
+        fold_unique_mor_thms;
+
+    val ctor_o_fold_thms =
+      let
+        val mor = mor_comp_thm OF [mor_fold_thm, mor_str_thm];
+      in
+        map2 (fn unique => fn fold_ctor =>
+          trans OF [mor RS unique, fold_ctor]) fold_unique_mor_thms fold_ctor_thms
+      end;
+
+    val timer = time (timer "fold definitions & thms");
+
+    val map_ctors = map2 (fn Ds => fn bnf =>
+      Term.list_comb (mk_map_of_bnf Ds (passiveAs @ FTs) (passiveAs @ Ts) bnf,
+        map HOLogic.id_const passiveAs @ ctors)) Dss bnfs;
+
+    fun dtor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtorN ^ "_");
+    val dtor_name = Binding.name_of o dtor_bind;
+    val dtor_def_bind = rpair [] o Binding.conceal o Thm.def_binding o dtor_bind;
+
+    fun dtor_spec i FT T =
+      let
+        val dtorT = T --> FT;
+
+        val lhs = Free (dtor_name i, dtorT);
+        val rhs = mk_fold Ts map_ctors i;
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((dtor_frees, (_, dtor_def_frees)), (lthy, lthy_old)) =
+      lthy
+      |> fold_map3 (fn i => fn FT => fn T =>
+        Specification.definition
+          (SOME (dtor_bind i, NONE, NoSyn), (dtor_def_bind i, dtor_spec i FT T))) ks FTs Ts
+      |>> apsnd split_list o split_list
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+    fun mk_dtors params =
+      map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ params) o Morphism.term phi)
+        dtor_frees;
+    val dtors = mk_dtors params';
+    val dtor_defs = map (Morphism.thm phi) dtor_def_frees;
+
+    val ctor_o_dtor_thms = map2 (fold_thms lthy o single) dtor_defs ctor_o_fold_thms;
+
+    val dtor_o_ctor_thms =
+      let
+        fun mk_goal dtor ctor FT =
+          mk_Trueprop_eq (HOLogic.mk_comp (dtor, ctor), HOLogic.id_const FT);
+        val goals = map3 mk_goal dtors ctors FTs;
+      in
+        map5 (fn goal => fn dtor_def => fn foldx => fn map_comp_id => fn map_cong0L =>
+          Goal.prove_sorry lthy [] [] goal
+            (K (mk_dtor_o_ctor_tac dtor_def foldx map_comp_id map_cong0L ctor_o_fold_thms))
+          |> Thm.close_derivation)
+        goals dtor_defs ctor_fold_thms map_comp_id_thms map_cong0L_thms
+      end;
+
+    val dtor_ctor_thms = map (fn thm => thm RS @{thm pointfree_idE}) dtor_o_ctor_thms;
+    val ctor_dtor_thms = map (fn thm => thm RS @{thm pointfree_idE}) ctor_o_dtor_thms;
+
+    val bij_dtor_thms =
+      map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) ctor_o_dtor_thms dtor_o_ctor_thms;
+    val inj_dtor_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_dtor_thms;
+    val surj_dtor_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_dtor_thms;
+    val dtor_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_dtor_thms;
+    val dtor_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_dtor_thms;
+    val dtor_exhaust_thms = map (fn thm => thm RS exE) dtor_nchotomy_thms;
+
+    val bij_ctor_thms =
+      map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) dtor_o_ctor_thms ctor_o_dtor_thms;
+    val inj_ctor_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_ctor_thms;
+    val surj_ctor_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_ctor_thms;
+    val ctor_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_ctor_thms;
+    val ctor_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_ctor_thms;
+    val ctor_exhaust_thms = map (fn thm => thm RS exE) ctor_nchotomy_thms;
+
+    val timer = time (timer "dtor definitions & thms");
+
+    val fst_rec_pair_thms =
+      let
+        val mor = mor_comp_thm OF [mor_fold_thm, mor_convol_thm];
+      in
+        map2 (fn unique => fn fold_ctor =>
+          trans OF [mor RS unique, fold_ctor]) fold_unique_mor_thms fold_ctor_thms
+      end;
+
+    fun rec_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctor_recN ^ "_");
+    val rec_name = Binding.name_of o rec_bind;
+    val rec_def_bind = rpair [] o Binding.conceal o Thm.def_binding o rec_bind;
+
+    val rec_strs =
+      map3 (fn ctor => fn prod_s => fn mapx =>
+        mk_convol (HOLogic.mk_comp (ctor, Term.list_comb (mapx, passive_ids @ rec_fsts)), prod_s))
+      ctors rec_ss rec_maps;
+
+    fun rec_spec i T AT =
+      let
+        val recT = Library.foldr (op -->) (rec_sTs, T --> AT);
+
+        val lhs = Term.list_comb (Free (rec_name i, recT), rec_ss);
+        val rhs = HOLogic.mk_comp (snd_const (HOLogic.mk_prodT (T, AT)), mk_fold Ts rec_strs i);
+      in
+        mk_Trueprop_eq (lhs, rhs)
+      end;
+
+    val ((rec_frees, (_, rec_def_frees)), (lthy, lthy_old)) =
+      lthy
+      |> fold_map3 (fn i => fn T => fn AT =>
+        Specification.definition
+          (SOME (rec_bind i, NONE, NoSyn), (rec_def_bind i, rec_spec i T AT)))
+          ks Ts activeAs
+      |>> apsnd split_list o split_list
+      ||> `Local_Theory.restore;
+
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+    val recs = map (Morphism.term phi) rec_frees;
+    val rec_names = map (fst o dest_Const) recs;
+    fun mk_rec ss i = Term.list_comb (Const (nth rec_names (i - 1), Library.foldr (op -->)
+      (map fastype_of ss, nth Ts (i - 1) --> range_type (fastype_of (nth ss (i - 1))))), ss);
+    val rec_defs = map (Morphism.thm phi) rec_def_frees;
+
+    val convols = map2 (fn T => fn i => mk_convol (HOLogic.id_const T, mk_rec rec_ss i)) Ts ks;
+    val ctor_rec_thms =
+      let
+        fun mk_goal i rec_s rec_map ctor x =
+          let
+            val lhs = mk_rec rec_ss i $ (ctor $ x);
+            val rhs = rec_s $ (Term.list_comb (rec_map, passive_ids @ convols) $ x);
+          in
+            fold_rev Logic.all (x :: rec_ss) (mk_Trueprop_eq (lhs, rhs))
+          end;
+        val goals = map5 mk_goal ks rec_ss rec_maps_rev ctors xFs;
+      in
+        map2 (fn goal => fn foldx =>
+          Goal.prove_sorry lthy [] [] goal (mk_rec_tac rec_defs foldx fst_rec_pair_thms)
+          |> Thm.close_derivation)
+        goals ctor_fold_thms
+      end;
+
+    val rec_unique_mor_thm =
+      let
+        val id_fs = map2 (fn T => fn f => mk_convol (HOLogic.id_const T, f)) Ts fs;
+        val prem = HOLogic.mk_Trueprop (mk_mor UNIVs ctors rec_UNIVs rec_strs id_fs);
+        fun mk_fun_eq f i = HOLogic.mk_eq (f, mk_rec rec_ss i);
+        val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_fun_eq fs ks));
+      in
+        Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (rec_ss @ fs) (Logic.mk_implies (prem, unique)))
+          (mk_rec_unique_mor_tac rec_defs fst_rec_pair_thms fold_unique_mor_thm)
+          |> Thm.close_derivation
+      end;
+
+    val (ctor_rec_unique_thms, ctor_rec_unique_thm) =
+      `split_conj_thm (split_conj_prems n
+        (mor_UNIV_thm RS iffD2 RS rec_unique_mor_thm)
+        |> Local_Defs.unfold lthy (@{thms convol_o o_id id_o o_assoc[symmetric] fst_convol} @
+           map_id0s @ sym_map_comps) OF replicate n @{thm arg_cong2[of _ _ _ _ convol, OF refl]});
+
+    val timer = time (timer "rec definitions & thms");
+
+    val (ctor_induct_thm, induct_params) =
+      let
+        fun mk_prem phi ctor sets x =
+          let
+            fun mk_IH phi set z =
+              let
+                val prem = HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set $ x));
+                val concl = HOLogic.mk_Trueprop (phi $ z);
+              in
+                Logic.all z (Logic.mk_implies (prem, concl))
+              end;
+
+            val IHs = map3 mk_IH phis (drop m sets) Izs;
+            val concl = HOLogic.mk_Trueprop (phi $ (ctor $ x));
+          in
+            Logic.all x (Logic.list_implies (IHs, concl))
+          end;
+
+        val prems = map4 mk_prem phis ctors FTs_setss xFs;
+
+        fun mk_concl phi z = phi $ z;
+        val concl =
+          HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_concl phis Izs));
+
+        val goal = Logic.list_implies (prems, concl);
+      in
+        (Goal.prove_sorry lthy [] []
+          (fold_rev Logic.all (phis @ Izs) goal)
+          (K (mk_ctor_induct_tac lthy m set_mapss init_induct_thm morE_thms mor_Abs_thm
+            Rep_inverses Abs_inverses Reps))
+        |> Thm.close_derivation,
+        rev (Term.add_tfrees goal []))
+      end;
+
+    val cTs = map (SOME o certifyT lthy o TFree) induct_params;
+
+    val weak_ctor_induct_thms =
+      let fun insts i = (replicate (i - 1) TrueI) @ (asm_rl :: replicate (n - i) TrueI);
+      in map (fn i => (ctor_induct_thm OF insts i) RS mk_conjunctN n i) ks end;
+
+    val (ctor_induct2_thm, induct2_params) =
+      let
+        fun mk_prem phi ctor ctor' sets sets' x y =
+          let
+            fun mk_IH phi set set' z1 z2 =
+              let
+                val prem1 = HOLogic.mk_Trueprop (HOLogic.mk_mem (z1, (set $ x)));
+                val prem2 = HOLogic.mk_Trueprop (HOLogic.mk_mem (z2, (set' $ y)));
+                val concl = HOLogic.mk_Trueprop (phi $ z1 $ z2);
+              in
+                fold_rev Logic.all [z1, z2] (Logic.list_implies ([prem1, prem2], concl))
+              end;
+
+            val IHs = map5 mk_IH phi2s (drop m sets) (drop m sets') Izs1 Izs2;
+            val concl = HOLogic.mk_Trueprop (phi $ (ctor $ x) $ (ctor' $ y));
+          in
+            fold_rev Logic.all [x, y] (Logic.list_implies (IHs, concl))
+          end;
+
+        val prems = map7 mk_prem phi2s ctors ctor's FTs_setss FTs'_setss xFs yFs;
+
+        fun mk_concl phi z1 z2 = phi $ z1 $ z2;
+        val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+          (map3 mk_concl phi2s Izs1 Izs2));
+        fun mk_t phi (z1, z1') (z2, z2') =
+          Term.absfree z1' (HOLogic.mk_all (fst z2', snd z2', phi $ z1 $ z2));
+        val cts = map3 (SOME o certify lthy ooo mk_t) phi2s (Izs1 ~~ Izs1') (Izs2 ~~ Izs2');
+        val goal = Logic.list_implies (prems, concl);
+      in
+        (singleton (Proof_Context.export names_lthy lthy)
+          (Goal.prove_sorry lthy [] [] goal
+            (mk_ctor_induct2_tac cTs cts ctor_induct_thm weak_ctor_induct_thms))
+          |> Thm.close_derivation,
+        rev (Term.add_tfrees goal []))
+      end;
+
+    val timer = time (timer "induction");
+
+    fun mk_ctor_map_DEADID_thm ctor_inject map_id0 =
+      trans OF [id_apply, iffD2 OF [ctor_inject, map_id0 RS sym]];
+
+    fun mk_ctor_Irel_DEADID_thm ctor_inject bnf =
+      trans OF [ctor_inject, rel_eq_of_bnf bnf RS @{thm predicate2_eqD} RS sym];
+
+    val IphiTs = map2 mk_pred2T passiveAs passiveBs;
+    val Ipsi1Ts = map2 mk_pred2T passiveAs passiveCs;
+    val Ipsi2Ts = map2 mk_pred2T passiveCs passiveBs;
+    val activephiTs = map2 mk_pred2T activeAs activeBs;
+    val activeIphiTs = map2 mk_pred2T Ts Ts';
+    val (((((Iphis, Ipsi1s), Ipsi2s), activephis), activeIphis), names_lthy) = names_lthy
+      |> mk_Frees "R" IphiTs
+      ||>> mk_Frees "R" Ipsi1Ts
+      ||>> mk_Frees "Q" Ipsi2Ts
+      ||>> mk_Frees "S" activephiTs
+      ||>> mk_Frees "IR" activeIphiTs;
+    val rels = map2 (fn Ds => mk_rel_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
+
+    (*register new datatypes as BNFs*)
+    val (timer, Ibnfs, (ctor_Imap_o_thms, ctor_Imap_thms), ctor_Iset_thmss',
+        ctor_Irel_thms, Ibnf_notes, lthy) =
+      if m = 0 then
+        (timer, replicate n DEADID_bnf,
+        map_split (`(mk_pointfree lthy)) (map2 mk_ctor_map_DEADID_thm ctor_inject_thms map_ids),
+        replicate n [], map2 mk_ctor_Irel_DEADID_thm ctor_inject_thms bnfs, [], lthy)
+      else let
+        val fTs = map2 (curry op -->) passiveAs passiveBs;
+        val uTs = map2 (curry op -->) Ts Ts';
+
+        val (((((fs, fs'), fs_copy), us), (ys, ys')),
+          names_lthy) = names_lthy
+          |> mk_Frees' "f" fTs
+          ||>> mk_Frees "f" fTs
+          ||>> mk_Frees "u" uTs
+          ||>> mk_Frees' "y" passiveAs;
+
+        val map_FTFT's = map2 (fn Ds =>
+          mk_map_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
+        fun mk_passive_maps ATs BTs Ts =
+          map2 (fn Ds => mk_map_of_bnf Ds (ATs @ Ts) (BTs @ Ts)) Dss bnfs;
+        fun mk_map_fold_arg fs Ts ctor fmap =
+          HOLogic.mk_comp (ctor, Term.list_comb (fmap, fs @ map HOLogic.id_const Ts));
+        fun mk_map Ts fs Ts' ctors mk_maps =
+          mk_fold Ts (map2 (mk_map_fold_arg fs Ts') ctors (mk_maps Ts'));
+        val pmapsABT' = mk_passive_maps passiveAs passiveBs;
+        val fs_maps = map (mk_map Ts fs Ts' ctor's pmapsABT') ks;
+
+        val ls = 1 upto m;
+        val setsss = map (mk_setss o mk_set_Ts) passiveAs;
+
+        fun mk_col l T z z' sets =
+          let
+            fun mk_UN set = mk_Union T $ (set $ z);
+          in
+            Term.absfree z'
+              (mk_union (nth sets (l - 1) $ z,
+                Library.foldl1 mk_union (map mk_UN (drop m sets))))
+          end;
+
+        val colss = map5 (fn l => fn T => map3 (mk_col l T)) ls passiveAs AFss AFss' setsss;
+        val setss_by_range = map (fn cols => map (mk_fold Ts cols) ks) colss;
+        val setss_by_bnf = transpose setss_by_range;
+
+        val set_bss =
+          map (flat o map2 (fn B => fn b =>
+            if member (op =) resDs (TFree B) then [] else [b]) resBs) set_bss0;
+
+        val ctor_witss =
+          let
+            val witss = map2 (fn Ds => fn bnf => mk_wits_of_bnf
+              (replicate (nwits_of_bnf bnf) Ds)
+              (replicate (nwits_of_bnf bnf) (passiveAs @ Ts)) bnf) Dss bnfs;
+            fun close_wit (I, wit) = fold_rev Term.absfree (map (nth ys') I) wit;
+            fun wit_apply (arg_I, arg_wit) (fun_I, fun_wit) =
+              (union (op =) arg_I fun_I, fun_wit $ arg_wit);
+
+            fun gen_arg support i =
+              if i < m then [([i], nth ys i)]
+              else maps (mk_wit support (nth ctors (i - m)) (i - m)) (nth support (i - m))
+            and mk_wit support ctor i (I, wit) =
+              let val args = map (gen_arg (nth_map i (remove (op =) (I, wit)) support)) I;
+              in
+                (args, [([], wit)])
+                |-> fold (map_product wit_apply)
+                |> map (apsnd (fn t => ctor $ t))
+                |> minimize_wits
+              end;
+          in
+            map3 (fn ctor => fn i => map close_wit o minimize_wits o maps (mk_wit witss ctor i))
+              ctors (0 upto n - 1) witss
+          end;
+
+        val (Ibnf_consts, lthy) =
+          fold_map8 (fn b => fn map_b => fn rel_b => fn set_bs => fn mapx => fn sets => fn wits =>
+              fn T => fn lthy =>
+            define_bnf_consts Dont_Inline (user_policy Note_Some lthy) (SOME deads)
+              map_b rel_b set_bs
+              ((((((b, T), fold_rev Term.absfree fs' mapx), sets), sum_bd), wits), NONE) lthy)
+          bs map_bs rel_bs set_bss fs_maps setss_by_bnf ctor_witss Ts lthy;
+
+        val (_, Iconsts, Iconst_defs, mk_Iconsts) = split_list4 Ibnf_consts;
+        val (_, Isetss, Ibds_Ds, Iwitss_Ds, _) = split_list5 Iconsts;
+        val (Imap_defs, Iset_defss, Ibd_defs, Iwit_defss, Irel_defs) = split_list5 Iconst_defs;
+        val (mk_Imaps_Ds, mk_It_Ds, _, mk_Irels_Ds, _) = split_list5 mk_Iconsts;
+
+        val Irel_unabs_defs = map (fn def => mk_unabs_def m (def RS meta_eq_to_obj_eq)) Irel_defs;
+        val Iset_defs = flat Iset_defss;
+
+        fun mk_Imaps As Bs = map (fn mk => mk deads As Bs) mk_Imaps_Ds;
+        fun mk_Isetss As = map2 (fn mk => fn Isets => map (mk deads As) Isets) mk_It_Ds Isetss;
+        val Ibds = map2 (fn mk => mk deads passiveAs) mk_It_Ds Ibds_Ds;
+        val Iwitss =
+          map2 (fn mk => fn Iwits => map (mk deads passiveAs o snd) Iwits) mk_It_Ds Iwitss_Ds;
+        fun mk_Irels As Bs = map (fn mk => mk deads As Bs) mk_Irels_Ds;
+
+        val Imaps = mk_Imaps passiveAs passiveBs;
+        val fs_Imaps = map (fn m => Term.list_comb (m, fs)) Imaps;
+        val fs_copy_Imaps = map (fn m => Term.list_comb (m, fs_copy)) Imaps;
+        val (Isetss_by_range, Isetss_by_bnf) = `transpose (mk_Isetss passiveAs);
+
+        val map_setss = map (fn T => map2 (fn Ds =>
+          mk_map_of_bnf Ds (passiveAs @ Ts) (mk_set_Ts T)) Dss bnfs) passiveAs;
+
+        val timer = time (timer "bnf constants for the new datatypes");
+
+        val (ctor_Imap_thms, ctor_Imap_o_thms) =
+          let
+            fun mk_goal fs_map map ctor ctor' = fold_rev Logic.all fs
+              (mk_Trueprop_eq (HOLogic.mk_comp (fs_map, ctor),
+                HOLogic.mk_comp (ctor', Term.list_comb (map, fs @ fs_Imaps))));
+            val goals = map4 mk_goal fs_Imaps map_FTFT's ctors ctor's;
+            val maps =
+              map4 (fn goal => fn foldx => fn map_comp_id => fn map_cong0 =>
+                Goal.prove_sorry lthy [] [] goal
+                  (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Imap_defs THEN
+                    mk_map_tac m n foldx map_comp_id map_cong0)
+                |> Thm.close_derivation)
+              goals ctor_fold_thms map_comp_id_thms map_cong0s;
+          in
+            `(map (fn thm => thm RS @{thm comp_eq_dest})) maps
+          end;
+
+        val (ctor_Imap_unique_thms, ctor_Imap_unique_thm) =
+          let
+            fun mk_prem u map ctor ctor' =
+              mk_Trueprop_eq (HOLogic.mk_comp (u, ctor),
+                HOLogic.mk_comp (ctor', Term.list_comb (map, fs @ us)));
+            val prems = map4 mk_prem us map_FTFT's ctors ctor's;
+            val goal =
+              HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+                (map2 (curry HOLogic.mk_eq) us fs_Imaps));
+            val unique = Goal.prove_sorry lthy [] []
+              (fold_rev Logic.all (us @ fs) (Logic.list_implies (prems, goal)))
+              (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Imap_defs THEN
+                mk_ctor_map_unique_tac ctor_fold_unique_thm sym_map_comps ctxt)
+              |> Thm.close_derivation;
+          in
+            `split_conj_thm unique
+          end;
+
+        val timer = time (timer "map functions for the new datatypes");
+
+        val ctor_Iset_thmss =
+          let
+            fun mk_goal sets ctor set col map =
+              mk_Trueprop_eq (HOLogic.mk_comp (set, ctor),
+                HOLogic.mk_comp (col, Term.list_comb (map, passive_ids @ sets)));
+            val goalss =
+              map3 (fn sets => map4 (mk_goal sets) ctors sets) Isetss_by_range colss map_setss;
+            val setss = map (map2 (fn foldx => fn goal =>
+                Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} =>
+                  unfold_thms_tac ctxt Iset_defs THEN mk_set_tac foldx)
+                |> Thm.close_derivation)
+              ctor_fold_thms) goalss;
+
+            fun mk_simp_goal pas_set act_sets sets ctor z set =
+              Logic.all z (mk_Trueprop_eq (set $ (ctor $ z),
+                mk_union (pas_set $ z,
+                  Library.foldl1 mk_union (map2 (fn X => mk_UNION (X $ z)) act_sets sets))));
+            val simp_goalss =
+              map2 (fn i => fn sets =>
+                map4 (fn Fsets => mk_simp_goal (nth Fsets (i - 1)) (drop m Fsets) sets)
+                  FTs_setss ctors xFs sets)
+                ls Isetss_by_range;
+
+            val ctor_setss = map3 (fn i => map3 (fn set_nats => fn goal => fn set =>
+                Goal.prove_sorry lthy [] [] goal
+                  (K (mk_ctor_set_tac set (nth set_nats (i - 1)) (drop m set_nats)))
+                |> Thm.close_derivation)
+              set_mapss) ls simp_goalss setss;
+          in
+            ctor_setss
+          end;
+
+        fun mk_set_thms ctor_set = (@{thm xt1(3)} OF [ctor_set, @{thm Un_upper1}]) ::
+          map (fn i => (@{thm xt1(3)} OF [ctor_set, @{thm Un_upper2}]) RS
+            (mk_Un_upper n i RS subset_trans) RSN
+            (2, @{thm UN_upper} RS subset_trans))
+            (1 upto n);
+        val set_Iset_thmsss = transpose (map (map mk_set_thms) ctor_Iset_thmss);
+
+        val timer = time (timer "set functions for the new datatypes");
+
+        val cxs = map (SOME o certify lthy) Izs;
+        val Isetss_by_range' =
+          map (map (Term.subst_atomic_types (passiveAs ~~ passiveBs))) Isetss_by_range;
+
+        val Iset_Imap0_thmss =
+          let
+            fun mk_set_map0 f map z set set' =
+              HOLogic.mk_eq (mk_image f $ (set $ z), set' $ (map $ z));
+
+            fun mk_cphi f map z set set' = certify lthy
+              (Term.absfree (dest_Free z) (mk_set_map0 f map z set set'));
+
+            val csetss = map (map (certify lthy)) Isetss_by_range';
+
+            val cphiss = map3 (fn f => fn sets => fn sets' =>
+              (map4 (mk_cphi f) fs_Imaps Izs sets sets')) fs Isetss_by_range Isetss_by_range';
+
+            val inducts = map (fn cphis =>
+              Drule.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm) cphiss;
+
+            val goals =
+              map3 (fn f => fn sets => fn sets' =>
+                HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+                  (map4 (mk_set_map0 f) fs_Imaps Izs sets sets')))
+                  fs Isetss_by_range Isetss_by_range';
+
+            fun mk_tac induct = mk_set_nat_tac m (rtac induct) set_mapss ctor_Imap_thms;
+            val thms =
+              map5 (fn goal => fn csets => fn ctor_sets => fn induct => fn i =>
+                singleton (Proof_Context.export names_lthy lthy)
+                  (Goal.prove_sorry lthy [] [] goal (mk_tac induct csets ctor_sets i))
+                |> Thm.close_derivation)
+              goals csetss ctor_Iset_thmss inducts ls;
+          in
+            map split_conj_thm thms
+          end;
+
+        val Iset_bd_thmss =
+          let
+            fun mk_set_bd z bd set = mk_ordLeq (mk_card_of (set $ z)) bd;
+
+            fun mk_cphi z set = certify lthy (Term.absfree (dest_Free z) (mk_set_bd z sum_bd set));
+
+            val cphiss = map (map2 mk_cphi Izs) Isetss_by_range;
+
+            val inducts = map (fn cphis =>
+              Drule.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm) cphiss;
+
+            val goals =
+              map (fn sets =>
+                HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+                  (map3 mk_set_bd Izs Ibds sets))) Isetss_by_range;
+
+            fun mk_tac induct = mk_set_bd_tac m (rtac induct) sum_Cinfinite set_bd_sumss;
+            val thms =
+              map4 (fn goal => fn ctor_sets => fn induct => fn i =>
+                singleton (Proof_Context.export names_lthy lthy)
+                  (Goal.prove_sorry lthy [] [] goal
+                    (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Ibd_defs THEN
+                      mk_tac induct ctor_sets i ctxt))
+                |> Thm.close_derivation)
+              goals ctor_Iset_thmss inducts ls;
+          in
+            map split_conj_thm thms
+          end;
+
+        val Imap_cong0_thms =
+          let
+            fun mk_prem z set f g y y' =
+              mk_Ball (set $ z) (Term.absfree y' (HOLogic.mk_eq (f $ y, g $ y)));
+
+            fun mk_map_cong0 sets z fmap gmap =
+              HOLogic.mk_imp
+                (Library.foldr1 HOLogic.mk_conj (map5 (mk_prem z) sets fs fs_copy ys ys'),
+                HOLogic.mk_eq (fmap $ z, gmap $ z));
+
+            fun mk_cphi sets z fmap gmap =
+              certify lthy (Term.absfree (dest_Free z) (mk_map_cong0 sets z fmap gmap));
+
+            val cphis = map4 mk_cphi Isetss_by_bnf Izs fs_Imaps fs_copy_Imaps;
+
+            val induct = Drule.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm;
+
+            val goal =
+              HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+                (map4 mk_map_cong0 Isetss_by_bnf Izs fs_Imaps fs_copy_Imaps));
+
+            val thm = singleton (Proof_Context.export names_lthy lthy)
+              (Goal.prove_sorry lthy [] [] goal
+              (mk_mcong_tac (rtac induct) set_Iset_thmsss map_cong0s ctor_Imap_thms))
+              |> Thm.close_derivation;
+          in
+            split_conj_thm thm
+          end;
+
+        val in_rels = map in_rel_of_bnf bnfs;
+        val in_Irels = map (fn def => trans OF [def, @{thm OO_Grp_alt}] RS @{thm predicate2_eqD})
+            Irel_unabs_defs;
+
+        val ctor_Iset_incl_thmss = map (map hd) set_Iset_thmsss;
+        val ctor_set_Iset_incl_thmsss = map (transpose o map tl) set_Iset_thmsss;
+        val ctor_Iset_thmss' = transpose ctor_Iset_thmss;
+
+        val Irels = mk_Irels passiveAs passiveBs;
+        val Irelphis = map (fn rel => Term.list_comb (rel, Iphis)) Irels;
+        val relphis = map (fn rel => Term.list_comb (rel, Iphis @ Irelphis)) rels;
+        val Irelpsi1s = map (fn rel => Term.list_comb (rel, Ipsi1s)) (mk_Irels passiveAs passiveCs);
+        val Irelpsi2s = map (fn rel => Term.list_comb (rel, Ipsi2s)) (mk_Irels passiveCs passiveBs);
+        val Irelpsi12s = map (fn rel =>
+            Term.list_comb (rel, map2 (curry mk_rel_compp) Ipsi1s Ipsi2s)) Irels;
+
+        val ctor_Irel_thms =
+          let
+            fun mk_goal xF yF ctor ctor' Irelphi relphi = fold_rev Logic.all (xF :: yF :: Iphis)
+              (mk_Trueprop_eq (Irelphi $ (ctor $ xF) $ (ctor' $ yF), relphi $ xF $ yF));
+            val goals = map6 mk_goal xFs yFs ctors ctor's Irelphis relphis;
+          in
+            map12 (fn i => fn goal => fn in_rel => fn map_comp0 => fn map_cong0 =>
+              fn ctor_map => fn ctor_sets => fn ctor_inject => fn ctor_dtor =>
+              fn set_map0s => fn ctor_set_incls => fn ctor_set_set_inclss =>
+              Goal.prove_sorry lthy [] [] goal
+               (K (mk_ctor_rel_tac lthy in_Irels i in_rel map_comp0 map_cong0 ctor_map ctor_sets
+                 ctor_inject ctor_dtor set_map0s ctor_set_incls ctor_set_set_inclss))
+              |> Thm.close_derivation)
+            ks goals in_rels map_comps map_cong0s ctor_Imap_thms ctor_Iset_thmss'
+              ctor_inject_thms ctor_dtor_thms set_mapss ctor_Iset_incl_thmss
+              ctor_set_Iset_incl_thmsss
+          end;
+
+        val le_Irel_OO_thm =
+          let
+            fun mk_le_Irel_OO Irelpsi1 Irelpsi2 Irelpsi12 Iz1 Iz2 =
+              HOLogic.mk_imp (mk_rel_compp (Irelpsi1, Irelpsi2) $ Iz1 $ Iz2,
+                Irelpsi12 $ Iz1 $ Iz2);
+            val goals = map5 mk_le_Irel_OO Irelpsi1s Irelpsi2s Irelpsi12s Izs1 Izs2;
+
+            val cTs = map (SOME o certifyT lthy o TFree) induct2_params;
+            val cxs = map (SOME o certify lthy) (splice Izs1 Izs2);
+            fun mk_cphi z1 z2 goal = SOME (certify lthy (Term.absfree z1 (Term.absfree z2 goal)));
+            val cphis = map3 mk_cphi Izs1' Izs2' goals;
+            val induct = Drule.instantiate' cTs (cphis @ cxs) ctor_induct2_thm;
+
+            val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj goals);
+          in
+            singleton (Proof_Context.export names_lthy lthy)
+              (Goal.prove_sorry lthy [] [] goal
+                (mk_le_rel_OO_tac m induct ctor_nchotomy_thms ctor_Irel_thms rel_mono_strongs
+                  rel_OOs))
+              |> Thm.close_derivation
+          end;
+
+        val timer = time (timer "helpers for BNF properties");
+
+        val map_id0_tacs = map (K o mk_map_id0_tac map_id0s) ctor_Imap_unique_thms;
+        val map_comp0_tacs =
+          map2 (K oo mk_map_comp0_tac map_comps ctor_Imap_thms) ctor_Imap_unique_thms ks;
+        val map_cong0_tacs = map (mk_map_cong0_tac m) Imap_cong0_thms;
+        val set_map0_tacss = map (map (K o mk_set_map0_tac)) (transpose Iset_Imap0_thmss);
+        val bd_co_tacs = replicate n (fn {context = ctxt, prems = _} =>
+          unfold_thms_tac ctxt Ibd_defs THEN mk_bd_card_order_tac bd_card_orders);
+        val bd_cinf_tacs = replicate n (fn {context = ctxt, prems = _} =>
+          unfold_thms_tac ctxt Ibd_defs THEN rtac (sum_Cinfinite RS conjunct1) 1);
+        val set_bd_tacss = map (map (fn thm => K (rtac thm 1))) (transpose Iset_bd_thmss);
+        val le_rel_OO_tacs = map (fn i =>
+          K ((rtac @{thm predicate2I} THEN' etac (le_Irel_OO_thm RS mk_conjunctN n i RS mp)) 1)) ks;
+
+        val rel_OO_Grp_tacs = map (fn def => K (rtac def 1)) Irel_unabs_defs;
+
+        val tacss = map9 zip_axioms map_id0_tacs map_comp0_tacs map_cong0_tacs set_map0_tacss
+          bd_co_tacs bd_cinf_tacs set_bd_tacss le_rel_OO_tacs rel_OO_Grp_tacs;
+
+        fun wit_tac {context = ctxt, prems = _} = unfold_thms_tac ctxt (flat Iwit_defss) THEN
+          mk_wit_tac ctxt n (flat ctor_Iset_thmss) (maps wit_thms_of_bnf bnfs);
+
+        val (Ibnfs, lthy) =
+          fold_map5 (fn tacs => fn map_b => fn rel_b => fn set_bs => fn consts => fn lthy =>
+            bnf_def Do_Inline (user_policy Note_Some) I tacs wit_tac (SOME deads)
+              map_b rel_b set_bs consts lthy
+            |> register_bnf (Local_Theory.full_name lthy b))
+          tacss map_bs rel_bs set_bss
+            ((((((bs ~~ Ts) ~~ Imaps) ~~ Isetss_by_bnf) ~~ Ibds) ~~ Iwitss) ~~ map SOME Irels)
+            lthy;
+
+        val timer = time (timer "registered new datatypes as BNFs");
+
+        val ls' = if m = 1 then [0] else ls
+
+        val Ibnf_common_notes =
+          [(ctor_map_uniqueN, [ctor_Imap_unique_thm])]
+          |> map (fn (thmN, thms) =>
+            ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]));
+
+        val Ibnf_notes =
+          [(ctor_mapN, map single ctor_Imap_thms),
+          (ctor_relN, map single ctor_Irel_thms),
+          (ctor_set_inclN, ctor_Iset_incl_thmss),
+          (ctor_set_set_inclN, map flat ctor_set_Iset_incl_thmsss)] @
+          map2 (fn i => fn thms => (mk_ctor_setN i, map single thms)) ls' ctor_Iset_thmss
+          |> maps (fn (thmN, thmss) =>
+            map2 (fn b => fn thms =>
+              ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]))
+            bs thmss)
+      in
+        (timer, Ibnfs, (ctor_Imap_o_thms, ctor_Imap_thms), ctor_Iset_thmss',
+          ctor_Irel_thms, Ibnf_common_notes @ Ibnf_notes, lthy)
+      end;
+
+      val ctor_fold_o_Imap_thms = mk_xtor_un_fold_o_map_thms Least_FP false m ctor_fold_unique_thm
+        ctor_Imap_o_thms (map (mk_pointfree lthy) ctor_fold_thms) sym_map_comps map_cong0s;
+      val ctor_rec_o_Imap_thms = mk_xtor_un_fold_o_map_thms Least_FP true m ctor_rec_unique_thm
+        ctor_Imap_o_thms (map (mk_pointfree lthy) ctor_rec_thms) sym_map_comps map_cong0s;
+
+      val Irels = if m = 0 then map HOLogic.eq_const Ts
+        else map (mk_rel_of_bnf deads passiveAs passiveBs) Ibnfs;
+      val Irel_induct_thm =
+        mk_rel_xtor_co_induct_thm Least_FP rels activeIphis Irels Iphis xFs yFs ctors ctor's
+          (mk_rel_induct_tac m ctor_induct2_thm ks ctor_Irel_thms rel_mono_strongs) lthy;
+
+      val rels = map2 (fn Ds => mk_rel_of_bnf Ds allAs allBs') Dss bnfs;
+      val ctor_fold_transfer_thms =
+        mk_un_fold_transfer_thms Least_FP rels activephis Irels Iphis
+          (mk_folds passiveAs activeAs) (mk_folds passiveBs activeBs)
+          (mk_fold_transfer_tac m Irel_induct_thm (map map_transfer_of_bnf bnfs) ctor_fold_thms)
+          lthy;
+
+      val timer = time (timer "relator induction");
+
+      val common_notes =
+        [(ctor_inductN, [ctor_induct_thm]),
+        (ctor_induct2N, [ctor_induct2_thm]),
+        (rel_inductN, [Irel_induct_thm]),
+        (ctor_fold_transferN, ctor_fold_transfer_thms)]
+        |> map (fn (thmN, thms) =>
+          ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]));
+
+      val notes =
+        [(ctor_dtorN, ctor_dtor_thms),
+        (ctor_exhaustN, ctor_exhaust_thms),
+        (ctor_foldN, ctor_fold_thms),
+        (ctor_fold_uniqueN, ctor_fold_unique_thms),
+        (ctor_rec_uniqueN, ctor_rec_unique_thms),
+        (ctor_fold_o_mapN, ctor_fold_o_Imap_thms),
+        (ctor_rec_o_mapN, ctor_rec_o_Imap_thms),
+        (ctor_injectN, ctor_inject_thms),
+        (ctor_recN, ctor_rec_thms),
+        (dtor_ctorN, dtor_ctor_thms),
+        (dtor_exhaustN, dtor_exhaust_thms),
+        (dtor_injectN, dtor_inject_thms)]
+        |> map (apsnd (map single))
+        |> maps (fn (thmN, thmss) =>
+          map2 (fn b => fn thms =>
+            ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]))
+          bs thmss);
+
+    (*FIXME: once the package exports all the necessary high-level characteristic theorems,
+       those should not only be concealed but rather not noted at all*)
+    val maybe_conceal_notes = note_all = false ? map (apfst (apfst Binding.conceal));
+  in
+    timer;
+    ({Ts = Ts, bnfs = Ibnfs, ctors = ctors, dtors = dtors, xtor_co_iterss = transpose [folds, recs],
+      xtor_co_induct = ctor_induct_thm, dtor_ctors = dtor_ctor_thms, ctor_dtors = ctor_dtor_thms,
+      ctor_injects = ctor_inject_thms, dtor_injects = dtor_inject_thms,
+      xtor_map_thms = ctor_Imap_thms, xtor_set_thmss = ctor_Iset_thmss',
+      xtor_rel_thms = ctor_Irel_thms,
+      xtor_co_iter_thmss = transpose [ctor_fold_thms, ctor_rec_thms],
+      xtor_co_iter_o_map_thmss = transpose [ctor_fold_o_Imap_thms, ctor_rec_o_Imap_thms],
+      rel_xtor_co_induct_thm = Irel_induct_thm},
+     lthy |> Local_Theory.notes (maybe_conceal_notes (common_notes @ notes @ Ibnf_notes)) |> snd)
+  end;
+
+val _ =
+  Outer_Syntax.local_theory @{command_spec "datatype_new"} "define new-style inductive datatypes"
+    (parse_co_datatype_cmd Least_FP construct_lfp);
+
+val _ = Outer_Syntax.local_theory @{command_spec "primrec_new"}
+  "define primitive recursive functions"
+  (Parse.fixes -- Parse_Spec.where_alt_specs >> (snd oo uncurry add_primrec_cmd));
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_lfp_compat.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,172 @@
+(*  Title:      HOL/BNF/Tools/bnf_lfp_compat.ML
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2013
+
+Compatibility layer with the old datatype package.
+*)
+
+signature BNF_LFP_COMPAT =
+sig
+  val datatype_new_compat_cmd : string list -> local_theory -> local_theory
+end;
+
+structure BNF_LFP_Compat : BNF_LFP_COMPAT =
+struct
+
+open Ctr_Sugar
+open BNF_Util
+open BNF_FP_Util
+open BNF_FP_Def_Sugar
+open BNF_FP_N2M_Sugar
+
+fun dtyp_of_typ _ (TFree a) = Datatype_Aux.DtTFree a
+  | dtyp_of_typ recTs (T as Type (s, Ts)) =
+    (case find_index (curry (op =) T) recTs of
+      ~1 => Datatype_Aux.DtType (s, map (dtyp_of_typ recTs) Ts)
+    | kk => Datatype_Aux.DtRec kk);
+
+val compatN = "compat_";
+
+(* TODO: graceful failure for local datatypes -- perhaps by making the command global *)
+fun datatype_new_compat_cmd raw_fpT_names lthy =
+  let
+    val thy = Proof_Context.theory_of lthy;
+
+    fun not_datatype s = error (quote s ^ " is not a new-style datatype");
+    fun not_mutually_recursive ss =
+      error ("{" ^ commas ss ^ "} is not a complete set of mutually recursive new-style datatypes");
+
+    val (fpT_names as fpT_name1 :: _) =
+      map (fst o dest_Type o Proof_Context.read_type_name_proper lthy false) raw_fpT_names;
+
+    fun lfp_sugar_of s =
+      (case fp_sugar_of lthy s of
+        SOME (fp_sugar as {fp = Least_FP, ...}) => fp_sugar
+      | _ => not_datatype s);
+
+    val {ctr_sugars, ...} = lfp_sugar_of fpT_name1;
+    val fpTs0 as Type (_, var_As) :: _ = map (body_type o fastype_of o hd o #ctrs) ctr_sugars;
+    val fpT_names' = map (fst o dest_Type) fpTs0;
+
+    val _ = eq_set (op =) (fpT_names, fpT_names') orelse not_mutually_recursive fpT_names;
+
+    val (unsorted_As, _) = lthy |> mk_TFrees (length var_As);
+    val As = map2 (resort_tfree o Type.sort_of_atyp) var_As unsorted_As;
+    val fpTs as fpT1 :: _ = map (fn s => Type (s, As)) fpT_names';
+
+    fun add_nested_types_of (T as Type (s, _)) seen =
+      if member (op =) seen T then
+        seen
+      else if s = @{type_name fun} then
+        (warning "Partial support for recursion through functions -- 'primrec' will fail"; seen)
+      else
+        (case try lfp_sugar_of s of
+          SOME ({T = T0, fp_res = {Ts = mutual_Ts0, ...}, ctr_sugars, ...}) =>
+          let
+            val rho = Vartab.fold (cons o apsnd snd) (Sign.typ_match thy (T0, T) Vartab.empty) [];
+            val substT = Term.typ_subst_TVars rho;
+
+            val mutual_Ts = map substT mutual_Ts0;
+
+            fun add_interesting_subtypes (U as Type (_, Us)) =
+                (case filter (exists_subtype_in mutual_Ts) Us of [] => I
+                | Us' => insert (op =) U #> fold add_interesting_subtypes Us')
+              | add_interesting_subtypes _ = I;
+
+            val ctrs = maps #ctrs ctr_sugars;
+            val ctr_Ts = maps (binder_types o substT o fastype_of) ctrs |> distinct (op =);
+            val subTs = fold add_interesting_subtypes ctr_Ts [];
+          in
+            fold add_nested_types_of subTs (seen @ mutual_Ts)
+          end
+        | NONE => error ("Unsupported recursion via type constructor " ^ quote s ^
+            " not corresponding to new-style datatype (cf. \"datatype_new\")"));
+
+    val Ts = add_nested_types_of fpT1 [];
+    val b_names = map base_name_of_typ Ts;
+    val compat_b_names = map (prefix compatN) b_names;
+    val compat_bs = map Binding.name compat_b_names;
+    val common_name = compatN ^ mk_common_name b_names;
+    val nn_fp = length fpTs;
+    val nn = length Ts;
+    val get_indices = K [];
+    val fp_sugars0 = map (lfp_sugar_of o fst o dest_Type) Ts;
+    val callssss = map (fn fp_sugar0 => indexify_callsss fp_sugar0 []) fp_sugars0;
+
+    val ((fp_sugars, (lfp_sugar_thms, _)), lthy) =
+      if nn > nn_fp then
+        mutualize_fp_sugars Least_FP compat_bs Ts get_indices callssss fp_sugars0 lthy
+      else
+        ((fp_sugars0, (NONE, NONE)), lthy);
+
+    val {ctr_sugars, co_inducts = [induct], co_iterss, co_iter_thmsss = iter_thmsss, ...} :: _ =
+      fp_sugars;
+    val inducts = conj_dests nn induct;
+
+    val mk_dtyp = dtyp_of_typ Ts;
+
+    fun mk_ctr_descr Ts = mk_ctr Ts #> dest_Const ##> (binder_types #> map mk_dtyp);
+    fun mk_typ_descr index (Type (T_name, Ts)) ({ctrs, ...} : ctr_sugar) =
+      (index, (T_name, map mk_dtyp Ts, map (mk_ctr_descr Ts) ctrs));
+
+    val descr = map3 mk_typ_descr (0 upto nn - 1) Ts ctr_sugars;
+    val recs = map (fst o dest_Const o co_rec_of) co_iterss;
+    val rec_thms = flat (map co_rec_of iter_thmsss);
+
+    fun mk_info ({T = Type (T_name0, _), index, ...} : fp_sugar) =
+      let
+        val {casex, exhaust, nchotomy, injects, distincts, case_thms, case_cong, weak_case_cong,
+          split, split_asm, ...} = nth ctr_sugars index;
+      in
+        (T_name0,
+         {index = index, descr = descr, inject = injects, distinct = distincts, induct = induct,
+         inducts = inducts, exhaust = exhaust, nchotomy = nchotomy, rec_names = recs,
+         rec_rewrites = rec_thms, case_name = fst (dest_Const casex), case_rewrites = case_thms,
+         case_cong = case_cong, weak_case_cong = weak_case_cong, split = split,
+         split_asm = split_asm})
+      end;
+
+    val infos = map mk_info (take nn_fp fp_sugars);
+
+    val all_notes =
+      (case lfp_sugar_thms of
+        NONE => []
+      | SOME ((induct_thms, induct_thm, induct_attrs), (fold_thmss, rec_thmss, _)) =>
+        let
+          val common_notes =
+            (if nn > 1 then [(inductN, [induct_thm], induct_attrs)] else [])
+            |> filter_out (null o #2)
+            |> map (fn (thmN, thms, attrs) =>
+              ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
+
+          val notes =
+            [(foldN, fold_thmss, []),
+             (inductN, map single induct_thms, induct_attrs),
+             (recN, rec_thmss, [])]
+            |> filter_out (null o #2)
+            |> maps (fn (thmN, thmss, attrs) =>
+              if forall null thmss then
+                []
+              else
+                map2 (fn b_name => fn thms =>
+                    ((Binding.qualify true b_name (Binding.name thmN), attrs), [(thms, [])]))
+                  compat_b_names thmss);
+        in
+          common_notes @ notes
+        end);
+
+    val register_interpret =
+      Datatype_Data.register infos
+      #> Datatype_Data.interpretation_data (Datatype_Aux.default_config, map fst infos)
+  in
+    lthy
+    |> Local_Theory.raw_theory register_interpret
+    |> Local_Theory.notes all_notes |> snd
+  end;
+
+val _ =
+  Outer_Syntax.local_theory @{command_spec "datatype_new_compat"}
+    "register new-style datatypes as old-style datatypes"
+    (Scan.repeat1 Parse.type_const >> datatype_new_compat_cmd);
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_lfp_rec_sugar.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,606 @@
+(*  Title:      HOL/BNF/Tools/bnf_lfp_rec_sugar.ML
+    Author:     Lorenz Panny, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2013
+
+Recursor sugar.
+*)
+
+signature BNF_LFP_REC_SUGAR =
+sig
+  val add_primrec: (binding * typ option * mixfix) list ->
+    (Attrib.binding * term) list -> local_theory -> (term list * thm list list) * local_theory
+  val add_primrec_cmd: (binding * string option * mixfix) list ->
+    (Attrib.binding * string) list -> local_theory -> (term list * thm list list) * local_theory
+  val add_primrec_global: (binding * typ option * mixfix) list ->
+    (Attrib.binding * term) list -> theory -> (term list * thm list list) * theory
+  val add_primrec_overloaded: (string * (string * typ) * bool) list ->
+    (binding * typ option * mixfix) list ->
+    (Attrib.binding * term) list -> theory -> (term list * thm list list) * theory
+  val add_primrec_simple: ((binding * typ) * mixfix) list -> term list ->
+    local_theory -> (string list * (term list * (int list list * thm list list))) * local_theory
+end;
+
+structure BNF_LFP_Rec_Sugar : BNF_LFP_REC_SUGAR =
+struct
+
+open Ctr_Sugar
+open BNF_Util
+open BNF_Tactics
+open BNF_Def
+open BNF_FP_Util
+open BNF_FP_Def_Sugar
+open BNF_FP_N2M_Sugar
+open BNF_FP_Rec_Sugar_Util
+
+val nitpicksimp_attrs = @{attributes [nitpick_simp]};
+val simp_attrs = @{attributes [simp]};
+val code_nitpicksimp_simp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs @ simp_attrs;
+
+exception Primrec_Error of string * term list;
+
+fun primrec_error str = raise Primrec_Error (str, []);
+fun primrec_error_eqn str eqn = raise Primrec_Error (str, [eqn]);
+fun primrec_error_eqns str eqns = raise Primrec_Error (str, eqns);
+
+datatype rec_call =
+  No_Rec of int * typ |
+  Mutual_Rec of (int * typ) * (int * typ) |
+  Nested_Rec of int * typ;
+
+type rec_ctr_spec =
+  {ctr: term,
+   offset: int,
+   calls: rec_call list,
+   rec_thm: thm};
+
+type rec_spec =
+  {recx: term,
+   nested_map_idents: thm list,
+   nested_map_comps: thm list,
+   ctr_specs: rec_ctr_spec list};
+
+exception AINT_NO_MAP of term;
+
+fun ill_formed_rec_call ctxt t =
+  error ("Ill-formed recursive call: " ^ quote (Syntax.string_of_term ctxt t));
+fun invalid_map ctxt t =
+  error ("Invalid map function in " ^ quote (Syntax.string_of_term ctxt t));
+fun unexpected_rec_call ctxt t =
+  error ("Unexpected recursive call: " ^ quote (Syntax.string_of_term ctxt t));
+
+fun massage_nested_rec_call ctxt has_call raw_massage_fun bound_Ts y y' =
+  let
+    fun check_no_call t = if has_call t then unexpected_rec_call ctxt t else ();
+
+    val typof = curry fastype_of1 bound_Ts;
+    val build_map_fst = build_map ctxt (fst_const o fst);
+
+    val yT = typof y;
+    val yU = typof y';
+
+    fun y_of_y' () = build_map_fst (yU, yT) $ y';
+    val elim_y = Term.map_aterms (fn t => if t = y then y_of_y' () else t);
+
+    fun massage_mutual_fun U T t =
+      (case t of
+        Const (@{const_name comp}, _) $ t1 $ t2 =>
+        mk_comp bound_Ts (tap check_no_call t1, massage_mutual_fun U T t2)
+      | _ =>
+        if has_call t then
+          (case try HOLogic.dest_prodT U of
+            SOME (U1, U2) => if U1 = T then raw_massage_fun T U2 t else invalid_map ctxt t
+          | NONE => invalid_map ctxt t)
+        else
+          mk_comp bound_Ts (t, build_map_fst (U, T)));
+
+    fun massage_map (Type (_, Us)) (Type (s, Ts)) t =
+        (case try (dest_map ctxt s) t of
+          SOME (map0, fs) =>
+          let
+            val Type (_, ran_Ts) = range_type (typof t);
+            val map' = mk_map (length fs) Us ran_Ts map0;
+            val fs' = map_flattened_map_args ctxt s (map3 massage_map_or_map_arg Us Ts) fs;
+          in
+            Term.list_comb (map', fs')
+          end
+        | NONE => raise AINT_NO_MAP t)
+      | massage_map _ _ t = raise AINT_NO_MAP t
+    and massage_map_or_map_arg U T t =
+      if T = U then
+        tap check_no_call t
+      else
+        massage_map U T t
+        handle AINT_NO_MAP _ => massage_mutual_fun U T t;
+
+    fun massage_call (t as t1 $ t2) =
+        if has_call t then
+          if t2 = y then
+            massage_map yU yT (elim_y t1) $ y'
+            handle AINT_NO_MAP t' => invalid_map ctxt t'
+          else
+            let val (g, xs) = Term.strip_comb t2 in
+              if g = y then
+                if exists has_call xs then unexpected_rec_call ctxt t2
+                else Term.list_comb (massage_call (mk_compN (length xs) bound_Ts (t1, y)), xs)
+              else
+                ill_formed_rec_call ctxt t
+            end
+        else
+          elim_y t
+      | massage_call t = if t = y then y_of_y' () else ill_formed_rec_call ctxt t;
+  in
+    massage_call
+  end;
+
+fun rec_specs_of bs arg_Ts res_Ts get_indices callssss0 lthy0 =
+  let
+    val thy = Proof_Context.theory_of lthy0;
+
+    val ((missing_arg_Ts, perm0_kks,
+          fp_sugars as {nested_bnfs, fp_res = {xtor_co_iterss = ctor_iters1 :: _, ...},
+            co_inducts = [induct_thm], ...} :: _, (lfp_sugar_thms, _)), lthy) =
+      nested_to_mutual_fps Least_FP bs arg_Ts get_indices callssss0 lthy0;
+
+    val perm_fp_sugars = sort (int_ord o pairself #index) fp_sugars;
+
+    val indices = map #index fp_sugars;
+    val perm_indices = map #index perm_fp_sugars;
+
+    val perm_ctrss = map (#ctrs o of_fp_sugar #ctr_sugars) perm_fp_sugars;
+    val perm_ctr_Tsss = map (map (binder_types o fastype_of)) perm_ctrss;
+    val perm_lfpTs = map (body_type o fastype_of o hd) perm_ctrss;
+
+    val nn0 = length arg_Ts;
+    val nn = length perm_lfpTs;
+    val kks = 0 upto nn - 1;
+    val perm_ns = map length perm_ctr_Tsss;
+    val perm_mss = map (map length) perm_ctr_Tsss;
+
+    val perm_Cs = map (body_type o fastype_of o co_rec_of o of_fp_sugar (#xtor_co_iterss o #fp_res))
+      perm_fp_sugars;
+    val perm_fun_arg_Tssss =
+      mk_iter_fun_arg_types perm_ctr_Tsss perm_ns perm_mss (co_rec_of ctor_iters1);
+
+    fun unpermute0 perm0_xs = permute_like (op =) perm0_kks kks perm0_xs;
+    fun unpermute perm_xs = permute_like (op =) perm_indices indices perm_xs;
+
+    val induct_thms = unpermute0 (conj_dests nn induct_thm);
+
+    val lfpTs = unpermute perm_lfpTs;
+    val Cs = unpermute perm_Cs;
+
+    val As_rho = tvar_subst thy (take nn0 lfpTs) arg_Ts;
+    val Cs_rho = map (fst o dest_TVar) Cs ~~ pad_list HOLogic.unitT nn res_Ts;
+
+    val substA = Term.subst_TVars As_rho;
+    val substAT = Term.typ_subst_TVars As_rho;
+    val substCT = Term.typ_subst_TVars Cs_rho;
+    val substACT = substAT o substCT;
+
+    val perm_Cs' = map substCT perm_Cs;
+
+    fun offset_of_ctr 0 _ = 0
+      | offset_of_ctr n (({ctrs, ...} : ctr_sugar) :: ctr_sugars) =
+        length ctrs + offset_of_ctr (n - 1) ctr_sugars;
+
+    fun call_of [i] [T] = (if exists_subtype_in Cs T then Nested_Rec else No_Rec) (i, substACT T)
+      | call_of [i, i'] [T, T'] = Mutual_Rec ((i, substACT T), (i', substACT T'));
+
+    fun mk_ctr_spec ctr offset fun_arg_Tss rec_thm =
+      let
+        val (fun_arg_hss, _) = indexedd fun_arg_Tss 0;
+        val fun_arg_hs = flat_rec_arg_args fun_arg_hss;
+        val fun_arg_iss = map (map (find_index_eq fun_arg_hs)) fun_arg_hss;
+      in
+        {ctr = substA ctr, offset = offset, calls = map2 call_of fun_arg_iss fun_arg_Tss,
+         rec_thm = rec_thm}
+      end;
+
+    fun mk_ctr_specs index (ctr_sugars : ctr_sugar list) iter_thmsss =
+      let
+        val ctrs = #ctrs (nth ctr_sugars index);
+        val rec_thms = co_rec_of (nth iter_thmsss index);
+        val k = offset_of_ctr index ctr_sugars;
+        val n = length ctrs;
+      in
+        map4 mk_ctr_spec ctrs (k upto k + n - 1) (nth perm_fun_arg_Tssss index) rec_thms
+      end;
+
+    fun mk_spec ({T, index, ctr_sugars, co_iterss = iterss, co_iter_thmsss = iter_thmsss, ...}
+      : fp_sugar) =
+      {recx = mk_co_iter thy Least_FP (substAT T) perm_Cs' (co_rec_of (nth iterss index)),
+       nested_map_idents = map (unfold_thms lthy @{thms id_def} o map_id0_of_bnf) nested_bnfs,
+       nested_map_comps = map map_comp_of_bnf nested_bnfs,
+       ctr_specs = mk_ctr_specs index ctr_sugars iter_thmsss};
+  in
+    ((is_some lfp_sugar_thms, map mk_spec fp_sugars, missing_arg_Ts, induct_thm, induct_thms), lthy)
+  end;
+
+val undef_const = Const (@{const_name undefined}, dummyT);
+
+fun permute_args n t =
+  list_comb (t, map Bound (0 :: (n downto 1))) |> fold (K (Term.abs (Name.uu, dummyT))) (0 upto n);
+
+type eqn_data = {
+  fun_name: string,
+  rec_type: typ,
+  ctr: term,
+  ctr_args: term list,
+  left_args: term list,
+  right_args: term list,
+  res_type: typ,
+  rhs_term: term,
+  user_eqn: term
+};
+
+fun dissect_eqn lthy fun_names eqn' =
+  let
+    val eqn = drop_all eqn' |> HOLogic.dest_Trueprop
+      handle TERM _ =>
+        primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
+    val (lhs, rhs) = HOLogic.dest_eq eqn
+        handle TERM _ =>
+          primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
+    val (fun_name, args) = strip_comb lhs
+      |>> (fn x => if is_Free x then fst (dest_Free x)
+          else primrec_error_eqn "malformed function equation (does not start with free)" eqn);
+    val (left_args, rest) = take_prefix is_Free args;
+    val (nonfrees, right_args) = take_suffix is_Free rest;
+    val num_nonfrees = length nonfrees;
+    val _ = num_nonfrees = 1 orelse if num_nonfrees = 0 then
+      primrec_error_eqn "constructor pattern missing in left-hand side" eqn else
+      primrec_error_eqn "more than one non-variable argument in left-hand side" eqn;
+    val _ = member (op =) fun_names fun_name orelse
+      primrec_error_eqn "malformed function equation (does not start with function name)" eqn
+
+    val (ctr, ctr_args) = strip_comb (the_single nonfrees);
+    val _ = try (num_binder_types o fastype_of) ctr = SOME (length ctr_args) orelse
+      primrec_error_eqn "partially applied constructor in pattern" eqn;
+    val _ = let val d = duplicates (op =) (left_args @ ctr_args @ right_args) in null d orelse
+      primrec_error_eqn ("duplicate variable \"" ^ Syntax.string_of_term lthy (hd d) ^
+        "\" in left-hand side") eqn end;
+    val _ = forall is_Free ctr_args orelse
+      primrec_error_eqn "non-primitive pattern in left-hand side" eqn;
+    val _ =
+      let val b = fold_aterms (fn x as Free (v, _) =>
+        if (not (member (op =) (left_args @ ctr_args @ right_args) x) andalso
+        not (member (op =) fun_names v) andalso
+        not (Variable.is_fixed lthy v)) then cons x else I | _ => I) rhs []
+      in
+        null b orelse
+        primrec_error_eqn ("extra variable(s) in right-hand side: " ^
+          commas (map (Syntax.string_of_term lthy) b)) eqn
+      end;
+  in
+    {fun_name = fun_name,
+     rec_type = body_type (type_of ctr),
+     ctr = ctr,
+     ctr_args = ctr_args,
+     left_args = left_args,
+     right_args = right_args,
+     res_type = map fastype_of (left_args @ right_args) ---> fastype_of rhs,
+     rhs_term = rhs,
+     user_eqn = eqn'}
+  end;
+
+fun rewrite_map_arg get_ctr_pos rec_type res_type =
+  let
+    val pT = HOLogic.mk_prodT (rec_type, res_type);
+
+    fun subst d (t as Bound d') = t |> d = SOME d' ? curry (op $) (fst_const pT)
+      | subst d (Abs (v, T, b)) =
+        Abs (v, if d = SOME ~1 then pT else T, subst (Option.map (Integer.add 1) d) b)
+      | subst d t =
+        let
+          val (u, vs) = strip_comb t;
+          val ctr_pos = try (get_ctr_pos o fst o dest_Free) u |> the_default ~1;
+        in
+          if ctr_pos >= 0 then
+            if d = SOME ~1 andalso length vs = ctr_pos then
+              list_comb (permute_args ctr_pos (snd_const pT), vs)
+            else if length vs > ctr_pos andalso is_some d
+                andalso d = try (fn Bound n => n) (nth vs ctr_pos) then
+              list_comb (snd_const pT $ nth vs ctr_pos, map (subst d) (nth_drop ctr_pos vs))
+            else
+              primrec_error_eqn ("recursive call not directly applied to constructor argument") t
+          else
+            list_comb (u, map (subst (d |> d = SOME ~1 ? K NONE)) vs)
+        end
+  in
+    subst (SOME ~1)
+  end;
+
+fun subst_rec_calls lthy get_ctr_pos has_call ctr_args mutual_calls nested_calls =
+  let
+    fun try_nested_rec bound_Ts y t =
+      AList.lookup (op =) nested_calls y
+      |> Option.map (fn y' =>
+        massage_nested_rec_call lthy has_call (rewrite_map_arg get_ctr_pos) bound_Ts y y' t);
+
+    fun subst bound_Ts (t as g' $ y) =
+        let
+          fun subst_rec () = subst bound_Ts g' $ subst bound_Ts y;
+          val y_head = head_of y;
+        in
+          if not (member (op =) ctr_args y_head) then
+            subst_rec ()
+          else
+            (case try_nested_rec bound_Ts y_head t of
+              SOME t' => t'
+            | NONE =>
+              let val (g, g_args) = strip_comb g' in
+                (case try (get_ctr_pos o fst o dest_Free) g of
+                  SOME ctr_pos =>
+                  (length g_args >= ctr_pos orelse
+                   primrec_error_eqn "too few arguments in recursive call" t;
+                   (case AList.lookup (op =) mutual_calls y of
+                     SOME y' => list_comb (y', g_args)
+                   | NONE => subst_rec ()))
+                | NONE => subst_rec ())
+              end)
+        end
+      | subst bound_Ts (Abs (v, T, b)) = Abs (v, T, subst (T :: bound_Ts) b)
+      | subst _ t = t
+
+    fun subst' t =
+      if has_call t then
+        (* FIXME detect this case earlier? *)
+        primrec_error_eqn "recursive call not directly applied to constructor argument" t
+      else
+        try_nested_rec [] (head_of t) t |> the_default t
+  in
+    subst' o subst []
+  end;
+
+fun build_rec_arg lthy (funs_data : eqn_data list list) has_call (ctr_spec : rec_ctr_spec)
+    (eqn_data_opt : eqn_data option) =
+  (case eqn_data_opt of
+    NONE => undef_const
+  | SOME {ctr_args, left_args, right_args, rhs_term = t, ...} =>
+    let
+      val calls = #calls ctr_spec;
+      val n_args = fold (Integer.add o (fn Mutual_Rec _ => 2 | _ => 1)) calls 0;
+
+      val no_calls' = tag_list 0 calls
+        |> map_filter (try (apsnd (fn No_Rec p => p | Mutual_Rec (p, _) => p)));
+      val mutual_calls' = tag_list 0 calls
+        |> map_filter (try (apsnd (fn Mutual_Rec (_, p) => p)));
+      val nested_calls' = tag_list 0 calls
+        |> map_filter (try (apsnd (fn Nested_Rec p => p)));
+
+      fun ensure_unique frees t =
+        if member (op =) frees t then Free (the_single (Term.variant_frees t [dest_Free t])) else t;
+
+      val args = replicate n_args ("", dummyT)
+        |> Term.rename_wrt_term t
+        |> map Free
+        |> fold (fn (ctr_arg_idx, (arg_idx, _)) =>
+            nth_map arg_idx (K (nth ctr_args ctr_arg_idx)))
+          no_calls'
+        |> fold (fn (ctr_arg_idx, (arg_idx, T)) => fn xs =>
+            nth_map arg_idx (K (ensure_unique xs (retype_free T (nth ctr_args ctr_arg_idx)))) xs)
+          mutual_calls'
+        |> fold (fn (ctr_arg_idx, (arg_idx, T)) =>
+            nth_map arg_idx (K (retype_free T (nth ctr_args ctr_arg_idx))))
+          nested_calls';
+
+      val fun_name_ctr_pos_list =
+        map (fn (x :: _) => (#fun_name x, length (#left_args x))) funs_data;
+      val get_ctr_pos = try (the o AList.lookup (op =) fun_name_ctr_pos_list) #> the_default ~1;
+      val mutual_calls = map (apfst (nth ctr_args) o apsnd (nth args o fst)) mutual_calls';
+      val nested_calls = map (apfst (nth ctr_args) o apsnd (nth args o fst)) nested_calls';
+    in
+      t
+      |> subst_rec_calls lthy get_ctr_pos has_call ctr_args mutual_calls nested_calls
+      |> fold_rev lambda (args @ left_args @ right_args)
+    end);
+
+fun build_defs lthy bs mxs (funs_data : eqn_data list list) (rec_specs : rec_spec list) has_call =
+  let
+    val n_funs = length funs_data;
+
+    val ctr_spec_eqn_data_list' =
+      (take n_funs rec_specs |> map #ctr_specs) ~~ funs_data
+      |> maps (uncurry (finds (fn (x, y) => #ctr x = #ctr y))
+          ##> (fn x => null x orelse
+            primrec_error_eqns "excess equations in definition" (map #rhs_term x)) #> fst);
+    val _ = ctr_spec_eqn_data_list' |> map (fn (_, x) => length x <= 1 orelse
+      primrec_error_eqns ("multiple equations for constructor") (map #user_eqn x));
+
+    val ctr_spec_eqn_data_list =
+      ctr_spec_eqn_data_list' @ (drop n_funs rec_specs |> maps #ctr_specs |> map (rpair []));
+
+    val recs = take n_funs rec_specs |> map #recx;
+    val rec_args = ctr_spec_eqn_data_list
+      |> sort ((op <) o pairself (#offset o fst) |> make_ord)
+      |> map (uncurry (build_rec_arg lthy funs_data has_call) o apsnd (try the_single));
+    val ctr_poss = map (fn x =>
+      if length (distinct ((op =) o pairself (length o #left_args)) x) <> 1 then
+        primrec_error ("inconstant constructor pattern position for function " ^
+          quote (#fun_name (hd x)))
+      else
+        hd x |> #left_args |> length) funs_data;
+  in
+    (recs, ctr_poss)
+    |-> map2 (fn recx => fn ctr_pos => list_comb (recx, rec_args) |> permute_args ctr_pos)
+    |> Syntax.check_terms lthy
+    |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.conceal (Thm.def_binding b), []), t)))
+      bs mxs
+  end;
+
+fun find_rec_calls has_call ({ctr, ctr_args, rhs_term, ...} : eqn_data) =
+  let
+    fun find bound_Ts (Abs (_, T, b)) ctr_arg = find (T :: bound_Ts) b ctr_arg
+      | find bound_Ts (t as _ $ _) ctr_arg =
+        let
+          val typof = curry fastype_of1 bound_Ts;
+          val (f', args') = strip_comb t;
+          val n = find_index (equal ctr_arg o head_of) args';
+        in
+          if n < 0 then
+            find bound_Ts f' ctr_arg @ maps (fn x => find bound_Ts x ctr_arg) args'
+          else
+            let
+              val (f, args as arg :: _) = chop n args' |>> curry list_comb f'
+              val (arg_head, arg_args) = Term.strip_comb arg;
+            in
+              if has_call f then
+                mk_partial_compN (length arg_args) (typof arg_head) f ::
+                maps (fn x => find bound_Ts x ctr_arg) args
+              else
+                find bound_Ts f ctr_arg @ maps (fn x => find bound_Ts x ctr_arg) args
+            end
+        end
+      | find _ _ _ = [];
+  in
+    map (find [] rhs_term) ctr_args
+    |> (fn [] => NONE | callss => SOME (ctr, callss))
+  end;
+
+fun mk_primrec_tac ctxt num_extra_args map_idents map_comps fun_defs recx =
+  unfold_thms_tac ctxt fun_defs THEN
+  HEADGOAL (rtac (funpow num_extra_args (fn thm => thm RS fun_cong) recx RS trans)) THEN
+  unfold_thms_tac ctxt (@{thms id_def split o_def fst_conv snd_conv} @ map_comps @ map_idents) THEN
+  HEADGOAL (rtac refl);
+
+fun prepare_primrec fixes specs lthy =
+  let
+    val thy = Proof_Context.theory_of lthy;
+
+    val (bs, mxs) = map_split (apfst fst) fixes;
+    val fun_names = map Binding.name_of bs;
+    val eqns_data = map (dissect_eqn lthy fun_names) specs;
+    val funs_data = eqns_data
+      |> partition_eq ((op =) o pairself #fun_name)
+      |> finds (fn (x, y) => x = #fun_name (hd y)) fun_names |> fst
+      |> map (fn (x, y) => the_single y handle List.Empty =>
+          primrec_error ("missing equations for function " ^ quote x));
+
+    val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
+    val arg_Ts = map (#rec_type o hd) funs_data;
+    val res_Ts = map (#res_type o hd) funs_data;
+    val callssss = funs_data
+      |> map (partition_eq ((op =) o pairself #ctr))
+      |> map (maps (map_filter (find_rec_calls has_call)));
+
+    val _ = (case filter_out (fn (_, T) => Sign.of_sort thy (T, HOLogic.typeS)) (bs ~~ res_Ts) of
+        [] => ()
+      | (b, _) :: _ => primrec_error ("type of " ^ Binding.print b ^ " contains top sort"));
+
+    val ((n2m, rec_specs, _, induct_thm, induct_thms), lthy') =
+      rec_specs_of bs arg_Ts res_Ts (get_indices fixes) callssss lthy;
+
+    val actual_nn = length funs_data;
+
+    val _ = let val ctrs = (maps (map #ctr o #ctr_specs) rec_specs) in
+      map (fn {ctr, user_eqn, ...} => member (op =) ctrs ctr orelse
+        primrec_error_eqn ("argument " ^ quote (Syntax.string_of_term lthy' ctr) ^
+          " is not a constructor in left-hand side") user_eqn) eqns_data end;
+
+    val defs = build_defs lthy' bs mxs funs_data rec_specs has_call;
+
+    fun prove lthy def_thms' ({ctr_specs, nested_map_idents, nested_map_comps, ...} : rec_spec)
+        (fun_data : eqn_data list) =
+      let
+        val def_thms = map (snd o snd) def_thms';
+        val simp_thmss = finds (fn (x, y) => #ctr x = #ctr y) fun_data ctr_specs
+          |> fst
+          |> map_filter (try (fn (x, [y]) =>
+            (#user_eqn x, length (#left_args x) + length (#right_args x), #rec_thm y)))
+          |> map (fn (user_eqn, num_extra_args, rec_thm) =>
+            mk_primrec_tac lthy num_extra_args nested_map_idents nested_map_comps def_thms rec_thm
+            |> K |> Goal.prove_sorry lthy [] [] user_eqn
+            |> Thm.close_derivation);
+        val poss = find_indices (op = o pairself #ctr) fun_data eqns_data;
+      in
+        (poss, simp_thmss)
+      end;
+
+    val notes =
+      (if n2m then map2 (fn name => fn thm =>
+        (name, inductN, [thm], [])) fun_names (take actual_nn induct_thms) else [])
+      |> map (fn (prefix, thmN, thms, attrs) =>
+        ((Binding.qualify true prefix (Binding.name thmN), attrs), [(thms, [])]));
+
+    val common_name = mk_common_name fun_names;
+
+    val common_notes =
+      (if n2m then [(inductN, [induct_thm], [])] else [])
+      |> map (fn (thmN, thms, attrs) =>
+        ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
+  in
+    (((fun_names, defs),
+      fn lthy => fn defs =>
+        split_list (map2 (prove lthy defs) (take actual_nn rec_specs) funs_data)),
+      lthy' |> Local_Theory.notes (notes @ common_notes) |> snd)
+  end;
+
+(* primrec definition *)
+
+fun add_primrec_simple fixes ts lthy =
+  let
+    val (((names, defs), prove), lthy) = prepare_primrec fixes ts lthy
+      handle ERROR str => primrec_error str;
+  in
+    lthy
+    |> fold_map Local_Theory.define defs
+    |-> (fn defs => `(fn lthy => (names, (map fst defs, prove lthy defs))))
+  end
+  handle Primrec_Error (str, eqns) =>
+    if null eqns
+    then error ("primrec_new error:\n  " ^ str)
+    else error ("primrec_new error:\n  " ^ str ^ "\nin\n  " ^
+      space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns));
+
+local
+
+fun gen_primrec prep_spec (raw_fixes : (binding * 'a option * mixfix) list) raw_spec lthy =
+  let
+    val d = duplicates (op =) (map (Binding.name_of o #1) raw_fixes)
+    val _ = null d orelse primrec_error ("duplicate function name(s): " ^ commas d);
+
+    val (fixes, specs) = fst (prep_spec raw_fixes raw_spec lthy);
+
+    val mk_notes =
+      flat ooo map3 (fn poss => fn prefix => fn thms =>
+        let
+          val (bs, attrss) = map_split (fst o nth specs) poss;
+          val notes =
+            map3 (fn b => fn attrs => fn thm =>
+              ((Binding.qualify false prefix b, code_nitpicksimp_simp_attrs @ attrs), [([thm], [])]))
+            bs attrss thms;
+        in
+          ((Binding.qualify true prefix (Binding.name simpsN), []), [(thms, [])]) :: notes
+        end);
+  in
+    lthy
+    |> add_primrec_simple fixes (map snd specs)
+    |-> (fn (names, (ts, (posss, simpss))) =>
+      Spec_Rules.add Spec_Rules.Equational (ts, flat simpss)
+      #> Local_Theory.notes (mk_notes posss names simpss)
+      #>> pair ts o map snd)
+  end;
+
+in
+
+val add_primrec = gen_primrec Specification.check_spec;
+val add_primrec_cmd = gen_primrec Specification.read_spec;
+
+end;
+
+fun add_primrec_global fixes specs thy =
+  let
+    val lthy = Named_Target.theory_init thy;
+    val ((ts, simps), lthy') = add_primrec fixes specs lthy;
+    val simps' = burrow (Proof_Context.export lthy' lthy) simps;
+  in ((ts, simps'), Local_Theory.exit_global lthy') end;
+
+fun add_primrec_overloaded ops fixes specs thy =
+  let
+    val lthy = Overloading.overloading ops thy;
+    val ((ts, simps), lthy') = add_primrec fixes specs lthy;
+    val simps' = burrow (Proof_Context.export lthy' lthy) simps;
+  in ((ts, simps'), Local_Theory.exit_global lthy') end;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_lfp_tactics.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,817 @@
+(*  Title:      HOL/BNF/Tools/bnf_lfp_tactics.ML
+    Author:     Dmitriy Traytel, TU Muenchen
+    Author:     Andrei Popescu, TU Muenchen
+    Copyright   2012
+
+Tactics for the datatype construction.
+*)
+
+signature BNF_LFP_TACTICS =
+sig
+  val mk_alg_min_alg_tac: int -> thm -> thm list -> thm -> thm -> thm list list -> thm list ->
+    thm list -> tactic
+  val mk_alg_not_empty_tac: Proof.context -> thm -> thm list -> thm list -> tactic
+  val mk_alg_select_tac: thm -> {prems: 'a, context: Proof.context} -> tactic
+  val mk_alg_set_tac: thm -> tactic
+  val mk_bd_card_order_tac: thm list -> tactic
+  val mk_bd_limit_tac: int -> thm -> tactic
+  val mk_card_of_min_alg_tac: thm -> thm -> thm -> thm -> thm -> tactic
+  val mk_copy_alg_tac: thm list list -> thm list -> thm -> thm -> thm -> tactic
+  val mk_copy_str_tac: thm list list -> thm -> thm list -> tactic
+  val mk_ctor_induct_tac: Proof.context -> int -> thm list list -> thm -> thm list -> thm ->
+    thm list -> thm list -> thm list -> tactic
+  val mk_ctor_induct2_tac: ctyp option list -> cterm option list -> thm -> thm list ->
+    {prems: 'a, context: Proof.context} -> tactic
+  val mk_ctor_set_tac: thm -> thm -> thm list -> tactic
+  val mk_ctor_rel_tac: Proof.context -> thm list -> int -> thm -> thm -> thm -> thm -> thm list ->
+    thm -> thm -> thm list -> thm list -> thm list list -> tactic
+  val mk_dtor_o_ctor_tac: thm -> thm -> thm -> thm -> thm list -> tactic
+  val mk_ex_copy_alg_tac: int -> thm -> thm -> tactic
+  val mk_init_ex_mor_tac: thm -> thm -> thm -> thm list -> thm -> thm -> thm ->
+    {prems: 'a, context: Proof.context} -> tactic
+  val mk_init_induct_tac: int -> thm -> thm -> thm list -> thm list -> tactic
+  val mk_init_unique_mor_tac: int -> thm -> thm -> thm list -> thm list -> thm list -> thm list ->
+    thm list -> tactic
+  val mk_iso_alt_tac: thm list -> thm -> tactic
+  val mk_fold_unique_mor_tac: thm list -> thm list -> thm list -> thm -> thm -> thm -> tactic
+  val mk_fold_transfer_tac: int -> thm -> thm list -> thm list ->
+    {prems: thm list, context: Proof.context} -> tactic
+  val mk_least_min_alg_tac: thm -> thm -> tactic
+  val mk_le_rel_OO_tac: int -> thm -> thm list -> thm list -> thm list -> thm list ->
+    {prems: 'a, context: Proof.context} -> tactic
+  val mk_map_comp0_tac: thm list -> thm list -> thm -> int -> tactic
+  val mk_map_id0_tac: thm list -> thm -> tactic
+  val mk_map_tac: int -> int -> thm -> thm -> thm -> tactic
+  val mk_ctor_map_unique_tac: thm -> thm list -> Proof.context -> tactic
+  val mk_mcong_tac: (int -> tactic) -> thm list list list -> thm list -> thm list ->
+    {prems: 'a, context: Proof.context} -> tactic
+  val mk_min_algs_card_of_tac: ctyp -> cterm -> int -> thm -> thm list -> thm list -> thm -> thm ->
+    thm -> thm -> thm -> thm -> thm -> tactic
+  val mk_min_algs_least_tac: ctyp -> cterm -> thm -> thm list -> thm list -> tactic
+  val mk_min_algs_mono_tac: Proof.context -> thm -> tactic
+  val mk_min_algs_tac: thm -> thm list -> tactic
+  val mk_mor_Abs_tac: thm -> thm list -> thm list -> tactic
+  val mk_mor_Rep_tac: thm list -> thm -> thm list -> thm list -> thm list ->
+    {prems: 'a, context: Proof.context} -> tactic
+  val mk_mor_UNIV_tac: int -> thm list -> thm -> tactic
+  val mk_mor_comp_tac: thm -> thm list list -> thm list -> tactic
+  val mk_mor_convol_tac: 'a list -> thm -> tactic
+  val mk_mor_elim_tac: thm -> tactic
+  val mk_mor_incl_tac: thm -> thm list -> tactic
+  val mk_mor_inv_tac: thm -> thm -> thm list list -> thm list -> thm list -> thm list -> tactic
+  val mk_mor_fold_tac: ctyp -> cterm -> thm list -> thm -> thm -> tactic
+  val mk_mor_select_tac: thm -> thm -> thm -> thm -> thm -> thm -> thm list -> thm list list ->
+    thm list -> tactic
+  val mk_mor_str_tac: 'a list -> thm -> tactic
+  val mk_rel_induct_tac: int -> thm -> int list -> thm list -> thm list ->
+    {prems: thm list, context: Proof.context} -> tactic
+  val mk_rec_tac: thm list -> thm -> thm list -> {prems: 'a, context: Proof.context} -> tactic
+  val mk_rec_unique_mor_tac: thm list -> thm list -> thm -> {prems: 'a, context: Proof.context} ->
+    tactic
+  val mk_set_bd_tac: int -> (int -> tactic) -> thm -> thm list list -> thm list -> int ->
+    Proof.context -> tactic
+  val mk_set_nat_tac: int -> (int -> tactic) -> thm list list -> thm list -> cterm list ->
+    thm list -> int -> {prems: 'a, context: Proof.context} -> tactic
+  val mk_set_map0_tac: thm -> tactic
+  val mk_set_tac: thm -> tactic
+  val mk_wit_tac: Proof.context -> int -> thm list -> thm list -> tactic
+end;
+
+structure BNF_LFP_Tactics : BNF_LFP_TACTICS =
+struct
+
+open BNF_Tactics
+open BNF_LFP_Util
+open BNF_Util
+
+val fst_snd_convs = @{thms fst_conv snd_conv};
+val ord_eq_le_trans = @{thm ord_eq_le_trans};
+val subset_trans = @{thm subset_trans};
+val trans_fun_cong_image_id_id_apply = @{thm trans[OF fun_cong[OF image_id] id_apply]};
+val rev_bspec = Drule.rotate_prems 1 bspec;
+val Un_cong = @{thm arg_cong2[of _ _ _ _ "op \<union>"]}
+
+fun mk_alg_set_tac alg_def =
+  dtac (alg_def RS iffD1) 1 THEN
+  REPEAT_DETERM (etac conjE 1) THEN
+  EVERY' [etac bspec, rtac CollectI] 1 THEN
+  REPEAT_DETERM (etac conjI 1) THEN atac 1;
+
+fun mk_alg_not_empty_tac ctxt alg_set alg_sets wits =
+  (EVERY' [rtac notI, hyp_subst_tac ctxt, ftac alg_set] THEN'
+  REPEAT_DETERM o FIRST'
+    [rtac subset_UNIV,
+    EVERY' [rtac @{thm subset_emptyI}, eresolve_tac wits],
+    EVERY' [rtac subsetI, rtac FalseE, eresolve_tac wits],
+    EVERY' [rtac subsetI, dresolve_tac wits, hyp_subst_tac ctxt,
+      FIRST' (map (fn thm => rtac thm THEN' atac) alg_sets)]] THEN'
+  etac @{thm emptyE}) 1;
+
+fun mk_mor_elim_tac mor_def =
+  (dtac (subst OF [mor_def]) THEN'
+  REPEAT o etac conjE THEN'
+  TRY o rtac @{thm image_subsetI} THEN'
+  etac bspec THEN'
+  atac) 1;
+
+fun mk_mor_incl_tac mor_def map_ids =
+  (stac mor_def THEN'
+  rtac conjI THEN'
+  CONJ_WRAP' (K (EVERY' [rtac ballI, etac set_mp, stac id_apply, atac])) map_ids THEN'
+  CONJ_WRAP' (fn thm =>
+   (EVERY' [rtac ballI, rtac trans, rtac id_apply, stac thm, rtac refl])) map_ids) 1;
+
+fun mk_mor_comp_tac mor_def set_maps map_comp_ids =
+  let
+    val fbetw_tac = EVERY' [rtac ballI, stac o_apply, etac bspec, etac bspec, atac];
+    fun mor_tac (set_map, map_comp_id) =
+      EVERY' [rtac ballI, stac o_apply, rtac trans,
+        rtac trans, dtac rev_bspec, atac, etac arg_cong,
+         REPEAT o eresolve_tac [CollectE, conjE], etac bspec, rtac CollectI] THEN'
+      CONJ_WRAP' (fn thm =>
+        FIRST' [rtac subset_UNIV,
+          (EVERY' [rtac ord_eq_le_trans, rtac thm, rtac @{thm image_subsetI},
+            etac bspec, etac set_mp, atac])]) set_map THEN'
+      rtac (map_comp_id RS arg_cong);
+  in
+    (dtac (mor_def RS subst) THEN' dtac (mor_def RS subst) THEN' stac mor_def THEN'
+    REPEAT o etac conjE THEN'
+    rtac conjI THEN'
+    CONJ_WRAP' (K fbetw_tac) set_maps THEN'
+    CONJ_WRAP' mor_tac (set_maps ~~ map_comp_ids)) 1
+  end;
+
+fun mk_mor_inv_tac alg_def mor_def set_maps morEs map_comp_ids map_cong0Ls =
+  let
+    val fbetw_tac = EVERY' [rtac ballI, etac set_mp, etac imageI];
+    fun Collect_tac set_map =
+      CONJ_WRAP' (fn thm =>
+        FIRST' [rtac subset_UNIV,
+          (EVERY' [rtac ord_eq_le_trans, rtac thm, rtac subset_trans,
+            etac @{thm image_mono}, atac])]) set_map;
+    fun mor_tac (set_map, ((morE, map_comp_id), map_cong0L)) =
+      EVERY' [rtac ballI, ftac rev_bspec, atac,
+         REPEAT o eresolve_tac [CollectE, conjE], rtac sym, rtac trans, rtac sym,
+         etac @{thm inverE}, etac bspec, rtac CollectI, Collect_tac set_map,
+         rtac trans, etac (morE RS arg_cong), rtac CollectI, Collect_tac set_map,
+         rtac trans, rtac (map_comp_id RS arg_cong), rtac (map_cong0L RS arg_cong),
+         REPEAT_DETERM_N (length morEs) o
+           (EVERY' [rtac subst, rtac @{thm inver_pointfree}, etac @{thm inver_mono}, atac])];
+  in
+    (stac mor_def THEN'
+    dtac (alg_def RS iffD1) THEN'
+    dtac (alg_def RS iffD1) THEN'
+    REPEAT o etac conjE THEN'
+    rtac conjI THEN'
+    CONJ_WRAP' (K fbetw_tac) set_maps THEN'
+    CONJ_WRAP' mor_tac (set_maps ~~ (morEs ~~ map_comp_ids ~~ map_cong0Ls))) 1
+  end;
+
+fun mk_mor_str_tac ks mor_def =
+  (stac mor_def THEN' rtac conjI THEN'
+  CONJ_WRAP' (K (EVERY' [rtac ballI, rtac UNIV_I])) ks THEN'
+  CONJ_WRAP' (K (EVERY' [rtac ballI, rtac refl])) ks) 1;
+
+fun mk_mor_convol_tac ks mor_def =
+  (stac mor_def THEN' rtac conjI THEN'
+  CONJ_WRAP' (K (EVERY' [rtac ballI, rtac UNIV_I])) ks THEN'
+  CONJ_WRAP' (K (EVERY' [rtac ballI, rtac trans, rtac @{thm fst_convol'}, rtac o_apply])) ks) 1;
+
+fun mk_mor_UNIV_tac m morEs mor_def =
+  let
+    val n = length morEs;
+    fun mor_tac morE = EVERY' [rtac ext, rtac trans, rtac o_apply, rtac trans, etac morE,
+      rtac CollectI, CONJ_WRAP' (K (rtac subset_UNIV)) (1 upto m + n),
+      rtac sym, rtac o_apply];
+  in
+    EVERY' [rtac iffI, CONJ_WRAP' mor_tac morEs,
+    stac mor_def, rtac conjI, CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) morEs,
+    REPEAT_DETERM o etac conjE, REPEAT_DETERM_N n o dtac (@{thm fun_eq_iff} RS subst),
+    CONJ_WRAP' (K (EVERY' [rtac ballI, REPEAT_DETERM o etac allE, rtac trans,
+      etac (o_apply RS subst), rtac o_apply])) morEs] 1
+  end;
+
+fun mk_iso_alt_tac mor_images mor_inv =
+  let
+    val n = length mor_images;
+    fun if_wrap_tac thm =
+      EVERY' [rtac ssubst, rtac @{thm bij_betw_iff_ex}, rtac exI, rtac conjI,
+        rtac @{thm inver_surj}, etac thm, etac thm, atac, etac conjI, atac]
+    val if_tac =
+      EVERY' [etac thin_rl, etac thin_rl, REPEAT o eresolve_tac [conjE, exE],
+        rtac conjI, atac, CONJ_WRAP' if_wrap_tac mor_images];
+    val only_if_tac =
+      EVERY' [rtac conjI, etac conjunct1, EVERY' (map (fn thm =>
+        EVERY' [rtac exE, rtac @{thm bij_betw_ex_weakE}, etac (conjunct2 RS thm)])
+        (map (mk_conjunctN n) (1 upto n))), REPEAT o rtac exI, rtac conjI, rtac mor_inv,
+        etac conjunct1, atac, atac, REPEAT_DETERM_N n o atac,
+        CONJ_WRAP' (K (etac conjunct2)) mor_images];
+  in
+    (rtac iffI THEN' if_tac THEN' only_if_tac) 1
+  end;
+
+fun mk_copy_str_tac set_maps alg_def alg_sets =
+  let
+    val n = length alg_sets;
+    val bij_betw_inv_tac =
+      EVERY' [etac thin_rl, REPEAT_DETERM_N n o EVERY' [dtac @{thm bij_betwI}, atac, atac],
+        REPEAT_DETERM_N (2 * n) o etac thin_rl, REPEAT_DETERM_N (n - 1) o etac conjI, atac];
+    fun set_tac thms =
+      EVERY' [rtac ord_eq_le_trans, resolve_tac thms, rtac subset_trans,
+          etac @{thm image_mono}, rtac equalityD1, etac @{thm bij_betw_imageE}];
+    val copy_str_tac =
+      CONJ_WRAP' (fn (thms, thm) =>
+        EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac set_mp,
+          rtac equalityD1, etac @{thm bij_betw_imageE}, rtac imageI, etac thm,
+          REPEAT_DETERM o rtac subset_UNIV, REPEAT_DETERM_N n o (set_tac thms)])
+      (set_maps ~~ alg_sets);
+  in
+    (rtac rev_mp THEN' DETERM o bij_betw_inv_tac THEN' rtac impI THEN'
+    stac alg_def THEN' copy_str_tac) 1
+  end;
+
+fun mk_copy_alg_tac set_maps alg_sets mor_def iso_alt copy_str =
+  let
+    val n = length alg_sets;
+    val fbetw_tac = CONJ_WRAP' (K (etac @{thm bij_betwE})) alg_sets;
+    fun set_tac thms =
+      EVERY' [rtac ord_eq_le_trans, resolve_tac thms, rtac subset_trans,
+        REPEAT_DETERM o etac conjE, etac @{thm image_mono},
+        rtac equalityD1, etac @{thm bij_betw_imageE}];
+    val mor_tac =
+      CONJ_WRAP' (fn (thms, thm) =>
+        EVERY' [rtac ballI, etac CollectE, etac @{thm inverE}, etac thm,
+          REPEAT_DETERM o rtac subset_UNIV, REPEAT_DETERM_N n o (set_tac thms)])
+      (set_maps ~~ alg_sets);
+  in
+    (rtac (iso_alt RS iffD2) THEN'
+    etac copy_str THEN' REPEAT_DETERM o atac THEN'
+    rtac conjI THEN' stac mor_def THEN' rtac conjI THEN' fbetw_tac THEN' mor_tac THEN'
+    CONJ_WRAP' (K atac) alg_sets) 1
+  end;
+
+fun mk_ex_copy_alg_tac n copy_str copy_alg =
+  EVERY' [REPEAT_DETERM_N n o rtac exI, rtac conjI, etac copy_str,
+    REPEAT_DETERM_N n o atac,
+    REPEAT_DETERM_N n o etac @{thm bij_betw_inver2},
+    REPEAT_DETERM_N n o etac @{thm bij_betw_inver1}, etac copy_alg,
+    REPEAT_DETERM_N n o atac,
+    REPEAT_DETERM_N n o etac @{thm bij_betw_inver2},
+    REPEAT_DETERM_N n o etac @{thm bij_betw_inver1}] 1;
+
+fun mk_bd_limit_tac n bd_Cinfinite =
+  EVERY' [REPEAT_DETERM o etac conjE, rtac rev_mp, rtac @{thm Cinfinite_limit_finite},
+    REPEAT_DETERM_N n o rtac @{thm finite.insertI}, rtac @{thm finite.emptyI},
+    REPEAT_DETERM_N n o etac @{thm insert_subsetI}, rtac @{thm empty_subsetI},
+    rtac bd_Cinfinite, rtac impI, etac bexE, rtac bexI,
+    CONJ_WRAP' (fn i =>
+      EVERY' [etac bspec, REPEAT_DETERM_N i o rtac @{thm insertI2}, rtac @{thm insertI1}])
+      (0 upto n - 1),
+    atac] 1;
+
+fun mk_min_algs_tac worel in_congs =
+  let
+    val minG_tac = EVERY' [rtac @{thm UN_cong}, rtac refl, dtac bspec, atac, etac arg_cong];
+    fun minH_tac thm =
+      EVERY' [rtac Un_cong, minG_tac, rtac @{thm image_cong}, rtac thm,
+        REPEAT_DETERM_N (length in_congs) o minG_tac, rtac refl];
+  in
+    (rtac (worel RS (@{thm wo_rel.worec_fixpoint} RS fun_cong)) THEN' rtac ssubst THEN'
+    rtac meta_eq_to_obj_eq THEN' rtac (worel RS @{thm wo_rel.adm_wo_def}) THEN'
+    REPEAT_DETERM_N 3 o rtac allI THEN' rtac impI THEN'
+    CONJ_WRAP_GEN' (EVERY' [rtac Pair_eqI, rtac conjI]) minH_tac in_congs) 1
+  end;
+
+fun mk_min_algs_mono_tac ctxt min_algs = EVERY' [stac @{thm relChain_def}, rtac allI, rtac allI,
+  rtac impI, rtac @{thm case_split}, rtac @{thm xt1(3)}, rtac min_algs, etac @{thm FieldI2},
+  rtac subsetI, rtac UnI1, rtac @{thm UN_I}, etac @{thm underS_I}, atac, atac,
+  rtac equalityD1, dtac @{thm notnotD}, hyp_subst_tac ctxt, rtac refl] 1;
+
+fun mk_min_algs_card_of_tac cT ct m worel min_algs in_bds bd_Card_order bd_Cnotzero
+  suc_Card_order suc_Cinfinite suc_Cnotzero suc_Asuc Asuc_Cinfinite =
+  let
+    val induct = worel RS
+      Drule.instantiate' [SOME cT] [NONE, SOME ct] @{thm well_order_induct_imp};
+    val src = 1 upto m + 1;
+    val dest = (m + 1) :: (1 upto m);
+    val absorbAs_tac = if m = 0 then K (all_tac)
+      else EVERY' [rtac @{thm ordIso_transitive}, rtac @{thm csum_cong1},
+        rtac @{thm ordIso_transitive},
+        BNF_Tactics.mk_rotate_eq_tac (rtac @{thm ordIso_refl} THEN'
+          FIRST' [rtac @{thm card_of_Card_order}, rtac @{thm Card_order_csum},
+            rtac @{thm Card_order_cexp}])
+        @{thm ordIso_transitive} @{thm csum_assoc} @{thm csum_com} @{thm csum_cong}
+        src dest,
+        rtac @{thm csum_absorb1}, rtac Asuc_Cinfinite, rtac ctrans, rtac @{thm ordLeq_csum1},
+        FIRST' [rtac @{thm Card_order_csum}, rtac @{thm card_of_Card_order}],
+        rtac @{thm ordLeq_cexp1}, rtac suc_Cnotzero, rtac @{thm Card_order_csum}];
+
+    val minG_tac = EVERY' [rtac @{thm UNION_Cinfinite_bound}, rtac @{thm ordLess_imp_ordLeq},
+      rtac @{thm ordLess_transitive}, rtac @{thm card_of_underS}, rtac suc_Card_order,
+      atac, rtac suc_Asuc, rtac ballI, etac allE, dtac mp, etac @{thm underS_E},
+      dtac mp, etac @{thm underS_Field}, REPEAT o etac conjE, atac, rtac Asuc_Cinfinite]
+
+    fun mk_minH_tac (min_alg, in_bd) = EVERY' [rtac @{thm ordIso_ordLeq_trans},
+      rtac @{thm card_of_ordIso_subst}, etac min_alg, rtac @{thm Un_Cinfinite_bound},
+      minG_tac, rtac ctrans, rtac @{thm card_of_image}, rtac ctrans, rtac in_bd, rtac ctrans,
+      rtac @{thm cexp_mono1}, rtac @{thm csum_mono1},
+      REPEAT_DETERM_N m o rtac @{thm csum_mono2},
+      CONJ_WRAP_GEN' (rtac @{thm csum_cinfinite_bound}) (K minG_tac) min_algs,
+      REPEAT_DETERM o FIRST'
+        [rtac @{thm card_of_Card_order}, rtac @{thm Card_order_csum},
+        rtac Asuc_Cinfinite, rtac bd_Card_order],
+      rtac @{thm ordIso_ordLeq_trans}, rtac @{thm cexp_cong1}, absorbAs_tac,
+      rtac @{thm csum_absorb1}, rtac Asuc_Cinfinite, rtac @{thm ctwo_ordLeq_Cinfinite},
+      rtac Asuc_Cinfinite, rtac bd_Card_order,
+      rtac @{thm ordIso_imp_ordLeq}, rtac @{thm cexp_cprod_ordLeq},
+      resolve_tac @{thms Card_order_csum Card_order_ctwo}, rtac suc_Cinfinite,
+      rtac bd_Cnotzero, rtac @{thm cardSuc_ordLeq}, rtac bd_Card_order, rtac Asuc_Cinfinite];
+  in
+    (rtac induct THEN'
+    rtac impI THEN'
+    CONJ_WRAP' mk_minH_tac (min_algs ~~ in_bds)) 1
+  end;
+
+fun mk_min_algs_least_tac cT ct worel min_algs alg_sets =
+  let
+    val induct = worel RS
+      Drule.instantiate' [SOME cT] [NONE, SOME ct] @{thm well_order_induct_imp};
+
+    val minG_tac = EVERY' [rtac @{thm UN_least}, etac allE, dtac mp, etac @{thm underS_E},
+      dtac mp, etac @{thm underS_Field}, REPEAT_DETERM o etac conjE, atac];
+
+    fun mk_minH_tac (min_alg, alg_set) = EVERY' [rtac ord_eq_le_trans, etac min_alg,
+      rtac @{thm Un_least}, minG_tac, rtac @{thm image_subsetI},
+      REPEAT_DETERM o eresolve_tac [CollectE, conjE], etac alg_set,
+      REPEAT_DETERM o FIRST' [atac, etac subset_trans THEN' minG_tac]];
+  in
+    (rtac induct THEN'
+    rtac impI THEN'
+    CONJ_WRAP' mk_minH_tac (min_algs ~~ alg_sets)) 1
+  end;
+
+fun mk_alg_min_alg_tac m alg_def min_alg_defs bd_limit bd_Cinfinite
+    set_bdss min_algs min_alg_monos =
+  let
+    val n = length min_algs;
+    fun mk_cardSuc_UNION_tac set_bds (mono, def) = EVERY'
+      [rtac bexE, rtac @{thm cardSuc_UNION_Cinfinite}, rtac bd_Cinfinite, rtac mono,
+       etac (def RSN (2, @{thm subset_trans[OF _ equalityD1]})), resolve_tac set_bds];
+    fun mk_conjunct_tac (set_bds, (min_alg, min_alg_def)) =
+      EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE],
+        EVERY' (map (mk_cardSuc_UNION_tac set_bds) (min_alg_monos ~~ min_alg_defs)), rtac bexE,
+        rtac bd_limit, REPEAT_DETERM_N (n - 1) o etac conjI, atac,
+        rtac (min_alg_def RS @{thm set_mp[OF equalityD2]}),
+        rtac @{thm UN_I}, REPEAT_DETERM_N (m + 3 * n) o etac thin_rl, atac, rtac set_mp,
+        rtac equalityD2, rtac min_alg, atac, rtac UnI2, rtac @{thm image_eqI}, rtac refl,
+        rtac CollectI, REPEAT_DETERM_N m o dtac asm_rl, REPEAT_DETERM_N n o etac thin_rl,
+        REPEAT_DETERM o etac conjE,
+        CONJ_WRAP' (K (FIRST' [atac,
+          EVERY' [etac subset_trans, rtac subsetI, rtac @{thm UN_I},
+            etac @{thm underS_I}, atac, atac]]))
+          set_bds];
+  in
+    (rtac (alg_def RS iffD2) THEN'
+    CONJ_WRAP' mk_conjunct_tac (set_bdss ~~ (min_algs ~~ min_alg_defs))) 1
+  end;
+
+fun mk_card_of_min_alg_tac min_alg_def card_of suc_Card_order suc_Asuc Asuc_Cinfinite =
+  EVERY' [stac min_alg_def, rtac @{thm UNION_Cinfinite_bound},
+    rtac @{thm ordIso_ordLeq_trans}, rtac @{thm card_of_Field_ordIso}, rtac suc_Card_order,
+    rtac @{thm ordLess_imp_ordLeq}, rtac suc_Asuc, rtac ballI, dtac rev_mp, rtac card_of,
+    REPEAT_DETERM o etac conjE, atac, rtac Asuc_Cinfinite] 1;
+
+fun mk_least_min_alg_tac min_alg_def least =
+  EVERY' [stac min_alg_def, rtac @{thm UN_least}, dtac least, dtac mp, atac,
+    REPEAT_DETERM o etac conjE, atac] 1;
+
+fun mk_alg_select_tac Abs_inverse {context = ctxt, prems = _} =
+  EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac ctxt] 1 THEN
+  unfold_thms_tac ctxt (Abs_inverse :: fst_snd_convs) THEN atac 1;
+
+fun mk_mor_select_tac mor_def mor_cong mor_comp mor_incl_min_alg alg_def alg_select alg_sets
+    set_maps str_init_defs =
+  let
+    val n = length alg_sets;
+    val fbetw_tac =
+      CONJ_WRAP' (K (EVERY' [rtac ballI, etac rev_bspec, etac CollectE, atac])) alg_sets;
+    val mor_tac =
+      CONJ_WRAP' (fn thm => EVERY' [rtac ballI, rtac thm]) str_init_defs;
+    fun alg_epi_tac ((alg_set, str_init_def), set_map) =
+      EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac CollectI,
+        rtac ballI, ftac (alg_select RS bspec), stac str_init_def, etac alg_set,
+        REPEAT_DETERM o FIRST' [rtac subset_UNIV,
+          EVERY' [rtac ord_eq_le_trans, resolve_tac set_map, rtac subset_trans,
+            etac @{thm image_mono}, rtac @{thm image_Collect_subsetI}, etac bspec, atac]]];
+  in
+    (rtac mor_cong THEN' REPEAT_DETERM_N n o (rtac sym THEN' rtac @{thm o_id}) THEN'
+    rtac (Thm.permute_prems 0 1 mor_comp) THEN' etac (Thm.permute_prems 0 1 mor_comp) THEN'
+    stac mor_def THEN' rtac conjI THEN' fbetw_tac THEN' mor_tac THEN' rtac mor_incl_min_alg THEN'
+    stac alg_def THEN' CONJ_WRAP' alg_epi_tac ((alg_sets ~~ str_init_defs) ~~ set_maps)) 1
+  end;
+
+fun mk_init_ex_mor_tac Abs_inverse copy_alg_ex alg_min_alg card_of_min_algs
+    mor_comp mor_select mor_incl_min_alg {context = ctxt, prems = _} =
+  let
+    val n = length card_of_min_algs;
+    val card_of_ordIso_tac = EVERY' [rtac ssubst, rtac @{thm card_of_ordIso},
+      rtac @{thm ordIso_symmetric}, rtac conjunct1, rtac conjunct2, atac];
+    fun internalize_tac card_of = EVERY' [rtac subst, rtac @{thm internalize_card_of_ordLeq2},
+      rtac @{thm ordLeq_ordIso_trans}, rtac card_of, rtac subst,
+      rtac @{thm Card_order_iff_ordIso_card_of}, rtac @{thm Card_order_cexp}];
+  in
+    (rtac rev_mp THEN'
+    REPEAT_DETERM_N (2 * n) o (rtac mp THEN' rtac @{thm ex_mono} THEN' rtac impI) THEN'
+    REPEAT_DETERM_N (n + 1) o etac thin_rl THEN' rtac (alg_min_alg RS copy_alg_ex) THEN'
+    REPEAT_DETERM_N n o atac THEN'
+    REPEAT_DETERM_N n o card_of_ordIso_tac THEN'
+    EVERY' (map internalize_tac card_of_min_algs) THEN'
+    rtac impI THEN'
+    REPEAT_DETERM o eresolve_tac [exE, conjE] THEN'
+    REPEAT_DETERM o rtac exI THEN'
+    rtac mor_select THEN' atac THEN' rtac CollectI THEN'
+    REPEAT_DETERM o rtac exI THEN'
+    rtac conjI THEN' rtac refl THEN' atac THEN'
+    K (unfold_thms_tac ctxt (Abs_inverse :: fst_snd_convs)) THEN'
+    etac mor_comp THEN' etac mor_incl_min_alg) 1
+  end;
+
+fun mk_init_unique_mor_tac m
+    alg_def alg_min_alg least_min_algs in_monos alg_sets morEs map_cong0s =
+  let
+    val n = length least_min_algs;
+    val ks = (1 upto n);
+
+    fun mor_tac morE in_mono = EVERY' [etac morE, rtac set_mp, rtac in_mono,
+      REPEAT_DETERM_N n o rtac @{thm Collect_restrict}, rtac CollectI,
+      REPEAT_DETERM_N (m + n) o (TRY o rtac conjI THEN' atac)];
+    fun cong_tac map_cong0 = EVERY' [rtac (map_cong0 RS arg_cong),
+      REPEAT_DETERM_N m o rtac refl,
+      REPEAT_DETERM_N n o (etac @{thm prop_restrict} THEN' atac)];
+
+    fun mk_alg_tac (alg_set, (in_mono, (morE, map_cong0))) = EVERY' [rtac ballI, rtac CollectI,
+      REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac conjI, rtac (alg_min_alg RS alg_set),
+      REPEAT_DETERM_N m o rtac subset_UNIV,
+      REPEAT_DETERM_N n o (etac subset_trans THEN' rtac @{thm Collect_restrict}),
+      rtac trans, mor_tac morE in_mono,
+      rtac trans, cong_tac map_cong0,
+      rtac sym, mor_tac morE in_mono];
+
+    fun mk_unique_tac (k, least_min_alg) =
+      select_prem_tac n (etac @{thm prop_restrict}) k THEN' rtac least_min_alg THEN'
+      stac alg_def THEN'
+      CONJ_WRAP' mk_alg_tac (alg_sets ~~ (in_monos ~~ (morEs ~~ map_cong0s)));
+  in
+    CONJ_WRAP' mk_unique_tac (ks ~~ least_min_algs) 1
+  end;
+
+fun mk_init_induct_tac m alg_def alg_min_alg least_min_algs alg_sets =
+  let
+    val n = length least_min_algs;
+
+    fun mk_alg_tac alg_set = EVERY' [rtac ballI, rtac CollectI,
+      REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac conjI, rtac (alg_min_alg RS alg_set),
+      REPEAT_DETERM_N m o rtac subset_UNIV,
+      REPEAT_DETERM_N n o (etac subset_trans THEN' rtac @{thm Collect_restrict}),
+      rtac mp, etac bspec, rtac CollectI,
+      REPEAT_DETERM_N m o (rtac conjI THEN' atac),
+      CONJ_WRAP' (K (etac subset_trans THEN' rtac @{thm Collect_restrict})) alg_sets,
+      CONJ_WRAP' (K (rtac ballI THEN' etac @{thm prop_restrict} THEN' atac)) alg_sets];
+
+    fun mk_induct_tac least_min_alg =
+      rtac ballI THEN' etac @{thm prop_restrict} THEN' rtac least_min_alg THEN'
+      stac alg_def THEN'
+      CONJ_WRAP' mk_alg_tac alg_sets;
+  in
+    CONJ_WRAP' mk_induct_tac least_min_algs 1
+  end;
+
+fun mk_mor_Rep_tac ctor_defs copy bijs inver_Abss inver_Reps {context = ctxt, prems = _} =
+  (K (unfold_thms_tac ctxt ctor_defs) THEN' rtac conjunct1 THEN' rtac copy THEN'
+  EVERY' (map (fn bij => EVERY' [rtac bij, atac, etac bexI, rtac UNIV_I]) bijs) THEN'
+  EVERY' (map rtac inver_Abss) THEN'
+  EVERY' (map rtac inver_Reps)) 1;
+
+fun mk_mor_Abs_tac inv inver_Abss inver_Reps =
+  (rtac inv THEN'
+  EVERY' (map2 (fn inver_Abs => fn inver_Rep =>
+    EVERY' [rtac conjI, rtac subset_UNIV, rtac conjI, rtac inver_Rep, rtac inver_Abs])
+    inver_Abss inver_Reps)) 1;
+
+fun mk_mor_fold_tac cT ct fold_defs ex_mor mor =
+  (EVERY' (map stac fold_defs) THEN' EVERY' [rtac rev_mp, rtac ex_mor, rtac impI] THEN'
+  REPEAT_DETERM_N (length fold_defs) o etac exE THEN'
+  rtac (Drule.instantiate' [SOME cT] [SOME ct] @{thm someI}) THEN' etac mor) 1;
+
+fun mk_fold_unique_mor_tac type_defs init_unique_mors Reps mor_comp mor_Abs mor_fold =
+  let
+    fun mk_unique type_def =
+      EVERY' [rtac @{thm surj_fun_eq}, rtac (type_def RS @{thm type_definition.Abs_image}),
+        rtac ballI, resolve_tac init_unique_mors,
+        EVERY' (map (fn thm => atac ORELSE' rtac thm) Reps),
+        rtac mor_comp, rtac mor_Abs, atac,
+        rtac mor_comp, rtac mor_Abs, rtac mor_fold];
+  in
+    CONJ_WRAP' mk_unique type_defs 1
+  end;
+
+fun mk_dtor_o_ctor_tac dtor_def foldx map_comp_id map_cong0L ctor_o_folds =
+  EVERY' [stac dtor_def, rtac ext, rtac trans, rtac o_apply, rtac trans, rtac foldx,
+    rtac trans, rtac map_comp_id, rtac trans, rtac map_cong0L,
+    EVERY' (map (fn thm => rtac ballI THEN' rtac (trans OF [thm RS fun_cong, id_apply]))
+      ctor_o_folds),
+    rtac sym, rtac id_apply] 1;
+
+fun mk_rec_tac rec_defs foldx fst_recs {context = ctxt, prems = _}=
+  unfold_thms_tac ctxt
+    (rec_defs @ map (fn thm => thm RS @{thm convol_expand_snd}) fst_recs) THEN
+  EVERY' [rtac trans, rtac o_apply, rtac trans, rtac (foldx RS @{thm arg_cong[of _ _ snd]}),
+    rtac @{thm snd_convol'}] 1;
+
+fun mk_rec_unique_mor_tac rec_defs fst_recs fold_unique_mor {context = ctxt, prems = _} =
+  unfold_thms_tac ctxt
+    (rec_defs @ map (fn thm => thm RS @{thm convol_expand_snd'}) fst_recs) THEN
+  etac fold_unique_mor 1;
+
+fun mk_ctor_induct_tac ctxt m set_mapss init_induct morEs mor_Abs Rep_invs Abs_invs Reps =
+  let
+    val n = length set_mapss;
+    val ks = 1 upto n;
+
+    fun mk_IH_tac Rep_inv Abs_inv set_map =
+      DETERM o EVERY' [dtac meta_mp, rtac (Rep_inv RS arg_cong RS subst), etac bspec,
+        dtac set_rev_mp, rtac equalityD1, rtac set_map, etac imageE,
+        hyp_subst_tac ctxt, rtac (Abs_inv RS ssubst), etac set_mp, atac, atac];
+
+    fun mk_closed_tac (k, (morE, set_maps)) =
+      EVERY' [select_prem_tac n (dtac asm_rl) k, rtac ballI, rtac impI,
+        rtac (mor_Abs RS morE RS arg_cong RS ssubst), atac,
+        REPEAT_DETERM o eresolve_tac [CollectE, conjE], dtac @{thm meta_spec},
+        EVERY' (map3 mk_IH_tac Rep_invs Abs_invs (drop m set_maps)), atac];
+
+    fun mk_induct_tac (Rep, Rep_inv) =
+      EVERY' [rtac (Rep_inv RS arg_cong RS subst), etac (Rep RSN (2, bspec))];
+  in
+    (rtac mp THEN' rtac impI THEN'
+    DETERM o CONJ_WRAP_GEN' (etac conjE THEN' rtac conjI) mk_induct_tac (Reps ~~ Rep_invs) THEN'
+    rtac init_induct THEN'
+    DETERM o CONJ_WRAP' mk_closed_tac (ks ~~ (morEs ~~ set_mapss))) 1
+  end;
+
+fun mk_ctor_induct2_tac cTs cts ctor_induct weak_ctor_inducts {context = ctxt, prems = _} =
+  let
+    val n = length weak_ctor_inducts;
+    val ks = 1 upto n;
+    fun mk_inner_induct_tac induct i =
+      EVERY' [rtac allI, fo_rtac induct ctxt,
+        select_prem_tac n (dtac @{thm meta_spec2}) i,
+        REPEAT_DETERM_N n o
+          EVERY' [dtac meta_mp THEN_ALL_NEW Goal.norm_hhf_tac ctxt,
+            REPEAT_DETERM o dtac @{thm meta_spec}, etac (spec RS meta_mp), atac],
+        atac];
+  in
+    EVERY' [rtac rev_mp, rtac (Drule.instantiate' cTs cts ctor_induct),
+      EVERY' (map2 mk_inner_induct_tac weak_ctor_inducts ks), rtac impI,
+      REPEAT_DETERM o eresolve_tac [conjE, allE],
+      CONJ_WRAP' (K atac) ks] 1
+  end;
+
+fun mk_map_tac m n foldx map_comp_id map_cong0 =
+  EVERY' [rtac ext, rtac trans, rtac o_apply, rtac trans, rtac foldx, rtac trans, rtac o_apply,
+    rtac trans, rtac (map_comp_id RS arg_cong), rtac trans, rtac (map_cong0 RS arg_cong),
+    REPEAT_DETERM_N m o rtac refl,
+    REPEAT_DETERM_N n o (EVERY' (map rtac [trans, o_apply, id_apply])),
+    rtac sym, rtac o_apply] 1;
+
+fun mk_ctor_map_unique_tac fold_unique sym_map_comps ctxt =
+  rtac fold_unique 1 THEN
+  unfold_thms_tac ctxt (sym_map_comps @ @{thms o_assoc[symmetric] id_o o_id}) THEN
+  ALLGOALS atac;
+
+fun mk_set_tac foldx = EVERY' [rtac ext, rtac trans, rtac o_apply,
+  rtac trans, rtac foldx, rtac sym, rtac o_apply] 1;
+
+fun mk_ctor_set_tac set set_map set_maps =
+  let
+    val n = length set_maps;
+    fun mk_UN thm = rtac (thm RS @{thm arg_cong[of _ _ Union]} RS trans) THEN'
+      rtac @{thm Union_image_eq};
+  in
+    EVERY' [rtac (set RS @{thm comp_eq_dest} RS trans), rtac Un_cong,
+      rtac (trans OF [set_map, trans_fun_cong_image_id_id_apply]),
+      REPEAT_DETERM_N (n - 1) o rtac Un_cong,
+      EVERY' (map mk_UN set_maps)] 1
+  end;
+
+fun mk_set_nat_tac m induct_tac set_mapss
+    ctor_maps csets ctor_sets i {context = ctxt, prems = _} =
+  let
+    val n = length ctor_maps;
+
+    fun useIH set_nat = EVERY' [rtac trans, rtac @{thm image_UN}, rtac trans, rtac @{thm UN_cong},
+      rtac refl, Goal.assume_rule_tac ctxt, rtac sym, rtac trans, rtac @{thm UN_cong},
+      rtac set_nat, rtac refl, rtac @{thm UN_simps(10)}];
+
+    fun mk_set_nat cset ctor_map ctor_set set_nats =
+      EVERY' [rtac trans, rtac @{thm image_cong}, rtac ctor_set, rtac refl,
+        rtac sym, rtac (trans OF [ctor_map RS HOL_arg_cong cset, ctor_set RS trans]),
+        rtac sym, EVERY' (map rtac [trans, @{thm image_Un}, Un_cong]),
+        rtac sym, rtac (nth set_nats (i - 1)),
+        REPEAT_DETERM_N (n - 1) o EVERY' (map rtac [trans, @{thm image_Un}, Un_cong]),
+        EVERY' (map useIH (drop m set_nats))];
+  in
+    (induct_tac THEN' EVERY' (map4 mk_set_nat csets ctor_maps ctor_sets set_mapss)) 1
+  end;
+
+fun mk_set_bd_tac m induct_tac bd_Cinfinite set_bdss ctor_sets i ctxt =
+  let
+    val n = length ctor_sets;
+
+    fun useIH set_bd = EVERY' [rtac @{thm UNION_Cinfinite_bound}, rtac set_bd, rtac ballI,
+      Goal.assume_rule_tac ctxt, rtac bd_Cinfinite];
+
+    fun mk_set_nat ctor_set set_bds =
+      EVERY' [rtac @{thm ordIso_ordLeq_trans}, rtac @{thm card_of_ordIso_subst}, rtac ctor_set,
+        rtac (bd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})), rtac (nth set_bds (i - 1)),
+        REPEAT_DETERM_N (n - 1) o rtac (bd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})),
+        EVERY' (map useIH (drop m set_bds))];
+  in
+    (induct_tac THEN' EVERY' (map2 mk_set_nat ctor_sets set_bdss)) 1
+  end;
+
+fun mk_mcong_tac induct_tac set_setsss map_cong0s ctor_maps {context = ctxt, prems = _} =
+  let
+    fun use_asm thm = EVERY' [etac bspec, etac set_rev_mp, rtac thm];
+
+    fun useIH set_sets = EVERY' [rtac mp, Goal.assume_rule_tac ctxt,
+      CONJ_WRAP' (fn thm =>
+        EVERY' [rtac ballI, etac bspec, etac set_rev_mp, etac thm]) set_sets];
+
+    fun mk_map_cong0 ctor_map map_cong0 set_setss =
+      EVERY' [rtac impI, REPEAT_DETERM o etac conjE,
+        rtac trans, rtac ctor_map, rtac trans, rtac (map_cong0 RS arg_cong),
+        EVERY' (map use_asm (map hd set_setss)),
+        EVERY' (map useIH (transpose (map tl set_setss))),
+        rtac sym, rtac ctor_map];
+  in
+    (induct_tac THEN' EVERY' (map3 mk_map_cong0 ctor_maps map_cong0s set_setsss)) 1
+  end;
+
+fun mk_le_rel_OO_tac m induct ctor_nchotomys ctor_Irels rel_mono_strongs rel_OOs 
+    {context = ctxt, prems = _} =
+  EVERY' (rtac induct ::
+  map4 (fn nchotomy => fn Irel => fn rel_mono => fn rel_OO =>
+    EVERY' [rtac impI, etac (nchotomy RS @{thm nchotomy_relcomppE}),
+      REPEAT_DETERM_N 2 o dtac (Irel RS iffD1), rtac (Irel RS iffD2),
+      rtac rel_mono, rtac (rel_OO RS @{thm predicate2_eqD} RS iffD2),
+      rtac @{thm relcomppI}, atac, atac,
+      REPEAT_DETERM_N m o EVERY' [rtac ballI, rtac ballI, rtac impI, atac],
+      REPEAT_DETERM_N (length rel_OOs) o
+        EVERY' [rtac ballI, rtac ballI, Goal.assume_rule_tac ctxt]])
+  ctor_nchotomys ctor_Irels rel_mono_strongs rel_OOs) 1;
+
+(* BNF tactics *)
+
+fun mk_map_id0_tac map_id0s unique =
+  (rtac sym THEN' rtac unique THEN'
+  EVERY' (map (fn thm =>
+    EVERY' [rtac trans, rtac @{thm id_o}, rtac trans, rtac sym, rtac @{thm o_id},
+      rtac (thm RS sym RS arg_cong)]) map_id0s)) 1;
+
+fun mk_map_comp0_tac map_comp0s ctor_maps unique iplus1 =
+  let
+    val i = iplus1 - 1;
+    val unique' = Thm.permute_prems 0 i unique;
+    val map_comp0s' = drop i map_comp0s @ take i map_comp0s;
+    val ctor_maps' = drop i ctor_maps @ take i ctor_maps;
+    fun mk_comp comp simp =
+      EVERY' [rtac ext, rtac trans, rtac o_apply, rtac trans, rtac o_apply,
+        rtac trans, rtac (simp RS arg_cong), rtac trans, rtac simp,
+        rtac trans, rtac (comp RS arg_cong), rtac sym, rtac o_apply];
+  in
+    (rtac sym THEN' rtac unique' THEN' EVERY' (map2 mk_comp map_comp0s' ctor_maps')) 1
+  end;
+
+fun mk_set_map0_tac set_nat =
+  EVERY' (map rtac [ext, trans, o_apply, sym, trans, o_apply, set_nat]) 1;
+
+fun mk_bd_card_order_tac bd_card_orders =
+  CONJ_WRAP_GEN' (rtac @{thm card_order_csum}) rtac bd_card_orders 1;
+
+fun mk_wit_tac ctxt n ctor_set wit =
+  REPEAT_DETERM (atac 1 ORELSE
+    EVERY' [dtac set_rev_mp, rtac equalityD1, resolve_tac ctor_set,
+    REPEAT_DETERM o
+      (TRY o REPEAT_DETERM o etac UnE THEN' TRY o etac @{thm UN_E} THEN'
+        (eresolve_tac wit ORELSE'
+        (dresolve_tac wit THEN'
+          (etac FalseE ORELSE'
+          EVERY' [hyp_subst_tac ctxt, dtac set_rev_mp, rtac equalityD1, resolve_tac ctor_set,
+            REPEAT_DETERM_N n o etac UnE]))))] 1);
+
+fun mk_ctor_rel_tac ctxt in_Irels i in_rel map_comp0 map_cong0 ctor_map ctor_sets ctor_inject
+  ctor_dtor set_map0s ctor_set_incls ctor_set_set_inclss =
+  let
+    val m = length ctor_set_incls;
+    val n = length ctor_set_set_inclss;
+
+    val (passive_set_map0s, active_set_map0s) = chop m set_map0s;
+    val in_Irel = nth in_Irels (i - 1);
+    val le_arg_cong_ctor_dtor = ctor_dtor RS arg_cong RS ord_eq_le_trans;
+    val eq_arg_cong_ctor_dtor = ctor_dtor RS arg_cong RS trans;
+    val if_tac =
+      EVERY' [dtac (in_Irel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
+        rtac (in_rel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
+        EVERY' (map2 (fn set_map0 => fn ctor_set_incl =>
+          EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac set_map0,
+            rtac ord_eq_le_trans, rtac trans_fun_cong_image_id_id_apply,
+            rtac (ctor_set_incl RS subset_trans), etac le_arg_cong_ctor_dtor])
+        passive_set_map0s ctor_set_incls),
+        CONJ_WRAP' (fn (in_Irel, (set_map0, ctor_set_set_incls)) =>
+          EVERY' [rtac ord_eq_le_trans, rtac set_map0, rtac @{thm image_subsetI}, rtac CollectI,
+            rtac @{thm prod_caseI}, rtac (in_Irel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
+            CONJ_WRAP' (fn thm =>
+              EVERY' (map etac [thm RS subset_trans, le_arg_cong_ctor_dtor]))
+            ctor_set_set_incls,
+            rtac conjI, rtac refl, rtac refl])
+        (in_Irels ~~ (active_set_map0s ~~ ctor_set_set_inclss)),
+        CONJ_WRAP' (fn conv =>
+          EVERY' [rtac trans, rtac map_comp0, rtac trans, rtac map_cong0,
+          REPEAT_DETERM_N m o rtac @{thm fun_cong[OF o_id]},
+          REPEAT_DETERM_N n o EVERY' (map rtac [trans, o_apply, conv]),
+          rtac (ctor_inject RS iffD1), rtac trans, rtac sym, rtac ctor_map,
+          etac eq_arg_cong_ctor_dtor])
+        fst_snd_convs];
+    val only_if_tac =
+      EVERY' [dtac (in_rel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
+        rtac (in_Irel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
+        CONJ_WRAP' (fn (ctor_set, passive_set_map0) =>
+          EVERY' [rtac ord_eq_le_trans, rtac ctor_set, rtac @{thm Un_least},
+            rtac ord_eq_le_trans, rtac @{thm box_equals[OF _ refl]},
+            rtac passive_set_map0, rtac trans_fun_cong_image_id_id_apply, atac,
+            CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least}))
+              (fn (active_set_map0, in_Irel) => EVERY' [rtac ord_eq_le_trans,
+                rtac @{thm UN_cong[OF _ refl]}, rtac active_set_map0, rtac @{thm UN_least},
+                dtac set_rev_mp, etac @{thm image_mono}, etac imageE,
+                dtac @{thm ssubst_mem[OF pair_collapse]},
+                REPEAT_DETERM o eresolve_tac (CollectE :: conjE ::
+                  @{thms prod_caseE iffD1[OF Pair_eq, elim_format]}),
+                hyp_subst_tac ctxt,
+                dtac (in_Irel RS iffD1), dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE,
+                TRY o
+                  EVERY' [dtac (Thm.permute_prems 0 1 @{thm ssubst_mem}), atac, hyp_subst_tac ctxt],
+                REPEAT_DETERM o eresolve_tac [CollectE, conjE], atac])
+            (rev (active_set_map0s ~~ in_Irels))])
+        (ctor_sets ~~ passive_set_map0s),
+        rtac conjI,
+        REPEAT_DETERM_N 2 o EVERY' [rtac trans, rtac ctor_map, rtac (ctor_inject RS iffD2),
+          rtac trans, rtac map_comp0, rtac trans, rtac map_cong0,
+          REPEAT_DETERM_N m o rtac @{thm fun_cong[OF o_id]},
+          EVERY' (map (fn in_Irel => EVERY' [rtac trans, rtac o_apply, dtac set_rev_mp, atac,
+            dtac @{thm ssubst_mem[OF pair_collapse]},
+            REPEAT_DETERM o
+              eresolve_tac (CollectE :: conjE :: @{thms prod_caseE iffD1[OF Pair_eq, elim_format]}),
+            hyp_subst_tac ctxt,
+            dtac (in_Irel RS iffD1), dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE, atac])
+          in_Irels),
+          atac]]
+  in
+    EVERY' [rtac iffI, if_tac, only_if_tac] 1
+  end;
+
+fun mk_rel_induct_tac m ctor_induct2 ks ctor_rels rel_mono_strongs {context = ctxt, prems = IHs} =
+  let val n = length ks;
+  in
+    unfold_tac ctxt @{thms le_fun_def le_bool_def all_simps(1,2)[symmetric]} THEN
+    EVERY' [REPEAT_DETERM o rtac allI, rtac ctor_induct2,
+      EVERY' (map3 (fn IH => fn ctor_rel => fn rel_mono_strong =>
+        EVERY' [rtac impI, dtac (ctor_rel RS iffD1), rtac (IH RS @{thm spec2} RS mp),
+          etac rel_mono_strong,
+          REPEAT_DETERM_N m o rtac @{thm ballI[OF ballI[OF imp_refl]]},
+          EVERY' (map (fn j =>
+            EVERY' [select_prem_tac n (dtac asm_rl) j, rtac @{thm ballI[OF ballI]},
+              Goal.assume_rule_tac ctxt]) ks)])
+      IHs ctor_rels rel_mono_strongs)] 1
+  end;
+
+fun mk_fold_transfer_tac m rel_induct map_transfers folds {context = ctxt, prems = _} =
+  let
+    val n = length map_transfers;
+  in
+    unfold_thms_tac ctxt
+      @{thms fun_rel_def_butlast all_conj_distrib[symmetric] imp_conjR[symmetric]} THEN
+    unfold_thms_tac ctxt @{thms fun_rel_iff_leq_vimage2p} THEN
+    HEADGOAL (EVERY'
+      [REPEAT_DETERM o resolve_tac [allI, impI], rtac rel_induct,
+      EVERY' (map (fn map_transfer => EVERY'
+        [REPEAT_DETERM o resolve_tac [allI, impI, @{thm vimage2pI}],
+        SELECT_GOAL (unfold_thms_tac ctxt folds),
+        etac @{thm predicate2D_vimage2p},
+        rtac (funpow (m + n + 1) (fn thm => thm RS @{thm fun_relD}) map_transfer),
+        REPEAT_DETERM_N m o rtac @{thm id_transfer},
+        REPEAT_DETERM_N n o rtac @{thm vimage2p_fun_rel},
+        atac])
+      map_transfers)])
+  end;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_lfp_util.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,63 @@
+(*  Title:      HOL/BNF/Tools/bnf_lfp_util.ML
+    Author:     Dmitriy Traytel, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2012
+
+Library for the datatype construction.
+*)
+
+signature BNF_LFP_UTIL =
+sig
+  val HOL_arg_cong: cterm -> thm
+
+  val mk_bij_betw: term -> term -> term -> term
+  val mk_cardSuc: term -> term
+  val mk_inver: term -> term -> term -> term
+  val mk_not_empty: term -> term
+  val mk_not_eq: term -> term -> term
+  val mk_rapp: term -> typ -> term
+  val mk_relChain: term -> term -> term
+  val mk_underS: term -> term
+  val mk_worec: term -> term -> term
+end;
+
+structure BNF_LFP_Util : BNF_LFP_UTIL =
+struct
+
+open BNF_Util
+
+fun HOL_arg_cong ct = Drule.instantiate'
+  (map SOME (Thm.dest_ctyp (Thm.ctyp_of_term ct))) [NONE, NONE, SOME ct] arg_cong;
+
+(*reverse application*)
+fun mk_rapp arg T = Term.absdummy (fastype_of arg --> T) (Bound 0 $ arg);
+
+fun mk_underS r =
+  let val T = fst (dest_relT (fastype_of r));
+  in Const (@{const_name underS}, mk_relT (T, T) --> T --> HOLogic.mk_setT T) $ r end;
+
+fun mk_worec r f =
+  let val (A, AB) = apfst domain_type (dest_funT (fastype_of f));
+  in Const (@{const_name wo_rel.worec}, mk_relT (A, A) --> (AB --> AB) --> AB) $ r $ f end;
+
+fun mk_relChain r f =
+  let val (A, AB) = `domain_type (fastype_of f);
+  in Const (@{const_name relChain}, mk_relT (A, A) --> AB --> HOLogic.boolT) $ r $ f end;
+
+fun mk_cardSuc r =
+  let val T = fst (dest_relT (fastype_of r));
+  in Const (@{const_name cardSuc}, mk_relT (T, T) --> mk_relT (`I (HOLogic.mk_setT T))) $ r end;
+
+fun mk_bij_betw f A B =
+ Const (@{const_name bij_betw},
+   fastype_of f --> fastype_of A --> fastype_of B --> HOLogic.boolT) $ f $ A $ B;
+
+fun mk_inver f g A =
+ Const (@{const_name inver}, fastype_of f --> fastype_of g --> fastype_of A --> HOLogic.boolT) $
+   f $ g $ A;
+
+fun mk_not_eq x y = HOLogic.mk_not (HOLogic.mk_eq (x, y));
+
+fun mk_not_empty B = mk_not_eq B (HOLogic.mk_set (HOLogic.dest_setT (fastype_of B)) []);
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_tactics.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,109 @@
+(*  Title:      HOL/BNF/Tools/bnf_tactics.ML
+    Author:     Dmitriy Traytel, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2012
+
+General tactics for bounded natural functors.
+*)
+
+signature BNF_TACTICS =
+sig
+  include CTR_SUGAR_GENERAL_TACTICS
+
+  val fo_rtac: thm -> Proof.context -> int -> tactic
+
+  val mk_flatten_assoc_tac: (int -> tactic) -> thm -> thm -> thm -> tactic
+  val mk_rotate_eq_tac: (int -> tactic) -> thm -> thm -> thm -> thm -> ''a list -> ''a list ->
+    int -> tactic
+
+  val mk_pointfree: Proof.context -> thm -> thm
+
+  val mk_Abs_bij_thm: Proof.context -> thm -> thm -> thm
+  val mk_Abs_inj_thm: thm -> thm
+
+  val mk_ctor_or_dtor_rel_tac:
+    thm -> thm list -> thm list -> thm -> {prems: 'a, context: Proof.context} -> tactic
+
+  val mk_map_comp_id_tac: thm -> tactic
+  val mk_map_cong0_tac: int -> thm -> {prems: 'a, context: Proof.context} -> tactic
+  val mk_map_cong0L_tac: int -> thm -> thm -> tactic
+end;
+
+structure BNF_Tactics : BNF_TACTICS =
+struct
+
+open Ctr_Sugar_General_Tactics
+open BNF_Util
+
+(*stolen from Christian Urban's Cookbook*)
+fun fo_rtac thm = Subgoal.FOCUS (fn {concl, ...} =>
+  let
+    val concl_pat = Drule.strip_imp_concl (cprop_of thm)
+    val insts = Thm.first_order_match (concl_pat, concl)
+  in
+    rtac (Drule.instantiate_normalize insts thm) 1
+  end);
+
+(*transforms f (g x) = h (k x) into f o g = h o k using first order matches for f, g, h, and k*)
+fun mk_pointfree ctxt thm = thm
+  |> Thm.prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq
+  |> pairself (dest_comb #> apsnd (dest_comb #> fst) #> HOLogic.mk_comp)
+  |> mk_Trueprop_eq
+  |> (fn goal => Goal.prove_sorry ctxt [] [] goal
+    (K (rtac ext 1 THEN unfold_thms_tac ctxt [o_apply, mk_sym thm] THEN rtac refl 1)))
+  |> Thm.close_derivation;
+
+
+(* Theorems for open typedefs with UNIV as representing set *)
+
+fun mk_Abs_inj_thm inj = inj OF (replicate 2 UNIV_I);
+fun mk_Abs_bij_thm ctxt Abs_inj_thm surj = rule_by_tactic ctxt ((rtac surj THEN' etac exI) 1)
+  (Abs_inj_thm RS @{thm bijI});
+
+
+
+(* General tactic generators *)
+
+(*applies assoc rule to the lhs of an equation as long as possible*)
+fun mk_flatten_assoc_tac refl_tac trans assoc cong = rtac trans 1 THEN
+  REPEAT_DETERM (CHANGED ((FIRST' [rtac trans THEN' rtac assoc, rtac cong THEN' refl_tac]) 1)) THEN
+  refl_tac 1;
+
+(*proves two sides of an equation to be equal assuming both are flattened and rhs can be obtained
+from lhs by the given permutation of monoms*)
+fun mk_rotate_eq_tac refl_tac trans assoc com cong =
+  let
+    fun gen_tac [] [] = K all_tac
+      | gen_tac [x] [y] = if x = y then refl_tac else error "mk_rotate_eq_tac: different lists"
+      | gen_tac (x :: xs) (y :: ys) = if x = y
+        then rtac cong THEN' refl_tac THEN' gen_tac xs ys
+        else rtac trans THEN' rtac com THEN'
+          K (mk_flatten_assoc_tac refl_tac trans assoc cong) THEN'
+          gen_tac (xs @ [x]) (y :: ys)
+      | gen_tac _ _ = error "mk_rotate_eq_tac: different lists";
+  in
+    gen_tac
+  end;
+
+fun mk_ctor_or_dtor_rel_tac srel_def IJrel_defs IJsrel_defs dtor_srel {context = ctxt, prems = _} =
+  unfold_thms_tac ctxt IJrel_defs THEN
+  rtac (unfold_thms ctxt (IJrel_defs @ IJsrel_defs @
+    @{thms Collect_mem_eq mem_Collect_eq fst_conv snd_conv}) dtor_srel RS trans) 1 THEN
+  unfold_thms_tac ctxt (srel_def ::
+    @{thms pair_collapse Collect_mem_eq mem_Collect_eq prod.cases fst_conv snd_conv
+      split_conv}) THEN
+  rtac refl 1;
+
+fun mk_map_comp_id_tac map_comp0 =
+  (rtac trans THEN' rtac map_comp0 THEN' REPEAT_DETERM o stac @{thm o_id} THEN' rtac refl) 1;
+
+fun mk_map_cong0_tac m map_cong0 {context = ctxt, prems = _} =
+  EVERY' [rtac mp, rtac map_cong0,
+    CONJ_WRAP' (K (rtac ballI THEN' Goal.assume_rule_tac ctxt)) (1 upto m)] 1;
+
+fun mk_map_cong0L_tac passive map_cong0 map_id =
+  (rtac trans THEN' rtac map_cong0 THEN' EVERY' (replicate passive (rtac refl))) 1 THEN
+  REPEAT_DETERM (EVERY' [rtac trans, etac bspec, atac, rtac sym, rtac @{thm id_apply}] 1) THEN
+  rtac map_id 1;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/BNF/bnf_util.ML	Mon Jan 20 18:24:56 2014 +0100
@@ -0,0 +1,583 @@
+(*  Title:      HOL/BNF/Tools/bnf_util.ML
+    Author:     Dmitriy Traytel, TU Muenchen
+    Copyright   2012
+
+Library for bounded natural functors.
+*)
+
+signature BNF_UTIL =
+sig
+  include CTR_SUGAR_UTIL
+
+  val map6: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g) ->
+    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list
+  val map7: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h) ->
+    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list
+  val map8: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i) ->
+    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list -> 'i list
+  val map9: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j) ->
+    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
+    'i list -> 'j list
+  val map10: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j -> 'k) ->
+    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
+    'i list -> 'j list -> 'k list
+  val map11: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j -> 'k -> 'l) ->
+    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
+    'i list -> 'j list -> 'k list -> 'l list
+  val map12: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j -> 'k -> 'l -> 'm) ->
+    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
+    'i list -> 'j list -> 'k list -> 'l list -> 'm list
+  val map13: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j -> 'k -> 'l -> 'm -> 'n) ->
+    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
+    'i list -> 'j list -> 'k list -> 'l list -> 'm list -> 'n list
+  val map14:
+    ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j -> 'k -> 'l -> 'm -> 'n -> 'o) ->
+    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
+    'i list -> 'j list -> 'k list -> 'l list -> 'm list -> 'n list -> 'o list
+  val fold_map4: ('a -> 'b -> 'c -> 'd -> 'e -> 'f * 'e) ->
+    'a list -> 'b list -> 'c list -> 'd list -> 'e -> 'f list * 'e
+  val fold_map5: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g * 'f) ->
+    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f -> 'g list * 'f
+  val fold_map6: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h * 'g) ->
+    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g -> 'h list * 'g
+  val fold_map7: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i * 'h) ->
+    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h -> 'i list * 'h
+  val fold_map8: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j * 'i) ->
+    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list -> 'i ->
+    'j list * 'i
+  val fold_map9: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j -> 'k * 'j) ->
+    'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
+    'i list -> 'j -> 'k list * 'j
+  val split_list4: ('a * 'b * 'c * 'd) list -> 'a list * 'b list * 'c list * 'd list
+  val split_list5: ('a * 'b * 'c * 'd * 'e) list -> 'a list * 'b list * 'c list * 'd list * 'e list
+  val find_indices: ('b * 'a -> bool) -> 'a list -> 'b list -> int list
+
+  val mk_TFreess: int list -> Proof.context -> typ list list * Proof.context
+  val mk_Freesss: string -> typ list list list -> Proof.context ->
+    term list list list * Proof.context
+  val mk_Freessss: string -> typ list list list list -> Proof.context ->
+    term list list list list * Proof.context
+  val nonzero_string_of_int: int -> string
+  val retype_free: typ -> term -> term
+
+  val binder_fun_types: typ -> typ list
+  val body_fun_type: typ -> typ
+  val num_binder_types: typ -> int
+  val strip_fun_type: typ -> typ list * typ
+  val strip_typeN: int -> typ -> typ list * typ
+
+  val mk_pred2T: typ -> typ -> typ
+  val mk_relT: typ * typ -> typ
+  val dest_relT: typ -> typ * typ
+  val dest_pred2T: typ -> typ * typ
+  val mk_sumT: typ * typ -> typ
+
+  val ctwo: term
+  val fst_const: typ -> term
+  val snd_const: typ -> term
+  val Id_const: typ -> term
+
+  val mk_Ball: term -> term -> term
+  val mk_Bex: term -> term -> term
+  val mk_Card_order: term -> term
+  val mk_Field: term -> term
+  val mk_Gr: term -> term -> term
+  val mk_Grp: term -> term -> term
+  val mk_UNION: term -> term -> term
+  val mk_Union: typ -> term
+  val mk_card_binop: string -> (typ * typ -> typ) -> term -> term -> term
+  val mk_card_of: term -> term
+  val mk_card_order: term -> term
+  val mk_cexp: term -> term -> term
+  val mk_cinfinite: term -> term
+  val mk_collect: term list -> typ -> term
+  val mk_converse: term -> term
+  val mk_conversep: term -> term
+  val mk_cprod: term -> term -> term
+  val mk_csum: term -> term -> term
+  val mk_dir_image: term -> term -> term
+  val mk_fun_rel: term -> term -> term
+  val mk_image: term -> term
+  val mk_in: term list -> term list -> typ -> term
+  val mk_leq: term -> term -> term
+  val mk_ordLeq: term -> term -> term
+  val mk_rel_comp: term * term -> term
+  val mk_rel_compp: term * term -> term
+
+  (*parameterized terms*)
+  val mk_nthN: int -> term -> int -> term
+
+  (*parameterized thms*)
+  val mk_Un_upper: int -> int -> thm
+  val mk_conjIN: int -> thm
+  val mk_conjunctN: int -> int -> thm
+  val conj_dests: int -> thm -> thm list
+  val mk_nthI: int -> int -> thm
+  val mk_nth_conv: int -> int -> thm
+  val mk_ordLeq_csum: int -> int -> thm -> thm
+  val mk_UnIN: int -> int -> thm
+
+  val Pair_eqD: thm
+  val Pair_eqI: thm
+  val ctrans: thm
+  val id_apply: thm
+  val meta_mp: thm
+  val meta_spec: thm
+  val o_apply: thm
+  val set_mp: thm
+  val set_rev_mp: thm
+  val subset_UNIV: thm
+  val mk_sym: thm -> thm
+  val mk_trans: thm -> thm -> thm
+
+  val is_refl: thm -> bool
+  val is_concl_refl: thm -> bool
+  val no_refl: thm list -> thm list
+  val no_reflexive: thm list -> thm list
+
+  val fold_thms: Proof.context -> thm list -> thm -> thm
+
+  val parse_binding_colon: binding parser
+  val parse_opt_binding_colon: binding parser
+  val parse_type_args_named_constrained: (binding option * (string * string option)) list parser
+  val parse_map_rel_bindings: (binding * binding) parser
+
+  val typedef: binding * (string * sort) list * mixfix -> term ->
+    (binding * binding) option -> tactic -> local_theory -> (string * Typedef.info) * local_theory
+end;
+
+structure BNF_Util : BNF_UTIL =
+struct
+
+open Ctr_Sugar_Util
+
+(* Library proper *)
+
+fun map6 _ [] [] [] [] [] [] = []
+  | map6 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) =
+    f x1 x2 x3 x4 x5 x6 :: map6 f x1s x2s x3s x4s x5s x6s
+  | map6 _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun map7 _ [] [] [] [] [] [] [] = []
+  | map7 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) =
+    f x1 x2 x3 x4 x5 x6 x7 :: map7 f x1s x2s x3s x4s x5s x6s x7s
+  | map7 _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun map8 _ [] [] [] [] [] [] [] [] = []
+  | map8 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) (x8::x8s) =
+    f x1 x2 x3 x4 x5 x6 x7 x8 :: map8 f x1s x2s x3s x4s x5s x6s x7s x8s
+  | map8 _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun map9 _ [] [] [] [] [] [] [] [] [] = []
+  | map9 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) (x8::x8s)
+      (x9::x9s) =
+    f x1 x2 x3 x4 x5 x6 x7 x8 x9 :: map9 f x1s x2s x3s x4s x5s x6s x7s x8s x9s
+  | map9 _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun map10 _ [] [] [] [] [] [] [] [] [] [] = []
+  | map10 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) (x8::x8s)
+      (x9::x9s) (x10::x10s) =
+    f x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 :: map10 f x1s x2s x3s x4s x5s x6s x7s x8s x9s x10s
+  | map10 _ _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun map11 _ [] [] [] [] [] [] [] [] [] [] [] = []
+  | map11 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) (x8::x8s)
+      (x9::x9s) (x10::x10s) (x11::x11s) =
+    f x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 :: map11 f x1s x2s x3s x4s x5s x6s x7s x8s x9s x10s x11s
+  | map11 _ _ _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun map12 _ [] [] [] [] [] [] [] [] [] [] [] [] = []
+  | map12 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) (x8::x8s)
+      (x9::x9s) (x10::x10s) (x11::x11s) (x12::x12s) =
+    f x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 ::
+      map12 f x1s x2s x3s x4s x5s x6s x7s x8s x9s x10s x11s x12s
+  | map12 _ _ _ _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun map13 _ [] [] [] [] [] [] [] [] [] [] [] [] [] = []
+  | map13 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) (x8::x8s)
+      (x9::x9s) (x10::x10s) (x11::x11s) (x12::x12s) (x13::x13s) =
+    f x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13 ::
+      map13 f x1s x2s x3s x4s x5s x6s x7s x8s x9s x10s x11s x12s x13s
+  | map13 _ _ _ _ _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun map14 _ [] [] [] [] [] [] [] [] [] [] [] [] [] [] = []
+  | map14 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) (x8::x8s)
+      (x9::x9s) (x10::x10s) (x11::x11s) (x12::x12s) (x13::x13s) (x14::x14s) =
+    f x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13 x14 ::
+      map14 f x1s x2s x3s x4s x5s x6s x7s x8s x9s x10s x11s x12s x13s x14s
+  | map14 _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun fold_map4 _ [] [] [] [] acc = ([], acc)
+  | fold_map4 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) acc =
+    let
+      val (x, acc') = f x1 x2 x3 x4 acc;
+      val (xs, acc'') = fold_map4 f x1s x2s x3s x4s acc';
+    in (x :: xs, acc'') end
+  | fold_map4 _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun fold_map5 _ [] [] [] [] [] acc = ([], acc)
+  | fold_map5 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) acc =
+    let
+      val (x, acc') = f x1 x2 x3 x4 x5 acc;
+      val (xs, acc'') = fold_map5 f x1s x2s x3s x4s x5s acc';
+    in (x :: xs, acc'') end
+  | fold_map5 _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun fold_map6 _ [] [] [] [] [] [] acc = ([], acc)
+  | fold_map6 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) acc =
+    let
+      val (x, acc') = f x1 x2 x3 x4 x5 x6 acc;
+      val (xs, acc'') = fold_map6 f x1s x2s x3s x4s x5s x6s acc';
+    in (x :: xs, acc'') end
+  | fold_map6 _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun fold_map7 _ [] [] [] [] [] [] [] acc = ([], acc)
+  | fold_map7 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) acc =
+    let
+      val (x, acc') = f x1 x2 x3 x4 x5 x6 x7 acc;
+      val (xs, acc'') = fold_map7 f x1s x2s x3s x4s x5s x6s x7s acc';
+    in (x :: xs, acc'') end
+  | fold_map7 _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun fold_map8 _ [] [] [] [] [] [] [] [] acc = ([], acc)
+  | fold_map8 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) (x8::x8s)
+      acc =
+    let
+      val (x, acc') = f x1 x2 x3 x4 x5 x6 x7 x8 acc;
+      val (xs, acc'') = fold_map8 f x1s x2s x3s x4s x5s x6s x7s x8s acc';
+    in (x :: xs, acc'') end
+  | fold_map8 _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun fold_map9 _ [] [] [] [] [] [] [] [] [] acc = ([], acc)
+  | fold_map9 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) (x8::x8s)
+      (x9::x9s) acc =
+    let
+      val (x, acc') = f x1 x2 x3 x4 x5 x6 x7 x8 x9 acc;
+      val (xs, acc'') = fold_map9 f x1s x2s x3s x4s x5s x6s x7s x8s x9s acc';
+    in (x :: xs, acc'') end
+  | fold_map9 _ _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun split_list4 [] = ([], [], [], [])
+  | split_list4 ((x1, x2, x3, x4) :: xs) =
+    let val (xs1, xs2, xs3, xs4) = split_list4 xs;
+    in (x1 :: xs1, x2 :: xs2, x3 :: xs3, x4 :: xs4) end;
+
+fun split_list5 [] = ([], [], [], [], [])
+  | split_list5 ((x1, x2, x3, x4, x5) :: xs) =
+    let val (xs1, xs2, xs3, xs4, xs5) = split_list5 xs;
+    in (x1 :: xs1, x2 :: xs2, x3 :: xs3, x4 :: xs4, x5 :: xs5) end;
+
+val parse_binding_colon = parse_binding --| @{keyword ":"};
+val parse_opt_binding_colon = Scan.optional parse_binding_colon Binding.empty;
+
+val parse_type_arg_constrained =
+  Parse.type_ident -- Scan.option (@{keyword "::"} |-- Parse.!!! Parse.sort);
+
+val parse_type_arg_named_constrained =
+   (Parse.minus --| @{keyword ":"} >> K NONE || parse_opt_binding_colon >> SOME) --
+   parse_type_arg_constrained;
+
+val parse_type_args_named_constrained =
+  parse_type_arg_constrained >> (single o pair (SOME Binding.empty)) ||
+  @{keyword "("} |-- Parse.!!! (Parse.list1 parse_type_arg_named_constrained --| @{keyword ")"}) ||
+  Scan.succeed [];
+
+val parse_map_rel_binding = Parse.short_ident --| @{keyword ":"} -- parse_binding;
+
+val no_map_rel = (Binding.empty, Binding.empty);
+
+fun extract_map_rel ("map", b) = apfst (K b)
+  | extract_map_rel ("rel", b) = apsnd (K b)
+  | extract_map_rel (s, _) = error ("Unknown label " ^ quote s ^ " (expected \"map\" or \"rel\")");
+
+val parse_map_rel_bindings =
+  @{keyword "("} |-- Scan.repeat parse_map_rel_binding --| @{keyword ")"}
+    >> (fn ps => fold extract_map_rel ps no_map_rel) ||
+  Scan.succeed no_map_rel;
+
+
+(*TODO: is this really different from Typedef.add_typedef_global?*)
+fun typedef (b, Ts, mx) set opt_morphs tac lthy =
+  let
+    (*Work around loss of qualification in "typedef" axioms by replicating it in the name*)
+    val b' = fold_rev Binding.prefix_name (map (suffix "_" o fst) (#2 (Binding.dest b))) b;
+    val ((name, info), (lthy, lthy_old)) =
+      lthy
+      |> Typedef.add_typedef (b', Ts, mx) set opt_morphs tac
+      ||> `Local_Theory.restore;
+    val phi = Proof_Context.export_morphism lthy_old lthy;
+  in
+    ((name, Typedef.transform_info phi info), lthy)
+  end;
+
+
+
+(* Term construction *)
+
+(** Fresh variables **)
+
+fun nonzero_string_of_int 0 = ""
+  | nonzero_string_of_int n = string_of_int n;
+
+val mk_TFreess = fold_map mk_TFrees;
+
+fun mk_Freesss x Tsss = fold_map2 mk_Freess (mk_names (length Tsss) x) Tsss;
+fun mk_Freessss x Tssss = fold_map2 mk_Freesss (mk_names (length Tssss) x) Tssss;
+
+fun retype_free T (Free (s, _)) = Free (s, T)
+  | retype_free _ t = raise TERM ("retype_free", [t]);
+
+
+(** Types **)
+
+(*stolen from ~~/src/HOL/Tools/Nitpick/nitpick_hol.ML*)
+fun num_binder_types (Type (@{type_name fun}, [_, T2])) =
+    1 + num_binder_types T2
+  | num_binder_types _ = 0
+
+(*maps [T1,...,Tn]--->T to ([T1,T2,...,Tn], T)*)
+fun strip_typeN 0 T = ([], T)
+  | strip_typeN n (Type (@{type_name fun}, [T, T'])) = strip_typeN (n - 1) T' |>> cons T
+  | strip_typeN _ T = raise TYPE ("strip_typeN", [T], []);
+
+(*maps [T1,...,Tn]--->T-->U to ([T1,T2,...,Tn], T-->U), where U is not a function type*)
+fun strip_fun_type T = strip_typeN (num_binder_types T - 1) T;
+
+val binder_fun_types = fst o strip_fun_type;
+val body_fun_type = snd o strip_fun_type;
+
+fun mk_pred2T T U = mk_predT [T, U];
+val mk_relT = HOLogic.mk_setT o HOLogic.mk_prodT;
+val dest_relT = HOLogic.dest_prodT o HOLogic.dest_setT;
+val dest_pred2T = apsnd Term.domain_type o Term.dest_funT;
+fun mk_sumT (LT, RT) = Type (@{type_name Sum_Type.sum}, [LT, RT]);
+
+
+(** Constants **)
+
+fun fst_const T = Const (@{const_name fst}, T --> fst (HOLogic.dest_prodT T));
+fun snd_const T = Const (@{const_name snd}, T --> snd (HOLogic.dest_prodT T));
+fun Id_const T = Const (@{const_name Id}, mk_relT (T, T));
+
+
+(** Operators **)
+
+fun mk_converse R =
+  let
+    val RT = dest_relT (fastype_of R);
+    val RST = mk_relT (snd RT, fst RT);
+  in Const (@{const_name converse}, fastype_of R --> RST) $ R end;
+
+fun mk_rel_comp (R, S) =
+  let
+    val RT = fastype_of R;
+    val ST = fastype_of S;
+    val RST = mk_relT (fst (dest_relT RT), snd (dest_relT ST));
+  in Const (@{const_name relcomp}, RT --> ST --> RST) $ R $ S end;
+
+fun mk_Gr A f =
+  let val ((AT, BT), FT) = `dest_funT (fastype_of f);
+  in Const (@{const_name Gr}, HOLogic.mk_setT AT --> FT --> mk_relT (AT, BT)) $ A $ f end;
+
+fun mk_conversep R =
+  let
+    val RT = dest_pred2T (fastype_of R);
+    val RST = mk_pred2T (snd RT) (fst RT);
+  in Const (@{const_name conversep}, fastype_of R --> RST) $ R end;
+
+fun mk_rel_compp (R, S) =
+  let
+    val RT = fastype_of R;
+    val ST = fastype_of S;
+    val RST = mk_pred2T (fst (dest_pred2T RT)) (snd (dest_pred2T ST));
+  in Const (@{const_name relcompp}, RT --> ST --> RST) $ R $ S end;
+
+fun mk_Grp A f =
+  let val ((AT, BT), FT) = `dest_funT (fastype_of f);
+  in Const (@{const_name Grp}, HOLogic.mk_setT AT --> FT --> mk_pred2T AT BT) $ A $ f end;
+
+fun mk_image f =
+  let val (T, U) = dest_funT (fastype_of f);
+  in Const (@{const_name image},
+    (T --> U) --> (HOLogic.mk_setT T) --> (HOLogic.mk_setT U)) $ f end;
+
+fun mk_Ball X f =
+  Const (@{const_name Ball}, fastype_of X --> fastype_of f --> HOLogic.boolT) $ X $ f;
+
+fun mk_Bex X f =
+  Const (@{const_name Bex}, fastype_of X --> fastype_of f --> HOLogic.boolT) $ X $ f;
+
+fun mk_UNION X f =
+  let val (T, U) = dest_funT (fastype_of f);
+  in Const (@{const_name SUPR}, fastype_of X --> (T --> U) --> U) $ X $ f end;
+
+fun mk_Union T =
+  Const (@{const_name Sup}, HOLogic.mk_setT (HOLogic.mk_setT T) --> HOLogic.mk_setT T);
+
+fun mk_Field r =
+  let val T = fst (dest_relT (fastype_of r));
+  in Const (@{const_name Field}, mk_relT (T, T) --> HOLogic.mk_setT T) $ r end;
+
+fun mk_card_order bd =
+  let
+    val T = fastype_of bd;
+    val AT = fst (dest_relT T);
+  in
+    Const (@{const_name card_order_on}, HOLogic.mk_setT AT --> T --> HOLogic.boolT) $
+      (HOLogic.mk_UNIV AT) $ bd
+  end;
+
+fun mk_Card_order bd =
+  let
+    val T = fastype_of bd;
+    val AT = fst (dest_relT T);
+  in
+    Const (@{const_name card_order_on}, HOLogic.mk_setT AT --> T --> HOLogic.boolT) $
+      mk_Field bd $ bd
+  end;
+
+fun mk_cinfinite bd =
+  Const (@{const_name cinfinite}, fastype_of bd --> HOLogic.boolT) $ bd;
+
+fun mk_ordLeq t1 t2 =
+  HOLogic.mk_mem (HOLogic.mk_prod (t1, t2),
+    Const (@{const_name ordLeq}, mk_relT (fastype_of t1, fastype_of t2)));
+
+fun mk_card_of A =
+  let
+    val AT = fastype_of A;
+    val T = HOLogic.dest_setT AT;
+  in
+    Const (@{const_name card_of}, AT --> mk_relT (T, T)) $ A
+  end;
+
+fun mk_dir_image r f =
+  let val (T, U) = dest_funT (fastype_of f);
+  in Const (@{const_name dir_image}, mk_relT (T, T) --> (T --> U) --> mk_relT (U, U)) $ r $ f end;
+
+fun mk_fun_rel R S =
+  let
+    val ((RA, RB), RT) = `dest_pred2T (fastype_of R);
+    val ((SA, SB), ST) = `dest_pred2T (fastype_of S);
+  in Const (@{const_name fun_rel}, RT --> ST --> mk_pred2T (RA --> SA) (RB --> SB)) $ R $ S end;
+
+(*FIXME: "x"?*)
+(*(nth sets i) must be of type "T --> 'ai set"*)
+fun mk_in As sets T =
+  let
+    fun in_single set A =
+      let val AT = fastype_of A;
+      in Const (@{const_name less_eq},
+        AT --> AT --> HOLogic.boolT) $ (set $ Free ("x", T)) $ A end;
+  in
+    if length sets > 0
+    then HOLogic.mk_Collect ("x", T, foldr1 (HOLogic.mk_conj) (map2 in_single sets As))
+    else HOLogic.mk_UNIV T
+  end;
+
+fun mk_leq t1 t2 =
+  Const (@{const_name less_eq}, (fastype_of t1) --> (fastype_of t2) --> HOLogic.boolT) $ t1 $ t2;
+
+fun mk_card_binop binop typop t1 t2 =
+  let
+    val (T1, relT1) = `(fst o dest_relT) (fastype_of t1);
+    val (T2, relT2) = `(fst o dest_relT) (fastype_of t2);
+  in
+    Const (binop, relT1 --> relT2 --> mk_relT (typop (T1, T2), typop (T1, T2))) $ t1 $ t2
+  end;
+
+val mk_csum = mk_card_binop @{const_name csum} mk_sumT;
+val mk_cprod = mk_card_binop @{const_name cprod} HOLogic.mk_prodT;
+val mk_cexp = mk_card_binop @{const_name cexp} (op --> o swap);
+val ctwo = @{term ctwo};
+
+fun mk_collect xs defT =
+  let val T = (case xs of [] => defT | (x::_) => fastype_of x);
+  in Const (@{const_name collect}, HOLogic.mk_setT T --> T) $ (HOLogic.mk_set T xs) end;
+
+fun find_indices eq xs ys = map_filter I
+  (map_index (fn (i, y) => if member eq xs y then SOME i else NONE) ys);
+
+fun mk_trans thm1 thm2 = trans OF [thm1, thm2];
+fun mk_sym thm = thm RS sym;
+
+(*TODO: antiquote heavily used theorems once*)
+val Pair_eqD = @{thm iffD1[OF Pair_eq]};
+val Pair_eqI = @{thm iffD2[OF Pair_eq]};
+val ctrans = @{thm ordLeq_transitive};
+val id_apply = @{thm id_apply};
+val meta_mp = @{thm meta_mp};
+val meta_spec = @{thm meta_spec};
+val o_apply = @{thm o_apply};
+val set_mp = @{thm set_mp};
+val set_rev_mp = @{thm set_rev_mp};
+val subset_UNIV = @{thm subset_UNIV};
+
+fun mk_nthN 1 t 1 = t
+  | mk_nthN _ t 1 = HOLogic.mk_fst t
+  | mk_nthN 2 t 2 = HOLogic.mk_snd t
+  | mk_nthN n t m = mk_nthN (n - 1) (HOLogic.mk_snd t) (m - 1);
+
+fun mk_nth_conv n m =
+  let
+    fun thm b = if b then @{thm fstI} else @{thm sndI}
+    fun mk_nth_conv _ 1 1 = refl
+      | mk_nth_conv _ _ 1 = @{thm fst_conv}
+      | mk_nth_conv _ 2 2 = @{thm snd_conv}
+      | mk_nth_conv b _ 2 = @{thm snd_conv} RS thm b
+      | mk_nth_conv b n m = mk_nth_conv false (n - 1) (m - 1) RS thm b;
+  in mk_nth_conv (not (m = n)) n m end;
+
+fun mk_nthI 1 1 = @{thm TrueE[OF TrueI]}
+  | mk_nthI n m = fold (curry op RS) (replicate (m - 1) @{thm sndI})
+    (if m = n then @{thm TrueE[OF TrueI]} else @{thm fstI});
+
+fun mk_conjunctN 1 1 = @{thm TrueE[OF TrueI]}
+  | mk_conjunctN _ 1 = conjunct1
+  | mk_conjunctN 2 2 = conjunct2
+  | mk_conjunctN n m = conjunct2 RS (mk_conjunctN (n - 1) (m - 1));
+
+fun conj_dests n thm = map (fn k => thm RS mk_conjunctN n k) (1 upto n);
+
+fun mk_conjIN 1 = @{thm TrueE[OF TrueI]}
+  | mk_conjIN n = mk_conjIN (n - 1) RSN (2, conjI);
+
+fun mk_ordLeq_csum 1 1 thm = thm
+  | mk_ordLeq_csum _ 1 thm = @{thm ordLeq_transitive} OF [thm, @{thm ordLeq_csum1}]
+  | mk_ordLeq_csum 2 2 thm = @{thm ordLeq_transitive} OF [thm, @{thm ordLeq_csum2}]
+  | mk_ordLeq_csum n m thm = @{thm ordLeq_transitive} OF
+    [mk_ordLeq_csum (n - 1) (m - 1) thm, @{thm ordLeq_csum2[OF Card_order_csum]}];
+
+local
+  fun mk_Un_upper' 0 = subset_refl
+    | mk_Un_upper' 1 = @{thm Un_upper1}
+    | mk_Un_upper' k = Library.foldr (op RS o swap)
+      (replicate (k - 1) @{thm subset_trans[OF Un_upper1]}, @{thm Un_upper1});
+in
+  fun mk_Un_upper 1 1 = subset_refl
+    | mk_Un_upper n 1 = mk_Un_upper' (n - 2) RS @{thm subset_trans[OF Un_upper1]}
+    | mk_Un_upper n m = mk_Un_upper' (n - m) RS @{thm subset_trans[OF Un_upper2]};
+end;
+
+local
+  fun mk_UnIN' 0 = @{thm UnI2}
+    | mk_UnIN' m = mk_UnIN' (m - 1) RS @{thm UnI1};
+in
+  fun mk_UnIN 1 1 = @{thm TrueE[OF TrueI]}
+    | mk_UnIN n 1 = Library.foldr1 (op RS o swap) (replicate (n - 1) @{thm UnI1})
+    | mk_UnIN n m = mk_UnIN' (n - m)
+end;
+
+fun is_refl_prop t =
+  op aconv (HOLogic.dest_eq (HOLogic.dest_Trueprop t))
+  handle TERM _ => false;
+
+val is_refl = is_refl_prop o Thm.prop_of;
+val is_concl_refl = is_refl_prop o Logic.strip_imp_concl o Thm.prop_of;
+
+val no_refl = filter_out is_refl;
+val no_reflexive = filter_out Thm.is_reflexive;
+
+fun fold_thms ctxt thms = Local_Defs.fold ctxt (distinct Thm.eq_thm_prop thms);
+
+end;