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