src/HOL/BNF/Tools/bnf_def.ML
author blanchet
Wed, 06 Nov 2013 23:05:44 +0100
changeset 55737 578371ba74cc
parent 55717 3e1d230f1c00
child 55794 632be352a5a3
permissions -rw-r--r--
reverted 3e1d230f1c00 -- pervasiveness is useful, cf. Coinductive_Nat in the AFP
     1 (*  Title:      HOL/BNF/Tools/bnf_def.ML
     2     Author:     Dmitriy Traytel, TU Muenchen
     3     Author:     Jasmin Blanchette, TU Muenchen
     4     Copyright   2012
     5 
     6 Definition of bounded natural functors.
     7 *)
     8 
     9 signature BNF_DEF =
    10 sig
    11   type bnf
    12   type nonemptiness_witness = {I: int list, wit: term, prop: thm list}
    13 
    14   val morph_bnf: morphism -> bnf -> bnf
    15   val eq_bnf: bnf * bnf -> bool
    16   val bnf_of: Proof.context -> string -> bnf option
    17   val register_bnf: string -> (bnf * local_theory) -> (bnf * local_theory)
    18 
    19   val name_of_bnf: bnf -> binding
    20   val T_of_bnf: bnf -> typ
    21   val live_of_bnf: bnf -> int
    22   val lives_of_bnf: bnf -> typ list
    23   val dead_of_bnf: bnf -> int
    24   val deads_of_bnf: bnf -> typ list
    25   val nwits_of_bnf: bnf -> int
    26 
    27   val mapN: string
    28   val relN: string
    29   val setN: string
    30   val mk_setN: int -> string
    31 
    32   val map_of_bnf: bnf -> term
    33   val sets_of_bnf: bnf -> term list
    34   val rel_of_bnf: bnf -> term
    35 
    36   val mk_T_of_bnf: typ list -> typ list -> bnf -> typ
    37   val mk_bd_of_bnf: typ list -> typ list -> bnf -> term
    38   val mk_map_of_bnf: typ list -> typ list -> typ list -> bnf -> term
    39   val mk_rel_of_bnf: typ list -> typ list -> typ list -> bnf -> term
    40   val mk_sets_of_bnf: typ list list -> typ list list -> bnf -> term list
    41   val mk_wits_of_bnf: typ list list -> typ list list -> bnf -> (int list * term) list
    42 
    43   val bd_Card_order_of_bnf: bnf -> thm
    44   val bd_Cinfinite_of_bnf: bnf -> thm
    45   val bd_Cnotzero_of_bnf: bnf -> thm
    46   val bd_card_order_of_bnf: bnf -> thm
    47   val bd_cinfinite_of_bnf: bnf -> thm
    48   val collect_set_map_of_bnf: bnf -> thm
    49   val in_bd_of_bnf: bnf -> thm
    50   val in_cong_of_bnf: bnf -> thm
    51   val in_mono_of_bnf: bnf -> thm
    52   val in_rel_of_bnf: bnf -> thm
    53   val map_comp0_of_bnf: bnf -> thm
    54   val map_comp_of_bnf: bnf -> thm
    55   val map_cong0_of_bnf: bnf -> thm
    56   val map_cong_of_bnf: bnf -> thm
    57   val map_def_of_bnf: bnf -> thm
    58   val map_id0_of_bnf: bnf -> thm
    59   val map_id_of_bnf: bnf -> thm
    60   val map_transfer_of_bnf: bnf -> thm
    61   val map_wppull_of_bnf: bnf -> thm
    62   val map_wpull_of_bnf: bnf -> thm
    63   val rel_def_of_bnf: bnf -> thm
    64   val rel_Grp_of_bnf: bnf -> thm
    65   val rel_OO_of_bnf: bnf -> thm
    66   val rel_OO_Grp_of_bnf: bnf -> thm
    67   val rel_cong_of_bnf: bnf -> thm
    68   val rel_conversep_of_bnf: bnf -> thm
    69   val rel_mono_of_bnf: bnf -> thm
    70   val rel_mono_strong_of_bnf: bnf -> thm
    71   val rel_eq_of_bnf: bnf -> thm
    72   val rel_flip_of_bnf: bnf -> thm
    73   val set_bd_of_bnf: bnf -> thm list
    74   val set_defs_of_bnf: bnf -> thm list
    75   val set_map0_of_bnf: bnf -> thm list
    76   val set_map_of_bnf: bnf -> thm list
    77   val wit_thms_of_bnf: bnf -> thm list
    78   val wit_thmss_of_bnf: bnf -> thm list list
    79 
    80   val mk_map: int -> typ list -> typ list -> term -> term
    81   val mk_rel: int -> typ list -> typ list -> term -> term
    82   val build_map: Proof.context -> (typ * typ -> term) -> typ * typ -> term
    83   val build_rel: Proof.context -> (typ * typ -> term) -> typ * typ -> term
    84   val flatten_type_args_of_bnf: bnf -> 'a -> 'a list -> 'a list
    85   val map_flattened_map_args: Proof.context -> string -> (term list -> 'a list) -> term list ->
    86     'a list
    87 
    88   val mk_witness: int list * term -> thm list -> nonemptiness_witness
    89   val minimize_wits: (''a list * 'b) list -> (''a list * 'b) list
    90   val wits_of_bnf: bnf -> nonemptiness_witness list
    91 
    92   val zip_axioms: 'a -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list
    93 
    94   datatype const_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline
    95   datatype fact_policy = Dont_Note | Note_Some | Note_All
    96 
    97   val bnf_note_all: bool Config.T
    98   val bnf_timing: bool Config.T
    99   val user_policy: fact_policy -> Proof.context -> fact_policy
   100   val note_bnf_thms: fact_policy -> (binding -> binding) -> binding -> bnf -> Proof.context ->
   101     Proof.context
   102 
   103   val print_bnfs: Proof.context -> unit
   104   val bnf_def: const_policy -> (Proof.context -> fact_policy) -> (binding -> binding) ->
   105     ({prems: thm list, context: Proof.context} -> tactic) list ->
   106     ({prems: thm list, context: Proof.context} -> tactic) -> typ list option -> binding ->
   107     binding -> binding list ->
   108     ((((binding * term) * term list) * term) * term list) * term option ->
   109     local_theory -> bnf * local_theory
   110 end;
   111 
   112 structure BNF_Def : BNF_DEF =
   113 struct
   114 
   115 open BNF_Util
   116 open BNF_Tactics
   117 open BNF_Def_Tactics
   118 
   119 val fundef_cong_attrs = @{attributes [fundef_cong]};
   120 
   121 type axioms = {
   122   map_id0: thm,
   123   map_comp0: thm,
   124   map_cong0: thm,
   125   set_map0: thm list,
   126   bd_card_order: thm,
   127   bd_cinfinite: thm,
   128   set_bd: thm list,
   129   map_wpull: thm,
   130   rel_OO_Grp: thm
   131 };
   132 
   133 fun mk_axioms' ((((((((id, comp), cong), map), c_o), cinf), set_bd), wpull), rel) =
   134   {map_id0 = id, map_comp0 = comp, map_cong0 = cong, set_map0 = map, bd_card_order = c_o,
   135    bd_cinfinite = cinf, set_bd = set_bd, map_wpull = wpull, rel_OO_Grp = rel};
   136 
   137 fun dest_cons [] = raise List.Empty
   138   | dest_cons (x :: xs) = (x, xs);
   139 
   140 fun mk_axioms n thms = thms
   141   |> map the_single
   142   |> dest_cons
   143   ||>> dest_cons
   144   ||>> dest_cons
   145   ||>> chop n
   146   ||>> dest_cons
   147   ||>> dest_cons
   148   ||>> chop n
   149   ||>> dest_cons
   150   ||> the_single
   151   |> mk_axioms';
   152 
   153 fun zip_axioms mid mcomp mcong smap bdco bdinf sbd wpull rel =
   154   [mid, mcomp, mcong] @ smap @ [bdco, bdinf] @ sbd @ [wpull, rel];
   155 
   156 fun dest_axioms {map_id0, map_comp0, map_cong0, set_map0, bd_card_order, bd_cinfinite, set_bd,
   157   map_wpull, rel_OO_Grp} =
   158   zip_axioms map_id0 map_comp0 map_cong0 set_map0 bd_card_order bd_cinfinite set_bd map_wpull
   159     rel_OO_Grp;
   160 
   161 fun map_axioms f {map_id0, map_comp0, map_cong0, set_map0, bd_card_order, bd_cinfinite, set_bd,
   162   map_wpull, rel_OO_Grp} =
   163   {map_id0 = f map_id0,
   164     map_comp0 = f map_comp0,
   165     map_cong0 = f map_cong0,
   166     set_map0 = map f set_map0,
   167     bd_card_order = f bd_card_order,
   168     bd_cinfinite = f bd_cinfinite,
   169     set_bd = map f set_bd,
   170     map_wpull = f map_wpull,
   171     rel_OO_Grp = f rel_OO_Grp};
   172 
   173 val morph_axioms = map_axioms o Morphism.thm;
   174 
   175 type defs = {
   176   map_def: thm,
   177   set_defs: thm list,
   178   rel_def: thm
   179 }
   180 
   181 fun mk_defs map sets rel = {map_def = map, set_defs = sets, rel_def = rel};
   182 
   183 fun map_defs f {map_def, set_defs, rel_def} =
   184   {map_def = f map_def, set_defs = map f set_defs, rel_def = f rel_def};
   185 
   186 val morph_defs = map_defs o Morphism.thm;
   187 
   188 type facts = {
   189   bd_Card_order: thm,
   190   bd_Cinfinite: thm,
   191   bd_Cnotzero: thm,
   192   collect_set_map: thm lazy,
   193   in_bd: thm lazy,
   194   in_cong: thm lazy,
   195   in_mono: thm lazy,
   196   in_rel: thm lazy,
   197   map_comp: thm lazy,
   198   map_cong: thm lazy,
   199   map_id: thm lazy,
   200   map_transfer: thm lazy,
   201   map_wppull: thm lazy,
   202   rel_eq: thm lazy,
   203   rel_flip: thm lazy,
   204   set_map: thm lazy list,
   205   rel_cong: thm lazy,
   206   rel_mono: thm lazy,
   207   rel_mono_strong: thm lazy,
   208   rel_Grp: thm lazy,
   209   rel_conversep: thm lazy,
   210   rel_OO: thm lazy
   211 };
   212 
   213 fun mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_map in_bd in_cong in_mono in_rel
   214     map_comp map_cong map_id map_transfer map_wppull rel_eq rel_flip set_map rel_cong rel_mono
   215     rel_mono_strong rel_Grp rel_conversep rel_OO = {
   216   bd_Card_order = bd_Card_order,
   217   bd_Cinfinite = bd_Cinfinite,
   218   bd_Cnotzero = bd_Cnotzero,
   219   collect_set_map = collect_set_map,
   220   in_bd = in_bd,
   221   in_cong = in_cong,
   222   in_mono = in_mono,
   223   in_rel = in_rel,
   224   map_comp = map_comp,
   225   map_cong = map_cong,
   226   map_id = map_id,
   227   map_transfer = map_transfer,
   228   map_wppull = map_wppull,
   229   rel_eq = rel_eq,
   230   rel_flip = rel_flip,
   231   set_map = set_map,
   232   rel_cong = rel_cong,
   233   rel_mono = rel_mono,
   234   rel_mono_strong = rel_mono_strong,
   235   rel_Grp = rel_Grp,
   236   rel_conversep = rel_conversep,
   237   rel_OO = rel_OO};
   238 
   239 fun map_facts f {
   240   bd_Card_order,
   241   bd_Cinfinite,
   242   bd_Cnotzero,
   243   collect_set_map,
   244   in_bd,
   245   in_cong,
   246   in_mono,
   247   in_rel,
   248   map_comp,
   249   map_cong,
   250   map_id,
   251   map_transfer,
   252   map_wppull,
   253   rel_eq,
   254   rel_flip,
   255   set_map,
   256   rel_cong,
   257   rel_mono,
   258   rel_mono_strong,
   259   rel_Grp,
   260   rel_conversep,
   261   rel_OO} =
   262   {bd_Card_order = f bd_Card_order,
   263     bd_Cinfinite = f bd_Cinfinite,
   264     bd_Cnotzero = f bd_Cnotzero,
   265     collect_set_map = Lazy.map f collect_set_map,
   266     in_bd = Lazy.map f in_bd,
   267     in_cong = Lazy.map f in_cong,
   268     in_mono = Lazy.map f in_mono,
   269     in_rel = Lazy.map f in_rel,
   270     map_comp = Lazy.map f map_comp,
   271     map_cong = Lazy.map f map_cong,
   272     map_id = Lazy.map f map_id,
   273     map_transfer = Lazy.map f map_transfer,
   274     map_wppull = Lazy.map f map_wppull,
   275     rel_eq = Lazy.map f rel_eq,
   276     rel_flip = Lazy.map f rel_flip,
   277     set_map = map (Lazy.map f) set_map,
   278     rel_cong = Lazy.map f rel_cong,
   279     rel_mono = Lazy.map f rel_mono,
   280     rel_mono_strong = Lazy.map f rel_mono_strong,
   281     rel_Grp = Lazy.map f rel_Grp,
   282     rel_conversep = Lazy.map f rel_conversep,
   283     rel_OO = Lazy.map f rel_OO};
   284 
   285 val morph_facts = map_facts o Morphism.thm;
   286 
   287 type nonemptiness_witness = {
   288   I: int list,
   289   wit: term,
   290   prop: thm list
   291 };
   292 
   293 fun mk_witness (I, wit) prop = {I = I, wit = wit, prop = prop};
   294 fun map_witness f g {I, wit, prop} = {I = I, wit = f wit, prop = map g prop};
   295 fun morph_witness phi = map_witness (Morphism.term phi) (Morphism.thm phi);
   296 
   297 datatype bnf = BNF of {
   298   name: binding,
   299   T: typ,
   300   live: int,
   301   lives: typ list, (*source type variables of map*)
   302   lives': typ list, (*target type variables of map*)
   303   dead: int,
   304   deads: typ list,
   305   map: term,
   306   sets: term list,
   307   bd: term,
   308   axioms: axioms,
   309   defs: defs,
   310   facts: facts,
   311   nwits: int,
   312   wits: nonemptiness_witness list,
   313   rel: term
   314 };
   315 
   316 (* getters *)
   317 
   318 fun rep_bnf (BNF bnf) = bnf;
   319 val name_of_bnf = #name o rep_bnf;
   320 val T_of_bnf = #T o rep_bnf;
   321 fun mk_T_of_bnf Ds Ts bnf =
   322   let val bnf_rep = rep_bnf bnf
   323   in Term.typ_subst_atomic ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#T bnf_rep) end;
   324 val live_of_bnf = #live o rep_bnf;
   325 val lives_of_bnf = #lives o rep_bnf;
   326 val dead_of_bnf = #dead o rep_bnf;
   327 val deads_of_bnf = #deads o rep_bnf;
   328 val axioms_of_bnf = #axioms o rep_bnf;
   329 val facts_of_bnf = #facts o rep_bnf;
   330 val nwits_of_bnf = #nwits o rep_bnf;
   331 val wits_of_bnf = #wits o rep_bnf;
   332 
   333 fun flatten_type_args_of_bnf bnf dead_x xs =
   334   let
   335     val Type (_, Ts) = T_of_bnf bnf;
   336     val lives = lives_of_bnf bnf;
   337     val deads = deads_of_bnf bnf;
   338   in
   339     permute_like (op =) (deads @ lives) Ts (replicate (length deads) dead_x @ xs)
   340   end;
   341 
   342 (*terms*)
   343 val map_of_bnf = #map o rep_bnf;
   344 val sets_of_bnf = #sets o rep_bnf;
   345 fun mk_map_of_bnf Ds Ts Us bnf =
   346   let val bnf_rep = rep_bnf bnf;
   347   in
   348     Term.subst_atomic_types
   349       ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#map bnf_rep)
   350   end;
   351 fun mk_sets_of_bnf Dss Tss bnf =
   352   let val bnf_rep = rep_bnf bnf;
   353   in
   354     map2 (fn (Ds, Ts) => Term.subst_atomic_types
   355       ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts))) (Dss ~~ Tss) (#sets bnf_rep)
   356   end;
   357 val bd_of_bnf = #bd o rep_bnf;
   358 fun mk_bd_of_bnf Ds Ts bnf =
   359   let val bnf_rep = rep_bnf bnf;
   360   in Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#bd bnf_rep) end;
   361 fun mk_wits_of_bnf Dss Tss bnf =
   362   let
   363     val bnf_rep = rep_bnf bnf;
   364     val wits = map (fn x => (#I x, #wit x)) (#wits bnf_rep);
   365   in
   366     map2 (fn (Ds, Ts) => apsnd (Term.subst_atomic_types
   367       ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)))) (Dss ~~ Tss) wits
   368   end;
   369 val rel_of_bnf = #rel o rep_bnf;
   370 fun mk_rel_of_bnf Ds Ts Us bnf =
   371   let val bnf_rep = rep_bnf bnf;
   372   in
   373     Term.subst_atomic_types
   374       ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#rel bnf_rep)
   375   end;
   376 
   377 (*thms*)
   378 val bd_card_order_of_bnf = #bd_card_order o #axioms o rep_bnf;
   379 val bd_cinfinite_of_bnf = #bd_cinfinite o #axioms o rep_bnf;
   380 val bd_Card_order_of_bnf = #bd_Card_order o #facts o rep_bnf;
   381 val bd_Cinfinite_of_bnf = #bd_Cinfinite o #facts o rep_bnf;
   382 val bd_Cnotzero_of_bnf = #bd_Cnotzero o #facts o rep_bnf;
   383 val collect_set_map_of_bnf = Lazy.force o #collect_set_map o #facts o rep_bnf;
   384 val in_bd_of_bnf = Lazy.force o #in_bd o #facts o rep_bnf;
   385 val in_cong_of_bnf = Lazy.force o #in_cong o #facts o rep_bnf;
   386 val in_mono_of_bnf = Lazy.force o #in_mono o #facts o rep_bnf;
   387 val in_rel_of_bnf = Lazy.force o #in_rel o #facts o rep_bnf;
   388 val map_def_of_bnf = #map_def o #defs o rep_bnf;
   389 val map_id0_of_bnf = #map_id0 o #axioms o rep_bnf;
   390 val map_id_of_bnf = Lazy.force o #map_id o #facts o rep_bnf;
   391 val map_comp0_of_bnf = #map_comp0 o #axioms o rep_bnf;
   392 val map_comp_of_bnf = Lazy.force o #map_comp o #facts o rep_bnf;
   393 val map_cong0_of_bnf = #map_cong0 o #axioms o rep_bnf;
   394 val map_cong_of_bnf = Lazy.force o #map_cong o #facts o rep_bnf;
   395 val map_transfer_of_bnf = Lazy.force o #map_transfer o #facts o rep_bnf;
   396 val map_wppull_of_bnf = Lazy.force o #map_wppull o #facts o rep_bnf;
   397 val map_wpull_of_bnf = #map_wpull o #axioms o rep_bnf;
   398 val rel_def_of_bnf = #rel_def o #defs o rep_bnf;
   399 val rel_eq_of_bnf = Lazy.force o #rel_eq o #facts o rep_bnf;
   400 val rel_flip_of_bnf = Lazy.force o #rel_flip o #facts o rep_bnf;
   401 val set_bd_of_bnf = #set_bd o #axioms o rep_bnf;
   402 val set_defs_of_bnf = #set_defs o #defs o rep_bnf;
   403 val set_map0_of_bnf = #set_map0 o #axioms o rep_bnf;
   404 val set_map_of_bnf = map Lazy.force o #set_map o #facts o rep_bnf;
   405 val rel_cong_of_bnf = Lazy.force o #rel_cong o #facts o rep_bnf;
   406 val rel_mono_of_bnf = Lazy.force o #rel_mono o #facts o rep_bnf;
   407 val rel_mono_strong_of_bnf = Lazy.force o #rel_mono_strong o #facts o rep_bnf;
   408 val rel_Grp_of_bnf = Lazy.force o #rel_Grp o #facts o rep_bnf;
   409 val rel_conversep_of_bnf = Lazy.force o #rel_conversep o #facts o rep_bnf;
   410 val rel_OO_of_bnf = Lazy.force o #rel_OO o #facts o rep_bnf;
   411 val rel_OO_Grp_of_bnf = #rel_OO_Grp o #axioms o rep_bnf;
   412 val wit_thms_of_bnf = maps #prop o wits_of_bnf;
   413 val wit_thmss_of_bnf = map #prop o wits_of_bnf;
   414 
   415 fun mk_bnf name T live lives lives' dead deads map sets bd axioms defs facts wits rel =
   416   BNF {name = name, T = T,
   417        live = live, lives = lives, lives' = lives', dead = dead, deads = deads,
   418        map = map, sets = sets, bd = bd,
   419        axioms = axioms, defs = defs, facts = facts,
   420        nwits = length wits, wits = wits, rel = rel};
   421 
   422 fun morph_bnf phi (BNF {name = name, T = T, live = live, lives = lives, lives' = lives',
   423   dead = dead, deads = deads, map = map, sets = sets, bd = bd,
   424   axioms = axioms, defs = defs, facts = facts,
   425   nwits = nwits, wits = wits, rel = rel}) =
   426   BNF {name = Morphism.binding phi name, T = Morphism.typ phi T,
   427     live = live, lives = List.map (Morphism.typ phi) lives,
   428     lives' = List.map (Morphism.typ phi) lives',
   429     dead = dead, deads = List.map (Morphism.typ phi) deads,
   430     map = Morphism.term phi map, sets = List.map (Morphism.term phi) sets,
   431     bd = Morphism.term phi bd,
   432     axioms = morph_axioms phi axioms,
   433     defs = morph_defs phi defs,
   434     facts = morph_facts phi facts,
   435     nwits = nwits,
   436     wits = List.map (morph_witness phi) wits,
   437     rel = Morphism.term phi rel};
   438 
   439 fun eq_bnf (BNF {T = T1, live = live1, dead = dead1, ...},
   440   BNF {T = T2, live = live2, dead = dead2, ...}) =
   441   Type.could_unify (T1, T2) andalso live1 = live2 andalso dead1 = dead2;
   442 
   443 structure Data = Generic_Data
   444 (
   445   type T = bnf Symtab.table;
   446   val empty = Symtab.empty;
   447   val extend = I;
   448   val merge = Symtab.merge eq_bnf;
   449 );
   450 
   451 fun bnf_of ctxt =
   452   Symtab.lookup (Data.get (Context.Proof ctxt))
   453   #> Option.map (morph_bnf (Morphism.thm_morphism (Thm.transfer (Proof_Context.theory_of ctxt))));
   454 
   455 
   456 (* Utilities *)
   457 
   458 fun normalize_set insts instA set =
   459   let
   460     val (T, T') = dest_funT (fastype_of set);
   461     val A = fst (Term.dest_TVar (HOLogic.dest_setT T'));
   462     val params = Term.add_tvar_namesT T [];
   463   in Term.subst_TVars ((A :: params) ~~ (instA :: insts)) set end;
   464 
   465 fun normalize_rel ctxt instTs instA instB rel =
   466   let
   467     val thy = Proof_Context.theory_of ctxt;
   468     val tyenv =
   469       Sign.typ_match thy (fastype_of rel, Library.foldr (op -->) (instTs, mk_pred2T instA instB))
   470         Vartab.empty;
   471   in Envir.subst_term (tyenv, Vartab.empty) rel end
   472   handle Type.TYPE_MATCH => error "Bad relator";
   473 
   474 fun normalize_wit insts CA As wit =
   475   let
   476     fun strip_param (Ts, T as Type (@{type_name fun}, [T1, T2])) =
   477         if Type.raw_instance (CA, T) then (Ts, T) else strip_param (T1 :: Ts, T2)
   478       | strip_param x = x;
   479     val (Ts, T) = strip_param ([], fastype_of wit);
   480     val subst = Term.add_tvar_namesT T [] ~~ insts;
   481     fun find y = find_index (fn x => x = y) As;
   482   in
   483     (map (find o Term.typ_subst_TVars subst) (rev Ts), Term.subst_TVars subst wit)
   484   end;
   485 
   486 fun minimize_wits wits =
   487  let
   488    fun minimize done [] = done
   489      | minimize done ((I, wit) :: todo) =
   490        if exists (fn (J, _) => subset (op =) (J, I)) (done @ todo)
   491        then minimize done todo
   492        else minimize ((I, wit) :: done) todo;
   493  in minimize [] wits end;
   494 
   495 fun mk_map live Ts Us t =
   496   let val (Type (_, Ts0), Type (_, Us0)) = strip_typeN (live + 1) (fastype_of t) |>> List.last in
   497     Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
   498   end;
   499 
   500 fun mk_rel live Ts Us t =
   501   let val [Type (_, Ts0), Type (_, Us0)] = binder_types (snd (strip_typeN live (fastype_of t))) in
   502     Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
   503   end;
   504 
   505 fun build_map_or_rel mk const of_bnf dest ctxt build_simple =
   506   let
   507     fun build (TU as (T, U)) =
   508       if T = U then
   509         const T
   510       else
   511         (case TU of
   512           (Type (s, Ts), Type (s', Us)) =>
   513           if s = s' then
   514             let
   515               val bnf = the (bnf_of ctxt s);
   516               val live = live_of_bnf bnf;
   517               val mapx = mk live Ts Us (of_bnf bnf);
   518               val TUs' = map dest (fst (strip_typeN live (fastype_of mapx)));
   519             in Term.list_comb (mapx, map build TUs') end
   520           else
   521             build_simple TU
   522         | _ => build_simple TU);
   523   in build end;
   524 
   525 val build_map = build_map_or_rel mk_map HOLogic.id_const map_of_bnf dest_funT;
   526 val build_rel = build_map_or_rel mk_rel HOLogic.eq_const rel_of_bnf dest_pred2T;
   527 
   528 fun map_flattened_map_args ctxt s map_args fs =
   529   let
   530     val flat_fs = flatten_type_args_of_bnf (the (bnf_of ctxt s)) Term.dummy fs;
   531     val flat_fs' = map_args flat_fs;
   532   in
   533     permute_like (op aconv) flat_fs fs flat_fs'
   534   end;
   535 
   536 
   537 (* Names *)
   538 
   539 val mapN = "map";
   540 val setN = "set";
   541 fun mk_setN i = setN ^ nonzero_string_of_int i;
   542 val bdN = "bd";
   543 val witN = "wit";
   544 fun mk_witN i = witN ^ nonzero_string_of_int i;
   545 val relN = "rel";
   546 
   547 val bd_card_orderN = "bd_card_order";
   548 val bd_cinfiniteN = "bd_cinfinite";
   549 val bd_Card_orderN = "bd_Card_order";
   550 val bd_CinfiniteN = "bd_Cinfinite";
   551 val bd_CnotzeroN = "bd_Cnotzero";
   552 val collect_set_mapN = "collect_set_map";
   553 val in_bdN = "in_bd";
   554 val in_monoN = "in_mono";
   555 val in_relN = "in_rel";
   556 val map_id0N = "map_id0";
   557 val map_idN = "map_id";
   558 val map_comp0N = "map_comp0";
   559 val map_compN = "map_comp";
   560 val map_cong0N = "map_cong0";
   561 val map_congN = "map_cong";
   562 val map_transferN = "map_transfer";
   563 val map_wpullN = "map_wpull";
   564 val rel_eqN = "rel_eq";
   565 val rel_flipN = "rel_flip";
   566 val set_map0N = "set_map0";
   567 val set_mapN = "set_map";
   568 val set_bdN = "set_bd";
   569 val rel_GrpN = "rel_Grp";
   570 val rel_conversepN = "rel_conversep";
   571 val rel_monoN = "rel_mono"
   572 val rel_mono_strongN = "rel_mono_strong"
   573 val rel_OON = "rel_compp";
   574 val rel_OO_GrpN = "rel_compp_Grp";
   575 
   576 datatype const_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline;
   577 
   578 datatype fact_policy = Dont_Note | Note_Some | Note_All;
   579 
   580 val bnf_note_all = Attrib.setup_config_bool @{binding bnf_note_all} (K false);
   581 val bnf_timing = Attrib.setup_config_bool @{binding bnf_timing} (K false);
   582 
   583 fun user_policy policy ctxt = if Config.get ctxt bnf_note_all then Note_All else policy;
   584 
   585 val smart_max_inline_size = 25; (*FUDGE*)
   586 
   587 fun note_bnf_thms fact_policy qualify' bnf_b bnf =
   588   let
   589     val axioms = axioms_of_bnf bnf;
   590     val facts = facts_of_bnf bnf;
   591     val wits = wits_of_bnf bnf;
   592     val qualify =
   593       let val (_, qs, _) = Binding.dest bnf_b;
   594       in fold_rev (fn (s, mand) => Binding.qualify mand s) qs #> qualify' end;
   595   in
   596     (if fact_policy = Note_All then
   597       let
   598         val witNs = if length wits = 1 then [witN] else map mk_witN (1 upto length wits);
   599         val notes =
   600           [(bd_card_orderN, [#bd_card_order axioms]),
   601             (bd_cinfiniteN, [#bd_cinfinite axioms]),
   602             (bd_Card_orderN, [#bd_Card_order facts]),
   603             (bd_CinfiniteN, [#bd_Cinfinite facts]),
   604             (bd_CnotzeroN, [#bd_Cnotzero facts]),
   605             (collect_set_mapN, [Lazy.force (#collect_set_map facts)]),
   606             (in_bdN, [Lazy.force (#in_bd facts)]),
   607             (in_monoN, [Lazy.force (#in_mono facts)]),
   608             (in_relN, [Lazy.force (#in_rel facts)]),
   609             (map_comp0N, [#map_comp0 axioms]),
   610             (map_id0N, [#map_id0 axioms]),
   611             (map_transferN, [Lazy.force (#map_transfer facts)]),
   612             (map_wpullN, [#map_wpull axioms]),
   613             (rel_mono_strongN, [Lazy.force (#rel_mono_strong facts)]),
   614             (set_map0N, #set_map0 axioms),
   615             (set_bdN, #set_bd axioms)] @
   616             (witNs ~~ wit_thmss_of_bnf bnf)
   617             |> map (fn (thmN, thms) =>
   618               ((qualify (Binding.qualify true (Binding.name_of bnf_b) (Binding.name thmN)), []),
   619               [(thms, [])]));
   620         in
   621           Local_Theory.notes notes #> snd
   622         end
   623       else
   624         I)
   625     #> (if fact_policy <> Dont_Note then
   626         let
   627           val notes =
   628             [(map_compN, [Lazy.force (#map_comp facts)], []),
   629             (map_cong0N, [#map_cong0 axioms], []),
   630             (map_congN, [Lazy.force (#map_cong facts)], fundef_cong_attrs),
   631             (map_idN, [Lazy.force (#map_id facts)], []),
   632             (rel_eqN, [Lazy.force (#rel_eq facts)], []),
   633             (rel_flipN, [Lazy.force (#rel_flip facts)], []),
   634             (set_mapN, map Lazy.force (#set_map facts), []),
   635             (rel_OO_GrpN, no_refl [#rel_OO_Grp axioms], []),
   636             (rel_GrpN, [Lazy.force (#rel_Grp facts)], []),
   637             (rel_conversepN, [Lazy.force (#rel_conversep facts)], []),
   638             (rel_monoN, [Lazy.force (#rel_mono facts)], []),
   639             (rel_OON, [Lazy.force (#rel_OO facts)], [])]
   640             |> filter_out (null o #2)
   641             |> map (fn (thmN, thms, attrs) =>
   642               ((qualify (Binding.qualify true (Binding.name_of bnf_b) (Binding.name thmN)),
   643                 attrs), [(thms, [])]));
   644         in
   645           Local_Theory.notes notes #> snd
   646         end
   647       else
   648         I)
   649   end;
   650 
   651 
   652 (* Define new BNFs *)
   653 
   654 fun prepare_def const_policy mk_fact_policy qualify prep_term Ds_opt map_b rel_b set_bs
   655   (((((raw_bnf_b, raw_map), raw_sets), raw_bd_Abs), raw_wits), raw_rel_opt) no_defs_lthy =
   656   let
   657     val fact_policy = mk_fact_policy no_defs_lthy;
   658     val bnf_b = qualify raw_bnf_b;
   659     val live = length raw_sets;
   660 
   661     val map_rhs = prep_term no_defs_lthy raw_map;
   662     val set_rhss = map (prep_term no_defs_lthy) raw_sets;
   663     val (bd_rhsT, bd_rhs) = (case prep_term no_defs_lthy raw_bd_Abs of
   664       Abs (_, T, t) => (T, t)
   665     | _ => error "Bad bound constant");
   666 
   667     fun err T =
   668       error ("Trying to register the type " ^ quote (Syntax.string_of_typ no_defs_lthy T) ^
   669         " as unnamed BNF");
   670 
   671     val (bnf_b, key) =
   672       if Binding.eq_name (bnf_b, Binding.empty) then
   673         (case bd_rhsT of
   674           Type (C, Ts) => if forall (can dest_TFree) Ts
   675             then (Binding.qualified_name C, C) else err bd_rhsT
   676         | T => err T)
   677       else (bnf_b, Local_Theory.full_name no_defs_lthy bnf_b);
   678 
   679     val def_qualify = Binding.conceal o Binding.qualify false (Binding.name_of bnf_b);
   680 
   681     fun mk_suffix_binding suf = Binding.suffix_name ("_" ^ suf) bnf_b;
   682 
   683     fun maybe_define user_specified (b, rhs) lthy =
   684       let
   685         val inline =
   686           (user_specified orelse fact_policy = Dont_Note) andalso
   687           (case const_policy of
   688             Dont_Inline => false
   689           | Hardly_Inline => Term.is_Free rhs orelse Term.is_Const rhs
   690           | Smart_Inline => Term.size_of_term rhs <= smart_max_inline_size
   691           | Do_Inline => true)
   692       in
   693         if inline then
   694           ((rhs, Drule.reflexive_thm), lthy)
   695         else
   696           let val b = b () in
   697             apfst (apsnd snd) (Local_Theory.define ((b, NoSyn), ((Thm.def_binding b, []), rhs))
   698               lthy)
   699           end
   700       end;
   701 
   702     fun maybe_restore lthy_old lthy =
   703       lthy |> not (pointer_eq (lthy_old, lthy)) ? Local_Theory.restore;
   704 
   705     val map_bind_def =
   706       (fn () => def_qualify (if Binding.is_empty map_b then mk_suffix_binding mapN else map_b),
   707          map_rhs);
   708     val set_binds_defs =
   709       let
   710         fun set_name i get_b =
   711           (case try (nth set_bs) (i - 1) of
   712             SOME b => if Binding.is_empty b then get_b else K b
   713           | NONE => get_b) #> def_qualify;
   714         val bs = if live = 1 then [set_name 1 (fn () => mk_suffix_binding setN)]
   715           else map (fn i => set_name i (fn () => mk_suffix_binding (mk_setN i))) (1 upto live);
   716       in bs ~~ set_rhss end;
   717     val bd_bind_def = (fn () => def_qualify (mk_suffix_binding bdN), bd_rhs);
   718 
   719     val ((((bnf_map_term, raw_map_def),
   720       (bnf_set_terms, raw_set_defs)),
   721       (bnf_bd_term, raw_bd_def)), (lthy, lthy_old)) =
   722         no_defs_lthy
   723         |> maybe_define true map_bind_def
   724         ||>> apfst split_list o fold_map (maybe_define true) set_binds_defs
   725         ||>> maybe_define true bd_bind_def
   726         ||> `(maybe_restore no_defs_lthy);
   727 
   728     val phi = Proof_Context.export_morphism lthy_old lthy;
   729 
   730     val bnf_map_def = Morphism.thm phi raw_map_def;
   731     val bnf_set_defs = map (Morphism.thm phi) raw_set_defs;
   732     val bnf_bd_def = Morphism.thm phi raw_bd_def;
   733 
   734     val bnf_map = Morphism.term phi bnf_map_term;
   735 
   736     (*TODO: handle errors*)
   737     (*simple shape analysis of a map function*)
   738     val ((alphas, betas), (CA, _)) =
   739       fastype_of bnf_map
   740       |> strip_typeN live
   741       |>> map_split dest_funT
   742       ||> dest_funT
   743       handle TYPE _ => error "Bad map function";
   744 
   745     val CA_params = map TVar (Term.add_tvarsT CA []);
   746 
   747     val bnf_sets = map2 (normalize_set CA_params) alphas (map (Morphism.term phi) bnf_set_terms);
   748     val bdT = Morphism.typ phi bd_rhsT;
   749     val bnf_bd =
   750       Term.subst_TVars (Term.add_tvar_namesT bdT [] ~~ CA_params) (Morphism.term phi bnf_bd_term);
   751 
   752     (*TODO: assert Ds = (TVars of bnf_map) \ (alphas @ betas) as sets*)
   753     val deads = (case Ds_opt of
   754       NONE => subtract (op =) (alphas @ betas) (map TVar (Term.add_tvars bnf_map []))
   755     | SOME Ds => map (Morphism.typ phi) Ds);
   756     val dead = length deads;
   757 
   758     (*TODO: further checks of type of bnf_map*)
   759     (*TODO: check types of bnf_sets*)
   760     (*TODO: check type of bnf_bd*)
   761     (*TODO: check type of bnf_rel*)
   762 
   763     val ((((((((((As', Bs'), Cs), Ds), B1Ts), B2Ts), domTs), ranTs), ranTs'), ranTs''),
   764       (Ts, T)) = lthy
   765       |> mk_TFrees live
   766       ||>> mk_TFrees live
   767       ||>> mk_TFrees live
   768       ||>> mk_TFrees dead
   769       ||>> mk_TFrees live
   770       ||>> mk_TFrees live
   771       ||>> mk_TFrees live
   772       ||>> mk_TFrees live
   773       ||>> mk_TFrees live
   774       ||>> mk_TFrees live
   775       ||> fst o mk_TFrees 1
   776       ||> the_single
   777       ||> `(replicate live);
   778 
   779     fun mk_bnf_map As' Bs' =
   780       Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As') @ (betas ~~ Bs')) bnf_map;
   781     fun mk_bnf_t As' = Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As'));
   782     fun mk_bnf_T As' = Term.typ_subst_atomic ((deads ~~ Ds) @ (alphas ~~ As'));
   783 
   784     val RTs = map HOLogic.mk_prodT (As' ~~ Bs');
   785     val pred2RTs = map2 mk_pred2T As' Bs';
   786     val pred2RTsAsCs = map2 mk_pred2T As' Cs;
   787     val pred2RTsBsCs = map2 mk_pred2T Bs' Cs;
   788     val pred2RT's = map2 mk_pred2T Bs' As';
   789     val self_pred2RTs = map2 mk_pred2T As' As';
   790     val transfer_domRTs = map2 mk_pred2T As' B1Ts;
   791     val transfer_ranRTs = map2 mk_pred2T Bs' B2Ts;
   792 
   793     val CA' = mk_bnf_T As' CA;
   794     val CB' = mk_bnf_T Bs' CA;
   795     val CC' = mk_bnf_T Cs CA;
   796     val CRs' = mk_bnf_T RTs CA;
   797     val CB1 = mk_bnf_T B1Ts CA;
   798     val CB2 = mk_bnf_T B2Ts CA;
   799 
   800     val bnf_map_AsAs = mk_bnf_map As' As';
   801     val bnf_map_AsBs = mk_bnf_map As' Bs';
   802     val bnf_map_AsCs = mk_bnf_map As' Cs;
   803     val bnf_map_BsCs = mk_bnf_map Bs' Cs;
   804     val bnf_sets_As = map (mk_bnf_t As') bnf_sets;
   805     val bnf_sets_Bs = map (mk_bnf_t Bs') bnf_sets;
   806     val bnf_bd_As = mk_bnf_t As' bnf_bd;
   807 
   808     val pre_names_lthy = lthy;
   809     val ((((((((((((((((((((((((fs, gs), hs), x), y), zs), ys), As),
   810       As_copy), Xs), B1s), B2s), f1s), f2s), e1s), e2s), p1s), p2s), bs), (Rs, Rs')), Rs_copy), Ss),
   811       transfer_domRs), transfer_ranRs), names_lthy) = pre_names_lthy
   812       |> mk_Frees "f" (map2 (curry op -->) As' Bs')
   813       ||>> mk_Frees "g" (map2 (curry op -->) Bs' Cs)
   814       ||>> mk_Frees "h" (map2 (curry op -->) As' Ts)
   815       ||>> yield_singleton (mk_Frees "x") CA'
   816       ||>> yield_singleton (mk_Frees "y") CB'
   817       ||>> mk_Frees "z" As'
   818       ||>> mk_Frees "y" Bs'
   819       ||>> mk_Frees "A" (map HOLogic.mk_setT As')
   820       ||>> mk_Frees "A" (map HOLogic.mk_setT As')
   821       ||>> mk_Frees "A" (map HOLogic.mk_setT domTs)
   822       ||>> mk_Frees "B1" (map HOLogic.mk_setT B1Ts)
   823       ||>> mk_Frees "B2" (map HOLogic.mk_setT B2Ts)
   824       ||>> mk_Frees "f1" (map2 (curry op -->) B1Ts ranTs)
   825       ||>> mk_Frees "f2" (map2 (curry op -->) B2Ts ranTs)
   826       ||>> mk_Frees "e1" (map2 (curry op -->) B1Ts ranTs')
   827       ||>> mk_Frees "e2" (map2 (curry op -->) B2Ts ranTs'')
   828       ||>> mk_Frees "p1" (map2 (curry op -->) domTs B1Ts)
   829       ||>> mk_Frees "p2" (map2 (curry op -->) domTs B2Ts)
   830       ||>> mk_Frees "b" As'
   831       ||>> mk_Frees' "R" pred2RTs
   832       ||>> mk_Frees "R" pred2RTs
   833       ||>> mk_Frees "S" pred2RTsBsCs
   834       ||>> mk_Frees "R" transfer_domRTs
   835       ||>> mk_Frees "S" transfer_ranRTs;
   836 
   837     val fs_copy = map2 (retype_free o fastype_of) fs gs;
   838     val x_copy = retype_free CA' y;
   839 
   840     val setRs =
   841       map3 (fn R => fn T => fn U =>
   842           HOLogic.Collect_const (HOLogic.mk_prodT (T, U)) $ HOLogic.mk_split R) Rs As' Bs';
   843 
   844     (*Grp (in (Collect (split R1) .. Collect (split Rn))) (map fst .. fst)^--1 OO
   845       Grp (in (Collect (split R1) .. Collect (split Rn))) (map snd .. snd)*)
   846     val OO_Grp =
   847       let
   848         val map1 = Term.list_comb (mk_bnf_map RTs As', map fst_const RTs);
   849         val map2 = Term.list_comb (mk_bnf_map RTs Bs', map snd_const RTs);
   850         val bnf_in = mk_in setRs (map (mk_bnf_t RTs) bnf_sets) CRs';
   851       in
   852         mk_rel_compp (mk_conversep (mk_Grp bnf_in map1), mk_Grp bnf_in map2)
   853       end;
   854 
   855     val rel_rhs = (case raw_rel_opt of
   856         NONE => fold_rev Term.absfree Rs' OO_Grp
   857       | SOME raw_rel => prep_term no_defs_lthy raw_rel);
   858 
   859     val rel_bind_def =
   860       (fn () => def_qualify (if Binding.is_empty rel_b then mk_suffix_binding relN else rel_b),
   861          rel_rhs);
   862 
   863     val wit_rhss =
   864       if null raw_wits then
   865         [fold_rev Term.absdummy As' (Term.list_comb (bnf_map_AsAs,
   866           map2 (fn T => fn i => Term.absdummy T (Bound i)) As' (live downto 1)) $
   867           Const (@{const_name undefined}, CA'))]
   868       else map (prep_term no_defs_lthy) raw_wits;
   869     val nwits = length wit_rhss;
   870     val wit_binds_defs =
   871       let
   872         val bs = if nwits = 1 then [fn () => def_qualify (mk_suffix_binding witN)]
   873           else map (fn i => fn () => def_qualify (mk_suffix_binding (mk_witN i))) (1 upto nwits);
   874       in bs ~~ wit_rhss end;
   875 
   876     val (((bnf_rel_term, raw_rel_def), (bnf_wit_terms, raw_wit_defs)), (lthy, lthy_old)) =
   877       lthy
   878       |> maybe_define (is_some raw_rel_opt) rel_bind_def
   879       ||>> apfst split_list o fold_map (maybe_define (not (null raw_wits))) wit_binds_defs
   880       ||> `(maybe_restore lthy);
   881 
   882     val phi = Proof_Context.export_morphism lthy_old lthy;
   883     val bnf_rel_def = Morphism.thm phi raw_rel_def;
   884     val bnf_rel = Morphism.term phi bnf_rel_term;
   885 
   886     fun mk_bnf_rel RTs CA' CB' = normalize_rel lthy RTs CA' CB' bnf_rel;
   887 
   888     val rel = mk_bnf_rel pred2RTs CA' CB';
   889     val relAsAs = mk_bnf_rel self_pred2RTs CA' CA';
   890 
   891     val bnf_wit_defs = map (Morphism.thm phi) raw_wit_defs;
   892     val bnf_wits = map (normalize_wit CA_params CA alphas o Morphism.term phi) bnf_wit_terms;
   893     val bnf_wit_As = map (apsnd (mk_bnf_t As')) bnf_wits;
   894 
   895     val map_id0_goal =
   896       let val bnf_map_app_id = Term.list_comb (bnf_map_AsAs, map HOLogic.id_const As') in
   897         mk_Trueprop_eq (bnf_map_app_id, HOLogic.id_const CA')
   898       end;
   899 
   900     val map_comp0_goal =
   901       let
   902         val bnf_map_app_comp = Term.list_comb (bnf_map_AsCs, map2 (curry HOLogic.mk_comp) gs fs);
   903         val comp_bnf_map_app = HOLogic.mk_comp
   904           (Term.list_comb (bnf_map_BsCs, gs), Term.list_comb (bnf_map_AsBs, fs));
   905       in
   906         fold_rev Logic.all (fs @ gs) (mk_Trueprop_eq (bnf_map_app_comp, comp_bnf_map_app))
   907       end;
   908 
   909     fun mk_map_cong_prem x z set f f_copy =
   910       Logic.all z (Logic.mk_implies
   911         (HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set $ x)),
   912         mk_Trueprop_eq (f $ z, f_copy $ z)));
   913 
   914     val map_cong0_goal =
   915       let
   916         val prems = map4 (mk_map_cong_prem x) zs bnf_sets_As fs fs_copy;
   917         val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x,
   918           Term.list_comb (bnf_map_AsBs, fs_copy) $ x);
   919       in
   920         fold_rev Logic.all (x :: fs @ fs_copy) (Logic.list_implies (prems, eq))
   921       end;
   922 
   923     val set_map0s_goal =
   924       let
   925         fun mk_goal setA setB f =
   926           let
   927             val set_comp_map =
   928               HOLogic.mk_comp (setB, Term.list_comb (bnf_map_AsBs, fs));
   929             val image_comp_set = HOLogic.mk_comp (mk_image f, setA);
   930           in
   931             fold_rev Logic.all fs (mk_Trueprop_eq (set_comp_map, image_comp_set))
   932           end;
   933       in
   934         map3 mk_goal bnf_sets_As bnf_sets_Bs fs
   935       end;
   936 
   937     val card_order_bd_goal = HOLogic.mk_Trueprop (mk_card_order bnf_bd_As);
   938 
   939     val cinfinite_bd_goal = HOLogic.mk_Trueprop (mk_cinfinite bnf_bd_As);
   940 
   941     val set_bds_goal =
   942       let
   943         fun mk_goal set =
   944           Logic.all x (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (set $ x)) bnf_bd_As));
   945       in
   946         map mk_goal bnf_sets_As
   947       end;
   948 
   949     val map_wpull_goal =
   950       let
   951         val prems = map HOLogic.mk_Trueprop
   952           (map8 mk_wpull Xs B1s B2s f1s f2s (replicate live NONE) p1s p2s);
   953         val CX = mk_bnf_T domTs CA;
   954         val bnf_sets_CX = map2 (normalize_set (map (mk_bnf_T domTs) CA_params)) domTs bnf_sets;
   955         val bnf_sets_CB1 = map2 (normalize_set (map (mk_bnf_T B1Ts) CA_params)) B1Ts bnf_sets;
   956         val bnf_sets_CB2 = map2 (normalize_set (map (mk_bnf_T B2Ts) CA_params)) B2Ts bnf_sets;
   957         val bnf_map_app_f1 = Term.list_comb (mk_bnf_map B1Ts ranTs, f1s);
   958         val bnf_map_app_f2 = Term.list_comb (mk_bnf_map B2Ts ranTs, f2s);
   959         val bnf_map_app_p1 = Term.list_comb (mk_bnf_map domTs B1Ts, p1s);
   960         val bnf_map_app_p2 = Term.list_comb (mk_bnf_map domTs B2Ts, p2s);
   961 
   962         val map_wpull = mk_wpull (mk_in Xs bnf_sets_CX CX)
   963           (mk_in B1s bnf_sets_CB1 CB1) (mk_in B2s bnf_sets_CB2 CB2)
   964           bnf_map_app_f1 bnf_map_app_f2 NONE bnf_map_app_p1 bnf_map_app_p2;
   965       in
   966         fold_rev Logic.all (Xs @ B1s @ B2s @ f1s @ f2s @ p1s @ p2s)
   967           (Logic.list_implies (prems, HOLogic.mk_Trueprop map_wpull))
   968       end;
   969 
   970     val rel_OO_Grp_goal = fold_rev Logic.all Rs (mk_Trueprop_eq (Term.list_comb (rel, Rs), OO_Grp));
   971 
   972     val goals = zip_axioms map_id0_goal map_comp0_goal map_cong0_goal set_map0s_goal
   973       card_order_bd_goal cinfinite_bd_goal set_bds_goal map_wpull_goal rel_OO_Grp_goal;
   974 
   975     fun mk_wit_goals (I, wit) =
   976       let
   977         val xs = map (nth bs) I;
   978         fun wit_goal i =
   979           let
   980             val z = nth zs i;
   981             val set_wit = nth bnf_sets_As i $ Term.list_comb (wit, xs);
   982             val concl = HOLogic.mk_Trueprop
   983               (if member (op =) I i then HOLogic.mk_eq (z, nth bs i)
   984               else @{term False});
   985           in
   986             fold_rev Logic.all (z :: xs)
   987               (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set_wit)), concl))
   988           end;
   989       in
   990         map wit_goal (0 upto live - 1)
   991       end;
   992 
   993     val trivial_wit_tac = mk_trivial_wit_tac bnf_wit_defs;
   994 
   995     val wit_goalss =
   996       (if null raw_wits then SOME trivial_wit_tac else NONE, map mk_wit_goals bnf_wit_As);
   997 
   998     fun after_qed mk_wit_thms thms lthy =
   999       let
  1000         val (axioms, nontriv_wit_thms) = apfst (mk_axioms live) (chop (length goals) thms);
  1001 
  1002         val bd_Card_order = #bd_card_order axioms RS @{thm conjunct2[OF card_order_on_Card_order]};
  1003         val bd_Cinfinite = @{thm conjI} OF [#bd_cinfinite axioms, bd_Card_order];
  1004         val bd_Cnotzero = bd_Cinfinite RS @{thm Cinfinite_Cnotzero};
  1005 
  1006         fun mk_collect_set_map () =
  1007           let
  1008             val defT = mk_bnf_T Ts CA --> HOLogic.mk_setT T;
  1009             val collect_map = HOLogic.mk_comp
  1010               (mk_collect (map (mk_bnf_t Ts) bnf_sets) defT,
  1011               Term.list_comb (mk_bnf_map As' Ts, hs));
  1012             val image_collect = mk_collect
  1013               (map2 (fn h => fn set => HOLogic.mk_comp (mk_image h, set)) hs bnf_sets_As)
  1014               defT;
  1015             (*collect {set1 ... setm} o map f1 ... fm = collect {f1` o set1 ... fm` o setm}*)
  1016             val goal = fold_rev Logic.all hs (mk_Trueprop_eq (collect_map, image_collect));
  1017           in
  1018             Goal.prove_sorry lthy [] [] goal (K (mk_collect_set_map_tac (#set_map0 axioms)))
  1019             |> Thm.close_derivation
  1020           end;
  1021 
  1022         val collect_set_map = Lazy.lazy mk_collect_set_map;
  1023 
  1024         fun mk_in_mono () =
  1025           let
  1026             val prems_mono = map2 (HOLogic.mk_Trueprop oo mk_leq) As As_copy;
  1027             val in_mono_goal =
  1028               fold_rev Logic.all (As @ As_copy)
  1029                 (Logic.list_implies (prems_mono, HOLogic.mk_Trueprop
  1030                   (mk_leq (mk_in As bnf_sets_As CA') (mk_in As_copy bnf_sets_As CA'))));
  1031           in
  1032             Goal.prove_sorry lthy [] [] in_mono_goal (K (mk_in_mono_tac live))
  1033             |> Thm.close_derivation
  1034           end;
  1035 
  1036         val in_mono = Lazy.lazy mk_in_mono;
  1037 
  1038         fun mk_in_cong () =
  1039           let
  1040             val prems_cong = map2 (curry mk_Trueprop_eq) As As_copy;
  1041             val in_cong_goal =
  1042               fold_rev Logic.all (As @ As_copy)
  1043                 (Logic.list_implies (prems_cong,
  1044                   mk_Trueprop_eq (mk_in As bnf_sets_As CA', mk_in As_copy bnf_sets_As CA')));
  1045           in
  1046             Goal.prove_sorry lthy [] [] in_cong_goal
  1047               (K ((TRY o hyp_subst_tac lthy THEN' rtac refl) 1))
  1048             |> Thm.close_derivation
  1049           end;
  1050 
  1051         val in_cong = Lazy.lazy mk_in_cong;
  1052 
  1053         val map_id = Lazy.lazy (fn () => mk_map_id (#map_id0 axioms));
  1054         val map_comp = Lazy.lazy (fn () => mk_map_comp (#map_comp0 axioms));
  1055 
  1056         fun mk_map_cong () =
  1057           let
  1058             val prem0 = mk_Trueprop_eq (x, x_copy);
  1059             val prems = map4 (mk_map_cong_prem x_copy) zs bnf_sets_As fs fs_copy;
  1060             val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x,
  1061               Term.list_comb (bnf_map_AsBs, fs_copy) $ x_copy);
  1062             val goal = fold_rev Logic.all (x :: x_copy :: fs @ fs_copy)
  1063               (Logic.list_implies (prem0 :: prems, eq));
  1064           in
  1065             Goal.prove_sorry lthy [] [] goal (fn _ => mk_map_cong_tac lthy (#map_cong0 axioms))
  1066             |> Thm.close_derivation
  1067           end;
  1068 
  1069         val map_cong = Lazy.lazy mk_map_cong;
  1070 
  1071         val set_map = map (fn thm => Lazy.lazy (fn () => mk_set_map thm)) (#set_map0 axioms);
  1072 
  1073         val wit_thms =
  1074           if null nontriv_wit_thms then mk_wit_thms (map Lazy.force set_map) else nontriv_wit_thms;
  1075 
  1076         fun mk_in_bd () =
  1077           let
  1078             val bdT = fst (dest_relT (fastype_of bnf_bd_As));
  1079             val bdTs = replicate live bdT;
  1080             val bd_bnfT = mk_bnf_T bdTs CA;
  1081             val surj_imp_ordLeq_inst = (if live = 0 then TrueI else
  1082               let
  1083                 val ranTs = map (fn AT => mk_sumT (AT, HOLogic.unitT)) As';
  1084                 val funTs = map (fn T => bdT --> T) ranTs;
  1085                 val ran_bnfT = mk_bnf_T ranTs CA;
  1086                 val (revTs, Ts) = `rev (bd_bnfT :: funTs);
  1087                 val cTs = map (SOME o certifyT lthy) [ran_bnfT, Library.foldr1 HOLogic.mk_prodT Ts];
  1088                 val tinst = fold (fn T => fn t => HOLogic.mk_split (Term.absdummy T t)) (tl revTs)
  1089                   (Term.absdummy (hd revTs) (Term.list_comb (mk_bnf_map bdTs ranTs,
  1090                     map Bound (live - 1 downto 0)) $ Bound live));
  1091                 val cts = [NONE, SOME (certify lthy tinst)];
  1092               in
  1093                 Drule.instantiate' cTs cts @{thm surj_imp_ordLeq}
  1094               end);
  1095             val bd = mk_cexp
  1096               (if live = 0 then ctwo
  1097                 else mk_csum (Library.foldr1 (uncurry mk_csum) (map mk_card_of As)) ctwo)
  1098               (mk_csum bnf_bd_As (mk_card_of (HOLogic.mk_UNIV bd_bnfT)));
  1099             val in_bd_goal =
  1100               fold_rev Logic.all As
  1101                 (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (mk_in As bnf_sets_As CA')) bd));
  1102           in
  1103             Goal.prove_sorry lthy [] [] in_bd_goal
  1104               (mk_in_bd_tac live surj_imp_ordLeq_inst
  1105                 (Lazy.force map_comp) (Lazy.force map_id) (#map_cong0 axioms)
  1106                 (map Lazy.force set_map) (#set_bd axioms) (#bd_card_order axioms)
  1107                 bd_Card_order bd_Cinfinite bd_Cnotzero)
  1108             |> Thm.close_derivation
  1109           end;
  1110 
  1111         val in_bd = Lazy.lazy mk_in_bd;
  1112 
  1113         fun mk_map_wppull () =
  1114           let
  1115             val prems = if live = 0 then [] else
  1116               [HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
  1117                 (map8 mk_wpull Xs B1s B2s f1s f2s (map SOME (e1s ~~ e2s)) p1s p2s))];
  1118             val CX = mk_bnf_T domTs CA;
  1119             val bnf_sets_CX =
  1120               map2 (normalize_set (map (mk_bnf_T domTs) CA_params)) domTs bnf_sets;
  1121             val bnf_sets_CB1 =
  1122               map2 (normalize_set (map (mk_bnf_T B1Ts) CA_params)) B1Ts bnf_sets;
  1123             val bnf_sets_CB2 =
  1124               map2 (normalize_set (map (mk_bnf_T B2Ts) CA_params)) B2Ts bnf_sets;
  1125             val bnf_map_app_f1 = Term.list_comb (mk_bnf_map B1Ts ranTs, f1s);
  1126             val bnf_map_app_f2 = Term.list_comb (mk_bnf_map B2Ts ranTs, f2s);
  1127             val bnf_map_app_e1 = Term.list_comb (mk_bnf_map B1Ts ranTs', e1s);
  1128             val bnf_map_app_e2 = Term.list_comb (mk_bnf_map B2Ts ranTs'', e2s);
  1129             val bnf_map_app_p1 = Term.list_comb (mk_bnf_map domTs B1Ts, p1s);
  1130             val bnf_map_app_p2 = Term.list_comb (mk_bnf_map domTs B2Ts, p2s);
  1131 
  1132             val concl = mk_wpull (mk_in Xs bnf_sets_CX CX)
  1133               (mk_in B1s bnf_sets_CB1 CB1) (mk_in B2s bnf_sets_CB2 CB2)
  1134               bnf_map_app_f1 bnf_map_app_f2 (SOME (bnf_map_app_e1, bnf_map_app_e2))
  1135               bnf_map_app_p1 bnf_map_app_p2;
  1136 
  1137             val goal =
  1138               fold_rev Logic.all (Xs @ B1s @ B2s @ f1s @ f2s @ e1s @ e2s @ p1s @ p2s)
  1139                 (Logic.list_implies (prems, HOLogic.mk_Trueprop concl))
  1140           in
  1141             Goal.prove_sorry lthy [] [] goal
  1142               (fn _ => mk_map_wppull_tac (#map_id0 axioms) (#map_cong0 axioms)
  1143                 (#map_wpull axioms) (Lazy.force map_comp) (map Lazy.force set_map))
  1144             |> Thm.close_derivation
  1145           end;
  1146 
  1147         val map_wppull = Lazy.lazy mk_map_wppull;
  1148 
  1149         val rel_OO_Grp = #rel_OO_Grp axioms;
  1150         val rel_OO_Grps = no_refl [rel_OO_Grp];
  1151 
  1152         fun mk_rel_Grp () =
  1153           let
  1154             val lhs = Term.list_comb (rel, map2 mk_Grp As fs);
  1155             val rhs = mk_Grp (mk_in As bnf_sets_As CA') (Term.list_comb (bnf_map_AsBs, fs));
  1156             val goal = fold_rev Logic.all (As @ fs) (mk_Trueprop_eq (lhs, rhs));
  1157           in
  1158             Goal.prove_sorry lthy [] [] goal
  1159               (mk_rel_Grp_tac rel_OO_Grps (#map_id0 axioms) (#map_cong0 axioms) (Lazy.force map_id)
  1160                 (Lazy.force map_comp) (map Lazy.force set_map))
  1161             |> Thm.close_derivation
  1162           end;
  1163 
  1164         val rel_Grp = Lazy.lazy mk_rel_Grp;
  1165 
  1166         fun mk_rel_prems f = map2 (HOLogic.mk_Trueprop oo f) Rs Rs_copy
  1167         fun mk_rel_concl f = HOLogic.mk_Trueprop
  1168           (f (Term.list_comb (rel, Rs), Term.list_comb (rel, Rs_copy)));
  1169 
  1170         fun mk_rel_mono () =
  1171           let
  1172             val mono_prems = mk_rel_prems mk_leq;
  1173             val mono_concl = mk_rel_concl (uncurry mk_leq);
  1174           in
  1175             Goal.prove_sorry lthy [] []
  1176               (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (mono_prems, mono_concl)))
  1177               (K (mk_rel_mono_tac rel_OO_Grps (Lazy.force in_mono)))
  1178             |> Thm.close_derivation
  1179           end;
  1180 
  1181         fun mk_rel_cong () =
  1182           let
  1183             val cong_prems = mk_rel_prems (curry HOLogic.mk_eq);
  1184             val cong_concl = mk_rel_concl HOLogic.mk_eq;
  1185           in
  1186             Goal.prove_sorry lthy [] []
  1187               (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (cong_prems, cong_concl)))
  1188               (fn _ => (TRY o hyp_subst_tac lthy THEN' rtac refl) 1)
  1189             |> Thm.close_derivation
  1190           end;
  1191 
  1192         val rel_mono = Lazy.lazy mk_rel_mono;
  1193         val rel_cong = Lazy.lazy mk_rel_cong;
  1194 
  1195         fun mk_rel_eq () =
  1196           Goal.prove_sorry lthy [] []
  1197             (mk_Trueprop_eq (Term.list_comb (relAsAs, map HOLogic.eq_const As'),
  1198               HOLogic.eq_const CA'))
  1199             (K (mk_rel_eq_tac live (Lazy.force rel_Grp) (Lazy.force rel_cong) (#map_id0 axioms)))
  1200           |> Thm.close_derivation;
  1201 
  1202         val rel_eq = Lazy.lazy mk_rel_eq;
  1203 
  1204         fun mk_rel_conversep () =
  1205           let
  1206             val relBsAs = mk_bnf_rel pred2RT's CB' CA';
  1207             val lhs = Term.list_comb (relBsAs, map mk_conversep Rs);
  1208             val rhs = mk_conversep (Term.list_comb (rel, Rs));
  1209             val le_goal = fold_rev Logic.all Rs (HOLogic.mk_Trueprop (mk_leq lhs rhs));
  1210             val le_thm = Goal.prove_sorry lthy [] [] le_goal
  1211               (mk_rel_conversep_le_tac rel_OO_Grps (Lazy.force rel_eq) (#map_cong0 axioms)
  1212                 (Lazy.force map_comp) (map Lazy.force set_map))
  1213               |> Thm.close_derivation
  1214             val goal = fold_rev Logic.all Rs (mk_Trueprop_eq (lhs, rhs));
  1215           in
  1216             Goal.prove_sorry lthy [] [] goal
  1217               (K (mk_rel_conversep_tac le_thm (Lazy.force rel_mono)))
  1218             |> Thm.close_derivation
  1219           end;
  1220 
  1221         val rel_conversep = Lazy.lazy mk_rel_conversep;
  1222 
  1223         fun mk_rel_OO () =
  1224           let
  1225             val relAsCs = mk_bnf_rel pred2RTsAsCs CA' CC';
  1226             val relBsCs = mk_bnf_rel pred2RTsBsCs CB' CC';
  1227             val lhs = Term.list_comb (relAsCs, map2 (curry mk_rel_compp) Rs Ss);
  1228             val rhs = mk_rel_compp (Term.list_comb (rel, Rs), Term.list_comb (relBsCs, Ss));
  1229             val goal = fold_rev Logic.all (Rs @ Ss) (mk_Trueprop_eq (lhs, rhs));
  1230           in
  1231             Goal.prove_sorry lthy [] [] goal
  1232               (mk_rel_OO_tac rel_OO_Grps (Lazy.force rel_eq) (#map_cong0 axioms)
  1233                 (Lazy.force map_wppull) (Lazy.force map_comp) (map Lazy.force set_map))
  1234             |> Thm.close_derivation
  1235           end;
  1236 
  1237         val rel_OO = Lazy.lazy mk_rel_OO;
  1238 
  1239         fun mk_in_rel () = trans OF [rel_OO_Grp, @{thm OO_Grp_alt}] RS @{thm predicate2_eqD};
  1240 
  1241         val in_rel = Lazy.lazy mk_in_rel;
  1242 
  1243         fun mk_rel_flip () =
  1244           let
  1245             val rel_conversep_thm = Lazy.force rel_conversep;
  1246             val cts = map (SOME o certify lthy) Rs;
  1247             val rel_conversep_thm' = cterm_instantiate_pos cts rel_conversep_thm;
  1248           in
  1249             unfold_thms lthy @{thms conversep_iff} (rel_conversep_thm' RS @{thm predicate2_eqD})
  1250             |> singleton (Proof_Context.export names_lthy pre_names_lthy)
  1251           end;
  1252 
  1253         val rel_flip = Lazy.lazy mk_rel_flip;
  1254 
  1255         fun mk_rel_mono_strong () =
  1256           let
  1257             fun mk_prem setA setB R S a b =
  1258               HOLogic.mk_Trueprop
  1259                 (mk_Ball (setA $ x) (Term.absfree (dest_Free a)
  1260                   (mk_Ball (setB $ y) (Term.absfree (dest_Free b)
  1261                     (HOLogic.mk_imp (R $ a $ b, S $ a $ b))))));
  1262             val prems = HOLogic.mk_Trueprop (Term.list_comb (rel, Rs) $ x $ y) :: 
  1263               map6 mk_prem bnf_sets_As bnf_sets_Bs Rs Rs_copy zs ys;
  1264             val concl = HOLogic.mk_Trueprop (Term.list_comb (rel, Rs_copy) $ x $ y);
  1265           in
  1266             Goal.prove_sorry lthy [] []
  1267               (fold_rev Logic.all (x :: y :: Rs @ Rs_copy) (Logic.list_implies (prems, concl)))
  1268               (mk_rel_mono_strong_tac (Lazy.force in_rel) (map Lazy.force set_map))
  1269             |> Thm.close_derivation
  1270           end;
  1271 
  1272         val rel_mono_strong = Lazy.lazy mk_rel_mono_strong;
  1273 
  1274         fun mk_map_transfer () =
  1275           let
  1276             val rels = map2 mk_fun_rel transfer_domRs transfer_ranRs;
  1277             val rel = mk_fun_rel
  1278               (Term.list_comb (mk_bnf_rel transfer_domRTs CA' CB1, transfer_domRs))
  1279               (Term.list_comb (mk_bnf_rel transfer_ranRTs CB' CB2, transfer_ranRs));
  1280             val concl = HOLogic.mk_Trueprop
  1281               (fold_rev mk_fun_rel rels rel $ bnf_map_AsBs $ mk_bnf_map B1Ts B2Ts);
  1282           in
  1283             Goal.prove_sorry lthy [] []
  1284               (fold_rev Logic.all (transfer_domRs @ transfer_ranRs) concl)
  1285               (mk_map_transfer_tac (Lazy.force rel_mono) (Lazy.force in_rel)
  1286                 (map Lazy.force set_map) (#map_cong0 axioms) (Lazy.force map_comp))
  1287             |> Thm.close_derivation
  1288           end;
  1289 
  1290         val map_transfer = Lazy.lazy mk_map_transfer;
  1291 
  1292         val defs = mk_defs bnf_map_def bnf_set_defs bnf_rel_def;
  1293 
  1294         val facts = mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_map in_bd in_cong
  1295           in_mono in_rel map_comp map_cong map_id map_transfer map_wppull rel_eq rel_flip set_map
  1296           rel_cong rel_mono rel_mono_strong rel_Grp rel_conversep rel_OO;
  1297 
  1298         val wits = map2 mk_witness bnf_wits wit_thms;
  1299 
  1300         val bnf_rel =
  1301           Term.subst_atomic_types ((Ds ~~ deads) @ (As' ~~ alphas) @ (Bs' ~~ betas)) rel;
  1302 
  1303         val bnf = mk_bnf bnf_b CA live alphas betas dead deads bnf_map bnf_sets bnf_bd axioms defs
  1304           facts wits bnf_rel;
  1305       in
  1306         (bnf, lthy |> note_bnf_thms fact_policy qualify bnf_b bnf)
  1307       end;
  1308 
  1309     val one_step_defs =
  1310       no_reflexive (bnf_map_def :: bnf_bd_def :: bnf_set_defs @ bnf_wit_defs @ [bnf_rel_def]);
  1311   in
  1312     (key, goals, wit_goalss, after_qed, lthy, one_step_defs)
  1313   end;
  1314 
  1315 fun register_bnf key (bnf, lthy) =
  1316   (bnf, Local_Theory.declaration {syntax = false, pervasive = true}
  1317     (fn phi => Data.map (Symtab.default (key, morph_bnf phi bnf))) lthy);
  1318 
  1319 fun bnf_def const_policy fact_policy qualify tacs wit_tac Ds map_b rel_b set_bs =
  1320   (fn (_, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, one_step_defs) =>
  1321   let
  1322     fun mk_wits_tac set_maps =
  1323       K (TRYALL Goal.conjunction_tac) THEN'
  1324       (case triv_tac_opt of
  1325         SOME tac => tac set_maps
  1326       | NONE => mk_unfold_thms_then_tac lthy one_step_defs wit_tac);
  1327     val wit_goals = map Logic.mk_conjunction_balanced wit_goalss;
  1328     fun mk_wit_thms set_maps =
  1329       Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals) (mk_wits_tac set_maps)
  1330         |> Conjunction.elim_balanced (length wit_goals)
  1331         |> map2 (Conjunction.elim_balanced o length) wit_goalss
  1332         |> map (map (Thm.close_derivation o Thm.forall_elim_vars 0));
  1333   in
  1334     map2 (Thm.close_derivation oo Goal.prove_sorry lthy [] [])
  1335       goals (map (mk_unfold_thms_then_tac lthy one_step_defs) tacs)
  1336     |> (fn thms => after_qed mk_wit_thms (map single thms) lthy)
  1337   end) oo prepare_def const_policy fact_policy qualify (K I) Ds map_b rel_b set_bs;
  1338 
  1339 val bnf_cmd = (fn (key, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, defs) =>
  1340   let
  1341     val wit_goals = map Logic.mk_conjunction_balanced wit_goalss;
  1342     fun mk_triv_wit_thms tac set_maps =
  1343       Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals)
  1344         (K (TRYALL Goal.conjunction_tac) THEN' tac set_maps)
  1345         |> Conjunction.elim_balanced (length wit_goals)
  1346         |> map2 (Conjunction.elim_balanced o length) wit_goalss
  1347         |> map (map (Thm.close_derivation o Thm.forall_elim_vars 0));
  1348     val (mk_wit_thms, nontriv_wit_goals) = 
  1349       (case triv_tac_opt of
  1350         NONE => (fn _ => [], map (map (rpair [])) wit_goalss)
  1351       | SOME tac => (mk_triv_wit_thms tac, []));
  1352   in
  1353     Proof.unfolding ([[(defs, [])]])
  1354       (Proof.theorem NONE (snd o register_bnf key oo after_qed mk_wit_thms)
  1355         (map (single o rpair []) goals @ nontriv_wit_goals) lthy)
  1356   end) oo prepare_def Do_Inline (user_policy Note_Some) I Syntax.read_term NONE
  1357     Binding.empty Binding.empty [];
  1358 
  1359 fun print_bnfs ctxt =
  1360   let
  1361     fun pretty_set sets i = Pretty.block
  1362       [Pretty.str (mk_setN (i + 1) ^ ":"), Pretty.brk 1,
  1363           Pretty.quote (Syntax.pretty_term ctxt (nth sets i))];
  1364 
  1365     fun pretty_bnf (key, BNF {T = T, map = map, sets = sets, bd = bd,
  1366       live = live, lives = lives, dead = dead, deads = deads, ...}) =
  1367       Pretty.big_list
  1368         (Pretty.string_of (Pretty.block [Pretty.str key, Pretty.str ":", Pretty.brk 1,
  1369           Pretty.quote (Syntax.pretty_typ ctxt T)]))
  1370         ([Pretty.block [Pretty.str "live:", Pretty.brk 1, Pretty.str (string_of_int live),
  1371             Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) lives)],
  1372           Pretty.block [Pretty.str "dead:", Pretty.brk 1, Pretty.str (string_of_int dead),
  1373             Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) deads)],
  1374           Pretty.block [Pretty.str (mapN ^ ":"), Pretty.brk 1,
  1375             Pretty.quote (Syntax.pretty_term ctxt map)]] @
  1376           List.map (pretty_set sets) (0 upto length sets - 1) @
  1377           [Pretty.block [Pretty.str (bdN ^ ":"), Pretty.brk 1,
  1378             Pretty.quote (Syntax.pretty_term ctxt bd)]]);
  1379   in
  1380     Pretty.big_list "BNFs:" (map pretty_bnf (Symtab.dest (Data.get (Context.Proof ctxt))))
  1381     |> Pretty.writeln
  1382   end;
  1383 
  1384 val _ =
  1385   Outer_Syntax.improper_command @{command_spec "print_bnfs"}
  1386     "print all bounded natural functors"
  1387     (Scan.succeed (Toplevel.keep (print_bnfs o Toplevel.context_of)));
  1388 
  1389 val _ =
  1390   Outer_Syntax.local_theory_to_proof @{command_spec "bnf"}
  1391     "register a type as a bounded natural functor"
  1392     ((parse_opt_binding_colon -- Parse.term --
  1393        (@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}) -- Parse.term --
  1394        (Scan.option ((@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}))
  1395          >> the_default []) --
  1396        Scan.option Parse.term)
  1397        >> bnf_cmd);
  1398 
  1399 end;