1(*  Title:      HOL/Tools/BNF/bnf_fp_util.ML
2    Author:     Dmitriy Traytel, TU Muenchen
3    Author:     Jasmin Blanchette, TU Muenchen
4    Author:     Martin Desharnais, TU Muenchen
5    Author:     Stefan Berghofer, TU Muenchen
6    Copyright   2012, 2013, 2014
7
8Shared library for the datatype and codatatype constructions.
9*)
10
11signature BNF_FP_UTIL =
12sig
13  exception EMPTY_DATATYPE of string
14
15  type fp_result =
16    {Ts: typ list,
17     bnfs: BNF_Def.bnf list,
18     pre_bnfs: BNF_Def.bnf list,
19     absT_infos: BNF_Comp.absT_info list,
20     ctors: term list,
21     dtors: term list,
22     xtor_un_folds: term list,
23     xtor_co_recs: term list,
24     xtor_co_induct: thm,
25     dtor_ctors: thm list,
26     ctor_dtors: thm list,
27     ctor_injects: thm list,
28     dtor_injects: thm list,
29     xtor_maps: thm list,
30     xtor_map_unique: thm,
31     xtor_setss: thm list list,
32     xtor_rels: thm list,
33     xtor_un_fold_thms: thm list,
34     xtor_co_rec_thms: thm list,
35     xtor_un_fold_unique: thm,
36     xtor_co_rec_unique: thm,
37     xtor_un_fold_o_maps: thm list,
38     xtor_co_rec_o_maps: thm list,
39     xtor_un_fold_transfers: thm list,
40     xtor_co_rec_transfers: thm list,
41     xtor_rel_co_induct: thm,
42     dtor_set_inducts: thm list}
43
44  val morph_fp_result: morphism -> fp_result -> fp_result
45
46  val time: Proof.context -> Timer.real_timer -> string -> Timer.real_timer
47
48  val fixpoint: ('a * 'a -> bool) -> ('a list -> 'a list) -> 'a list -> 'a list
49
50  val IITN: string
51  val LevN: string
52  val algN: string
53  val behN: string
54  val bisN: string
55  val carTN: string
56  val caseN: string
57  val coN: string
58  val coinductN: string
59  val coinduct_strongN: string
60  val corecN: string
61  val corec_discN: string
62  val corec_disc_iffN: string
63  val ctorN: string
64  val ctor_dtorN: string
65  val ctor_exhaustN: string
66  val ctor_induct2N: string
67  val ctor_inductN: string
68  val ctor_injectN: string
69  val ctor_foldN: string
70  val ctor_fold_o_mapN: string
71  val ctor_fold_transferN: string
72  val ctor_fold_uniqueN: string
73  val ctor_mapN: string
74  val ctor_map_uniqueN: string
75  val ctor_recN: string
76  val ctor_rec_o_mapN: string
77  val ctor_rec_transferN: string
78  val ctor_rec_uniqueN: string
79  val ctor_relN: string
80  val ctor_rel_inductN: string
81  val ctor_set_inclN: string
82  val ctor_set_set_inclN: string
83  val dtorN: string
84  val dtor_coinductN: string
85  val dtor_corecN: string
86  val dtor_corec_o_mapN: string
87  val dtor_corec_transferN: string
88  val dtor_corec_uniqueN: string
89  val dtor_ctorN: string
90  val dtor_exhaustN: string
91  val dtor_injectN: string
92  val dtor_mapN: string
93  val dtor_map_coinductN: string
94  val dtor_map_coinduct_strongN: string
95  val dtor_map_uniqueN: string
96  val dtor_relN: string
97  val dtor_rel_coinductN: string
98  val dtor_set_inclN: string
99  val dtor_set_set_inclN: string
100  val dtor_coinduct_strongN: string
101  val dtor_unfoldN: string
102  val dtor_unfold_o_mapN: string
103  val dtor_unfold_transferN: string
104  val dtor_unfold_uniqueN: string
105  val exhaustN: string
106  val colN: string
107  val inductN: string
108  val injectN: string
109  val isNodeN: string
110  val lsbisN: string
111  val mapN: string
112  val map_uniqueN: string
113  val min_algN: string
114  val morN: string
115  val nchotomyN: string
116  val recN: string
117  val rel_casesN: string
118  val rel_coinductN: string
119  val rel_inductN: string
120  val rel_injectN: string
121  val rel_introsN: string
122  val rel_distinctN: string
123  val rel_selN: string
124  val rvN: string
125  val corec_selN: string
126  val set_inclN: string
127  val set_set_inclN: string
128  val setN: string
129  val simpsN: string
130  val strTN: string
131  val str_initN: string
132  val sum_bdN: string
133  val sum_bdTN: string
134  val uniqueN: string
135
136  (* TODO: Don't index set facts. Isabelle packages traditionally generate uniform names. *)
137  val mk_ctor_setN: int -> string
138  val mk_dtor_setN: int -> string
139  val mk_dtor_set_inductN: int -> string
140  val mk_set_inductN: int -> string
141
142  val co_prefix: BNF_Util.fp_kind -> string
143
144  val split_conj_thm: thm -> thm list
145  val split_conj_prems: int -> thm -> thm
146
147  val mk_sumTN: typ list -> typ
148  val mk_sumTN_balanced: typ list -> typ
149  val mk_tupleT_balanced: typ list -> typ
150  val mk_sumprodT_balanced: typ list list -> typ
151
152  val mk_proj: typ -> int -> int -> term
153
154  val mk_convol: term * term -> term
155  val mk_rel_prod: term -> term -> term
156  val mk_rel_sum: term -> term -> term
157
158  val Inl_const: typ -> typ -> term
159  val Inr_const: typ -> typ -> term
160  val mk_tuple_balanced: term list -> term
161  val mk_tuple1_balanced: typ list -> term list -> term
162  val abs_curried_balanced: typ list -> term -> term
163  val mk_tupled_fun: term -> term -> term list -> term
164
165  val mk_case_sum: term * term -> term
166  val mk_case_sumN: term list -> term
167  val mk_case_absumprod: typ -> term -> term list -> term list list -> term list list -> term
168
169  val mk_Inl: typ -> term -> term
170  val mk_Inr: typ -> term -> term
171  val mk_sumprod_balanced: typ -> int -> int -> term list -> term
172  val mk_absumprod: typ -> term -> int -> int -> term list -> term
173
174  val dest_sumT: typ -> typ * typ
175  val dest_sumTN_balanced: int -> typ -> typ list
176  val dest_tupleT_balanced: int -> typ -> typ list
177  val dest_absumprodT: typ -> typ -> int -> int list -> typ -> typ list list
178
179  val If_const: typ -> term
180
181  val mk_Field: term -> term
182  val mk_If: term -> term -> term -> term
183
184  val mk_absumprodE: thm -> int list -> thm
185
186  val mk_sum_caseN: int -> int -> thm
187  val mk_sum_caseN_balanced: int -> int -> thm
188
189  val mk_sum_Cinfinite: thm list -> thm
190  val mk_sum_card_order: thm list -> thm
191
192  val force_typ: Proof.context -> typ -> term -> term
193
194  val mk_xtor_rel_co_induct_thm: BNF_Util.fp_kind -> term list -> term list -> term list ->
195    term list -> term list -> term list -> term list -> term list ->
196    ({prems: thm list, context: Proof.context} -> tactic) -> Proof.context -> thm
197  val mk_xtor_co_iter_transfer_thms: BNF_Util.fp_kind -> term list -> term list -> term list ->
198    term list -> term list -> term list -> term list ->
199    ({prems: thm list, context: Proof.context} -> tactic) -> Proof.context -> thm list
200  val mk_xtor_co_iter_o_map_thms: BNF_Util.fp_kind -> bool -> int -> thm -> thm list -> thm list ->
201    thm list -> thm list -> thm list
202  val derive_xtor_co_recs: BNF_Util.fp_kind -> binding list -> (typ list -> typ list) ->
203    (typ list list * typ list) -> BNF_Def.bnf list -> term list -> term list ->
204    thm -> thm list -> thm list -> thm list -> thm list ->
205    (BNF_Comp.absT_info * BNF_Comp.absT_info) option list ->
206    local_theory ->
207    (term list * (thm list * thm * thm list * thm list)) * local_theory
208  val raw_qualify: (binding -> binding) -> binding -> binding -> binding
209
210  val fixpoint_bnf: bool -> (binding -> binding) ->
211      (binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list ->
212      BNF_Comp.absT_info list -> local_theory -> 'a) ->
213    binding list -> (string * sort) list -> (string * sort) list -> (string * sort) list ->
214    typ list -> BNF_Comp.comp_cache -> local_theory ->
215    ((BNF_Def.bnf list * BNF_Comp.absT_info list) * BNF_Comp.comp_cache) * 'a
216end;
217
218structure BNF_FP_Util : BNF_FP_UTIL =
219struct
220
221open Ctr_Sugar
222open BNF_Comp
223open BNF_Def
224open BNF_Util
225open BNF_FP_Util_Tactics
226
227exception EMPTY_DATATYPE of string;
228
229type fp_result =
230  {Ts: typ list,
231   bnfs: bnf list,
232   pre_bnfs: BNF_Def.bnf list,
233   absT_infos: BNF_Comp.absT_info list,
234   ctors: term list,
235   dtors: term list,
236   xtor_un_folds: term list,
237   xtor_co_recs: term list,
238   xtor_co_induct: thm,
239   dtor_ctors: thm list,
240   ctor_dtors: thm list,
241   ctor_injects: thm list,
242   dtor_injects: thm list,
243   xtor_maps: thm list,
244   xtor_map_unique: thm,
245   xtor_setss: thm list list,
246   xtor_rels: thm list,
247   xtor_un_fold_thms: thm list,
248   xtor_co_rec_thms: thm list,
249   xtor_un_fold_unique: thm,
250   xtor_co_rec_unique: thm,
251   xtor_un_fold_o_maps: thm list,
252   xtor_co_rec_o_maps: thm list,
253   xtor_un_fold_transfers: thm list,
254   xtor_co_rec_transfers: thm list,
255   xtor_rel_co_induct: thm,
256   dtor_set_inducts: thm list};
257
258fun morph_fp_result phi {Ts, bnfs, pre_bnfs, absT_infos, ctors, dtors, xtor_un_folds,
259    xtor_co_recs, xtor_co_induct, dtor_ctors, ctor_dtors, ctor_injects, dtor_injects, xtor_maps,
260    xtor_map_unique, xtor_setss, xtor_rels, xtor_un_fold_thms, xtor_co_rec_thms,
261    xtor_un_fold_unique, xtor_co_rec_unique, xtor_un_fold_o_maps,
262    xtor_co_rec_o_maps, xtor_un_fold_transfers, xtor_co_rec_transfers, xtor_rel_co_induct,
263    dtor_set_inducts} =
264  {Ts = map (Morphism.typ phi) Ts,
265   bnfs = map (morph_bnf phi) bnfs,
266   pre_bnfs = map (morph_bnf phi) pre_bnfs,
267   absT_infos = map (morph_absT_info phi) absT_infos,
268   ctors = map (Morphism.term phi) ctors,
269   dtors = map (Morphism.term phi) dtors,
270   xtor_un_folds = map (Morphism.term phi) xtor_un_folds,
271   xtor_co_recs = map (Morphism.term phi) xtor_co_recs,
272   xtor_co_induct = Morphism.thm phi xtor_co_induct,
273   dtor_ctors = map (Morphism.thm phi) dtor_ctors,
274   ctor_dtors = map (Morphism.thm phi) ctor_dtors,
275   ctor_injects = map (Morphism.thm phi) ctor_injects,
276   dtor_injects = map (Morphism.thm phi) dtor_injects,
277   xtor_maps = map (Morphism.thm phi) xtor_maps,
278   xtor_map_unique = Morphism.thm phi xtor_map_unique,
279   xtor_setss = map (map (Morphism.thm phi)) xtor_setss,
280   xtor_rels = map (Morphism.thm phi) xtor_rels,
281   xtor_un_fold_thms = map (Morphism.thm phi) xtor_un_fold_thms,
282   xtor_co_rec_thms = map (Morphism.thm phi) xtor_co_rec_thms,
283   xtor_un_fold_unique = Morphism.thm phi xtor_un_fold_unique,
284   xtor_co_rec_unique = Morphism.thm phi xtor_co_rec_unique,
285   xtor_un_fold_o_maps = map (Morphism.thm phi) xtor_un_fold_o_maps,
286   xtor_co_rec_o_maps = map (Morphism.thm phi) xtor_co_rec_o_maps,
287   xtor_un_fold_transfers = map (Morphism.thm phi) xtor_un_fold_transfers,
288   xtor_co_rec_transfers = map (Morphism.thm phi) xtor_co_rec_transfers,
289   xtor_rel_co_induct = Morphism.thm phi xtor_rel_co_induct,
290   dtor_set_inducts = map (Morphism.thm phi) dtor_set_inducts};
291
292fun time ctxt timer msg = (if Config.get ctxt bnf_timing
293  then warning (msg ^ ": " ^ string_of_int (Time.toMilliseconds (Timer.checkRealTimer timer)) ^
294    " ms")
295  else (); Timer.startRealTimer ());
296
297val preN = "pre_"
298val rawN = "raw_"
299
300val coN = "co"
301val unN = "un"
302val algN = "alg"
303val IITN = "IITN"
304val foldN = "fold"
305val unfoldN = unN ^ foldN
306val uniqueN = "unique"
307val transferN = "transfer"
308val simpsN = "simps"
309val ctorN = "ctor"
310val dtorN = "dtor"
311val ctor_foldN = ctorN ^ "_" ^ foldN
312val dtor_unfoldN = dtorN ^ "_" ^ unfoldN
313val ctor_fold_uniqueN = ctor_foldN ^ "_" ^ uniqueN
314val ctor_fold_o_mapN = ctor_foldN ^ "_o_" ^ mapN
315val dtor_unfold_uniqueN = dtor_unfoldN ^ "_" ^ uniqueN
316val dtor_unfold_o_mapN = dtor_unfoldN ^ "_o_" ^ mapN
317val ctor_fold_transferN = ctor_foldN ^ "_" ^ transferN
318val dtor_unfold_transferN = dtor_unfoldN ^ "_" ^ transferN
319val ctor_mapN = ctorN ^ "_" ^ mapN
320val dtor_mapN = dtorN ^ "_" ^ mapN
321val map_uniqueN = mapN ^ "_" ^ uniqueN
322val ctor_map_uniqueN = ctorN ^ "_" ^ map_uniqueN
323val dtor_map_uniqueN = dtorN ^ "_" ^ map_uniqueN
324val min_algN = "min_alg"
325val morN = "mor"
326val bisN = "bis"
327val lsbisN = "lsbis"
328val sum_bdTN = "sbdT"
329val sum_bdN = "sbd"
330val carTN = "carT"
331val strTN = "strT"
332val isNodeN = "isNode"
333val LevN = "Lev"
334val rvN = "recover"
335val behN = "beh"
336val setN = "set"
337val mk_ctor_setN = prefix (ctorN ^ "_") o mk_setN
338val mk_dtor_setN = prefix (dtorN ^ "_") o mk_setN
339fun mk_set_inductN i = mk_setN i ^ "_induct"
340val mk_dtor_set_inductN = prefix (dtorN ^ "_") o mk_set_inductN
341
342val str_initN = "str_init"
343val recN = "rec"
344val corecN = coN ^ recN
345val ctor_recN = ctorN ^ "_" ^ recN
346val ctor_rec_o_mapN = ctor_recN ^ "_o_" ^ mapN
347val ctor_rec_transferN = ctor_recN ^ "_" ^ transferN
348val ctor_rec_uniqueN = ctor_recN ^ "_" ^ uniqueN
349val dtor_corecN = dtorN ^ "_" ^ corecN
350val dtor_corec_o_mapN = dtor_corecN ^ "_o_" ^ mapN
351val dtor_corec_transferN = dtor_corecN ^ "_" ^ transferN
352val dtor_corec_uniqueN = dtor_corecN ^ "_" ^ uniqueN
353
354val ctor_dtorN = ctorN ^ "_" ^ dtorN
355val dtor_ctorN = dtorN ^ "_" ^ ctorN
356val nchotomyN = "nchotomy"
357val injectN = "inject"
358val exhaustN = "exhaust"
359val ctor_injectN = ctorN ^ "_" ^ injectN
360val ctor_exhaustN = ctorN ^ "_" ^ exhaustN
361val dtor_injectN = dtorN ^ "_" ^ injectN
362val dtor_exhaustN = dtorN ^ "_" ^ exhaustN
363val ctor_relN = ctorN ^ "_" ^ relN
364val dtor_relN = dtorN ^ "_" ^ relN
365val inductN = "induct"
366val coinductN = coN ^ inductN
367val ctor_inductN = ctorN ^ "_" ^ inductN
368val ctor_induct2N = ctor_inductN ^ "2"
369val dtor_map_coinductN = dtor_mapN ^ "_" ^ coinductN
370val dtor_coinductN = dtorN ^ "_" ^ coinductN
371val coinduct_strongN = coinductN ^ "_strong"
372val dtor_map_coinduct_strongN = dtor_mapN ^ "_" ^ coinduct_strongN
373val dtor_coinduct_strongN = dtorN ^ "_" ^ coinduct_strongN
374val colN = "col"
375val set_inclN = "set_incl"
376val ctor_set_inclN = ctorN ^ "_" ^ set_inclN
377val dtor_set_inclN = dtorN ^ "_" ^ set_inclN
378val set_set_inclN = "set_set_incl"
379val ctor_set_set_inclN = ctorN ^ "_" ^ set_set_inclN
380val dtor_set_set_inclN = dtorN ^ "_" ^ set_set_inclN
381
382val caseN = "case"
383val discN = "disc"
384val corec_discN = corecN ^ "_" ^ discN
385val iffN = "_iff"
386val corec_disc_iffN = corec_discN ^ iffN
387val distinctN = "distinct"
388val rel_distinctN = relN ^ "_" ^ distinctN
389val injectN = "inject"
390val rel_casesN = relN ^ "_cases"
391val rel_injectN = relN ^ "_" ^ injectN
392val rel_introsN = relN ^ "_intros"
393val rel_coinductN = relN ^ "_" ^ coinductN
394val rel_selN = relN ^ "_sel"
395val dtor_rel_coinductN = dtorN ^ "_" ^ rel_coinductN
396val rel_inductN = relN ^ "_" ^ inductN
397val ctor_rel_inductN = ctorN ^ "_" ^ rel_inductN
398val selN = "sel"
399val corec_selN = corecN ^ "_" ^ selN
400
401fun co_prefix fp = case_fp fp "" "co";
402
403fun dest_sumT (Type (\<^type_name>\<open>sum\<close>, [T, T'])) = (T, T');
404
405val dest_sumTN_balanced = Balanced_Tree.dest dest_sumT;
406
407fun dest_tupleT_balanced 0 \<^typ>\<open>unit\<close> = []
408  | dest_tupleT_balanced n T = Balanced_Tree.dest HOLogic.dest_prodT n T;
409
410fun dest_absumprodT absT repT n ms =
411  map2 dest_tupleT_balanced ms o dest_sumTN_balanced n o mk_repT absT repT;
412
413val mk_sumTN = Library.foldr1 mk_sumT;
414val mk_sumTN_balanced = Balanced_Tree.make mk_sumT;
415
416fun mk_tupleT_balanced [] = HOLogic.unitT
417  | mk_tupleT_balanced Ts = Balanced_Tree.make HOLogic.mk_prodT Ts;
418
419val mk_sumprodT_balanced = mk_sumTN_balanced o map mk_tupleT_balanced;
420
421fun mk_proj T n k =
422  let val (binders, _) = strip_typeN n T in
423    fold_rev (fn T => fn t => Abs (Name.uu, T, t)) binders (Bound (n - k - 1))
424  end;
425
426fun mk_convol (f, g) =
427  let
428    val (fU, fTU) = `range_type (fastype_of f);
429    val ((gT, gU), gTU) = `dest_funT (fastype_of g);
430    val convolT = fTU --> gTU --> gT --> HOLogic.mk_prodT (fU, gU);
431  in Const (\<^const_name>\<open>convol\<close>, convolT) $ f $ g end;
432
433fun mk_rel_prod R S =
434  let
435    val ((A1, A2), RT) = `dest_pred2T (fastype_of R);
436    val ((B1, B2), ST) = `dest_pred2T (fastype_of S);
437    val rel_prodT = RT --> ST --> mk_pred2T (HOLogic.mk_prodT (A1, B1)) (HOLogic.mk_prodT (A2, B2));
438  in Const (\<^const_name>\<open>rel_prod\<close>, rel_prodT) $ R $ S end;
439
440fun mk_rel_sum R S =
441  let
442    val ((A1, A2), RT) = `dest_pred2T (fastype_of R);
443    val ((B1, B2), ST) = `dest_pred2T (fastype_of S);
444    val rel_sumT = RT --> ST --> mk_pred2T (mk_sumT (A1, B1)) (mk_sumT (A2, B2));
445  in Const (\<^const_name>\<open>rel_sum\<close>, rel_sumT) $ R $ S end;
446
447fun Inl_const LT RT = Const (\<^const_name>\<open>Inl\<close>, LT --> mk_sumT (LT, RT));
448fun mk_Inl RT t = Inl_const (fastype_of t) RT $ t;
449
450fun Inr_const LT RT = Const (\<^const_name>\<open>Inr\<close>, RT --> mk_sumT (LT, RT));
451fun mk_Inr LT t = Inr_const LT (fastype_of t) $ t;
452
453fun mk_prod1 bound_Ts (t, u) =
454  HOLogic.pair_const (fastype_of1 (bound_Ts, t)) (fastype_of1 (bound_Ts, u)) $ t $ u;
455
456fun mk_tuple1_balanced _ [] = HOLogic.unit
457  | mk_tuple1_balanced bound_Ts ts = Balanced_Tree.make (mk_prod1 bound_Ts) ts;
458
459val mk_tuple_balanced = mk_tuple1_balanced [];
460
461fun abs_curried_balanced Ts t =
462  t $ mk_tuple1_balanced (List.rev Ts) (map Bound (length Ts - 1 downto 0))
463  |> fold_rev (Term.abs o pair Name.uu) Ts;
464
465fun mk_sumprod_balanced T n k ts = Sum_Tree.mk_inj T n k (mk_tuple_balanced ts);
466
467fun mk_absumprod absT abs0 n k ts =
468  let val abs = mk_abs absT abs0;
469  in abs $ mk_sumprod_balanced (domain_type (fastype_of abs)) n k ts end;
470
471fun mk_case_sum (f, g) =
472  let
473    val (fT, T') = dest_funT (fastype_of f);
474    val (gT, _) = dest_funT (fastype_of g);
475  in
476    Sum_Tree.mk_sumcase fT gT T' f g
477  end;
478
479val mk_case_sumN = Library.foldr1 mk_case_sum;
480val mk_case_sumN_balanced = Balanced_Tree.make mk_case_sum;
481
482fun mk_tupled_fun f x xs =
483  if xs = [x] then f else HOLogic.tupled_lambda x (Term.list_comb (f, xs));
484
485fun mk_case_absumprod absT rep fs xss xss' =
486  HOLogic.mk_comp (mk_case_sumN_balanced
487    (@{map 3} mk_tupled_fun fs (map mk_tuple_balanced xss) xss'), mk_rep absT rep);
488
489fun If_const T = Const (\<^const_name>\<open>If\<close>, HOLogic.boolT --> T --> T --> T);
490fun mk_If p t f = let val T = fastype_of t in If_const T $ p $ t $ f end;
491
492fun mk_Field r =
493  let val T = fst (dest_relT (fastype_of r));
494  in Const (\<^const_name>\<open>Field\<close>, mk_relT (T, T) --> HOLogic.mk_setT T) $ r end;
495
496(*dangerous; use with monotonic, converging functions only!*)
497fun fixpoint eq f X = if subset eq (f X, X) then X else fixpoint eq f (f X);
498
499(* stolen from "~~/src/HOL/Tools/Datatype/datatype_aux.ML" *)
500fun split_conj_thm th =
501  ((th RS conjunct1) :: split_conj_thm (th RS conjunct2)) handle THM _ => [th];
502
503fun split_conj_prems limit th =
504  let
505    fun split n i th =
506      if i = n then th else split n (i + 1) (conjI RSN (i, th)) handle THM _ => th;
507  in split limit 1 th end;
508
509fun mk_obj_sumEN_balanced n =
510  Balanced_Tree.make (fn (thm1, thm2) => thm1 RSN (1, thm2 RSN (2, @{thm obj_sumE_f})))
511    (replicate n asm_rl);
512
513fun mk_tupled_allIN_balanced 0 = @{thm unit_all_impI}
514  | mk_tupled_allIN_balanced n =
515    let
516      val (tfrees, _) = BNF_Util.mk_TFrees n \<^context>;
517      val T = mk_tupleT_balanced tfrees;
518    in
519      @{thm asm_rl[of "\<forall>x. P x \<longrightarrow> Q x" for P Q]}
520      |> Thm.instantiate' [SOME (Thm.ctyp_of \<^context> T)] []
521      |> Raw_Simplifier.rewrite_goals_rule \<^context> @{thms split_paired_All[THEN eq_reflection]}
522      |> (fn thm => impI RS funpow n (fn th => allI RS th) thm)
523      |> Thm.varifyT_global
524    end;
525
526fun mk_absumprodE type_definition ms =
527  let val n = length ms in
528    mk_obj_sumEN_balanced n OF map mk_tupled_allIN_balanced ms RS
529      (type_definition RS @{thm type_copy_obj_one_point_absE})
530  end;
531
532fun mk_sum_caseN 1 1 = refl
533  | mk_sum_caseN _ 1 = @{thm sum.case(1)}
534  | mk_sum_caseN 2 2 = @{thm sum.case(2)}
535  | mk_sum_caseN n k = trans OF [@{thm case_sum_step(2)}, mk_sum_caseN (n - 1) (k - 1)];
536
537fun mk_sum_step base step thm =
538  if Thm.eq_thm_prop (thm, refl) then base else trans OF [step, thm];
539
540fun mk_sum_caseN_balanced 1 1 = refl
541  | mk_sum_caseN_balanced n k =
542    Balanced_Tree.access {left = mk_sum_step @{thm sum.case(1)} @{thm case_sum_step(1)},
543      right = mk_sum_step @{thm sum.case(2)} @{thm case_sum_step(2)}, init = refl} n k;
544
545fun mk_sum_Cinfinite [thm] = thm
546  | mk_sum_Cinfinite (thm :: thms) = @{thm Cinfinite_csum_weak} OF [thm, mk_sum_Cinfinite thms];
547
548fun mk_sum_card_order [thm] = thm
549  | mk_sum_card_order (thm :: thms) = @{thm card_order_csum} OF [thm, mk_sum_card_order thms];
550
551fun mk_xtor_rel_co_induct_thm fp pre_rels pre_phis rels phis xs ys xtors xtor's tac lthy =
552  let
553    val pre_relphis = map (fn rel => Term.list_comb (rel, phis @ pre_phis)) pre_rels;
554    val relphis = map (fn rel => Term.list_comb (rel, phis)) rels;
555    fun mk_xtor fp' xtor x = if fp = fp' then xtor $ x else x;
556    val dtor = mk_xtor Greatest_FP;
557    val ctor = mk_xtor Least_FP;
558    fun flip f x y = if fp = Greatest_FP then f y x else f x y;
559
560    fun mk_prem pre_relphi phi x y xtor xtor' =
561      HOLogic.mk_Trueprop (list_all_free [x, y] (flip (curry HOLogic.mk_imp)
562        (pre_relphi $ (dtor xtor x) $ (dtor xtor' y)) (phi $ (ctor xtor x) $ (ctor xtor' y))));
563    val prems = @{map 6} mk_prem pre_relphis pre_phis xs ys xtors xtor's;
564
565    val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
566      (map2 (flip mk_leq) relphis pre_phis));
567  in
568    Goal.prove_sorry lthy (map (fst o dest_Free) (phis @ pre_phis)) prems concl tac
569    |> Thm.close_derivation \<^here>
570    |> (fn thm => thm OF (replicate (length pre_rels) @{thm allI[OF allI[OF impI]]}))
571  end;
572
573fun mk_xtor_co_iter_transfer_thms fp pre_rels pre_iphis pre_ophis rels phis un_folds un_folds' tac lthy =
574  let
575    val pre_relphis = map (fn rel => Term.list_comb (rel, phis @ pre_iphis)) pre_rels;
576    val relphis = map (fn rel => Term.list_comb (rel, phis)) rels;
577    fun flip f x y = if fp = Greatest_FP then f y x else f x y;
578
579    val arg_rels = map2 (flip mk_rel_fun) pre_relphis pre_ophis;
580    fun mk_transfer relphi pre_phi un_fold un_fold' =
581      fold_rev mk_rel_fun arg_rels (flip mk_rel_fun relphi pre_phi) $ un_fold $ un_fold';
582    val transfers = @{map 4} mk_transfer relphis pre_ophis un_folds un_folds';
583
584    val goal = fold_rev Logic.all (phis @ pre_ophis)
585      (HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj transfers));
586  in
587    Goal.prove_sorry lthy [] [] goal tac
588    |> Thm.close_derivation \<^here>
589    |> split_conj_thm
590  end;
591
592fun mk_xtor_co_iter_o_map_thms fp is_rec m un_fold_unique xtor_maps xtor_un_folds sym_map_comps
593    map_cong0s =
594  let
595    val n = length sym_map_comps;
596    val rewrite_comp_comp2 = case_fp fp @{thm rewriteR_comp_comp2} @{thm rewriteL_comp_comp2};
597    val rewrite_comp_comp = case_fp fp @{thm rewriteR_comp_comp} @{thm rewriteL_comp_comp};
598    val map_cong_passive_args1 = replicate m (case_fp fp @{thm id_comp} @{thm comp_id} RS fun_cong);
599    val map_cong_active_args1 = replicate n (if is_rec
600      then case_fp fp @{thm convol_o} @{thm o_case_sum} RS fun_cong
601      else refl);
602    val map_cong_passive_args2 = replicate m (case_fp fp @{thm comp_id} @{thm id_comp} RS fun_cong);
603    val map_cong_active_args2 = replicate n (if is_rec
604      then case_fp fp @{thm map_prod_o_convol_id} @{thm case_sum_o_map_sum_id}
605      else case_fp fp @{thm id_comp} @{thm comp_id} RS fun_cong);
606    fun mk_map_congs passive active =
607      map (fn thm => thm OF (passive @ active) RS @{thm ext}) map_cong0s;
608    val map_cong1s = mk_map_congs map_cong_passive_args1 map_cong_active_args1;
609    val map_cong2s = mk_map_congs map_cong_passive_args2 map_cong_active_args2;
610
611    fun mk_rewrites map_congs = map2 (fn sym_map_comp => fn map_cong =>
612      mk_trans sym_map_comp map_cong RS rewrite_comp_comp) sym_map_comps map_congs;
613    val rewrite1s = mk_rewrites map_cong1s;
614    val rewrite2s = mk_rewrites map_cong2s;
615    val unique_prems =
616      @{map 4} (fn xtor_map => fn un_fold => fn rewrite1 => fn rewrite2 =>
617        mk_trans (rewrite_comp_comp2 OF [xtor_map, un_fold])
618          (mk_trans rewrite1 (mk_sym rewrite2)))
619      xtor_maps xtor_un_folds rewrite1s rewrite2s;
620  in
621    split_conj_thm (un_fold_unique OF map (case_fp fp I mk_sym) unique_prems)
622  end;
623
624fun force_typ ctxt T =
625  Term.map_types Type_Infer.paramify_vars
626  #> Type.constraint T
627  #> Syntax.check_term ctxt
628  #> singleton (Variable.polymorphic ctxt);
629
630fun absT_info_encodeT thy (SOME (src : absT_info, dst : absT_info)) src_absT =
631    let
632      val src_repT = mk_repT (#absT src) (#repT src) src_absT;
633      val dst_absT = mk_absT thy (#repT dst) (#absT dst) src_repT;
634    in
635      dst_absT
636    end
637  | absT_info_encodeT _ NONE T = T;
638
639fun absT_info_decodeT thy = absT_info_encodeT thy o Option.map swap;
640
641fun absT_info_encode thy fp (opt as SOME (src : absT_info, dst : absT_info)) t =
642    let
643      val co_alg_funT = case_fp fp domain_type range_type;
644      fun co_swap pair = case_fp fp I swap pair;
645      val mk_co_comp = curry (HOLogic.mk_comp o co_swap);
646      val mk_co_abs = case_fp fp mk_abs mk_rep;
647      val mk_co_rep = case_fp fp mk_rep mk_abs;
648      val co_abs = case_fp fp #abs #rep;
649      val co_rep = case_fp fp #rep #abs;
650      val src_absT = co_alg_funT (fastype_of t);
651      val dst_absT = absT_info_encodeT thy opt src_absT;
652      val co_src_abs = mk_co_abs src_absT (co_abs src);
653      val co_dst_rep = mk_co_rep dst_absT (co_rep dst);
654    in
655      mk_co_comp (mk_co_comp t co_src_abs) co_dst_rep
656    end
657  | absT_info_encode _ _ NONE t = t;
658
659fun absT_info_decode thy fp = absT_info_encode thy fp o Option.map swap;
660
661fun mk_xtor_un_fold_xtor_thms ctxt fp un_folds xtors xtor_un_fold_unique map_id0s
662    absT_info_opts =
663  let
664    val thy = Proof_Context.theory_of ctxt;
665    fun mk_goal un_fold =
666      let
667        val rhs = list_comb (un_fold, @{map 2} (absT_info_encode thy fp) absT_info_opts xtors);
668        val T = range_type (fastype_of rhs);
669      in
670        HOLogic.mk_eq (HOLogic.id_const T, rhs)
671      end;
672    val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map mk_goal un_folds));
673    fun mk_inverses NONE = []
674      | mk_inverses (SOME (src, dst)) =
675        [#type_definition dst RS @{thm type_definition.Abs_inverse[OF _ UNIV_I]},
676         #type_definition src RS @{thm type_definition.Rep_inverse}];
677    val inverses = maps mk_inverses absT_info_opts;
678  in
679    Goal.prove_sorry ctxt [] [] goal (fn {context = ctxt, prems = _} =>
680      mk_xtor_un_fold_xtor_tac ctxt xtor_un_fold_unique map_id0s inverses)
681    |> split_conj_thm |> map mk_sym
682  end;
683
684fun derive_xtor_co_recs fp bs mk_Ts (Dss, resDs) pre_bnfs xtors0 un_folds0
685    xtor_un_fold_unique xtor_un_folds xtor_un_fold_transfers xtor_maps xtor_rels
686    absT_info_opts lthy =
687  let
688    val thy = Proof_Context.theory_of lthy;
689    fun co_swap pair = case_fp fp I swap pair;
690    val mk_co_comp = curry (HOLogic.mk_comp o co_swap);
691    fun mk_co_algT T U = case_fp fp (T --> U) (U --> T);
692    val co_alg_funT = case_fp fp domain_type range_type;
693    val mk_co_product = curry (case_fp fp mk_convol mk_case_sum);
694    val co_proj1_const = case_fp fp fst_const (uncurry Inl_const o dest_sumT) o co_alg_funT;
695    val co_proj2_const = case_fp fp snd_const (uncurry Inr_const o dest_sumT) o co_alg_funT;
696    val mk_co_productT = curry (case_fp fp HOLogic.mk_prodT mk_sumT);
697    val rewrite_comp_comp = case_fp fp @{thm rewriteL_comp_comp} @{thm rewriteR_comp_comp};
698
699    val n = length pre_bnfs;
700    val live = live_of_bnf (hd pre_bnfs);
701    val m = live - n;
702    val ks = 1 upto n;
703
704    val map_id0s = map map_id0_of_bnf pre_bnfs;
705    val map_comps = map map_comp_of_bnf pre_bnfs;
706    val map_cong0s = map map_cong0_of_bnf pre_bnfs;
707    val map_transfers = map map_transfer_of_bnf pre_bnfs;
708    val sym_map_comp0s = map (mk_sym o map_comp0_of_bnf) pre_bnfs;
709
710    val deads = fold (union (op =)) Dss resDs;
711    val ((((As, Bs), Xs), Ys), names_lthy) = lthy
712      |> fold Variable.declare_typ deads
713      |> mk_TFrees m
714      ||>> mk_TFrees m
715      ||>> mk_TFrees n
716      ||>> mk_TFrees n;
717
718    val XFTs = @{map 2} (fn Ds => mk_T_of_bnf Ds (As @ Xs)) Dss pre_bnfs;
719    val co_algXFTs = @{map 2} mk_co_algT XFTs Xs;
720    val Ts = mk_Ts As;
721    val un_foldTs = @{map 2} (fn T => fn X => co_algXFTs ---> mk_co_algT T X) Ts Xs;
722    val un_folds = @{map 2} (force_typ names_lthy) un_foldTs un_folds0;
723    val ABs = As ~~ Bs;
724    val XYs = Xs ~~ Ys;
725
726    val Us = map (typ_subst_atomic ABs) Ts;
727
728    val TFTs = @{map 2} (fn Ds => mk_T_of_bnf Ds (As @ Ts)) Dss pre_bnfs;
729
730    val TFTs' = @{map 2} (absT_info_decodeT thy) absT_info_opts TFTs;
731    val xtors = @{map 3} (force_typ names_lthy oo mk_co_algT) TFTs' Ts xtors0;
732
733    val ids = map HOLogic.id_const As;
734    val co_rec_Xs = @{map 2} mk_co_productT Ts Xs;
735    val co_rec_Ys = @{map 2} mk_co_productT Us Ys;
736    val co_rec_algXs = @{map 2} mk_co_algT co_rec_Xs Xs;
737    val co_proj1s = map co_proj1_const co_rec_algXs;
738    val co_rec_maps = @{map 2} (fn Ds =>
739      mk_map_of_bnf Ds (As @ case_fp fp co_rec_Xs Ts) (As @ case_fp fp Ts co_rec_Xs)) Dss pre_bnfs;
740    val co_rec_Ts = @{map 2} (fn Ds => mk_T_of_bnf Ds (As @ co_rec_Xs)) Dss pre_bnfs
741    val co_rec_argTs = @{map 2} mk_co_algT co_rec_Ts Xs;
742    val co_rec_resTs = @{map 2} mk_co_algT Ts Xs;
743
744    val (((co_rec_ss, fs), xs), names_lthy) = names_lthy
745      |> mk_Frees "s" co_rec_argTs
746      ||>> mk_Frees "f" co_rec_resTs
747      ||>> mk_Frees "x" (case_fp fp TFTs' Xs);
748
749    val co_rec_strs =
750      @{map 4} (fn xtor => fn s => fn mapx => fn absT_info_opt =>
751        mk_co_product (mk_co_comp (absT_info_encode thy fp absT_info_opt xtor)
752          (list_comb (mapx, ids @ co_proj1s))) s)
753      xtors co_rec_ss co_rec_maps absT_info_opts;
754
755    val theta = Xs ~~ co_rec_Xs;
756    val co_rec_un_folds = map (subst_atomic_types theta) un_folds;
757
758    val co_rec_spec0s = map (fn un_fold => list_comb (un_fold, co_rec_strs)) co_rec_un_folds;
759
760    val co_rec_ids = @{map 2} (mk_co_comp o co_proj1_const) co_rec_algXs co_rec_spec0s;
761    val co_rec_specs = @{map 2} (mk_co_comp o co_proj2_const) co_rec_algXs co_rec_spec0s;
762
763    val co_recN = case_fp fp ctor_recN dtor_corecN;
764    fun co_rec_bind i = nth bs (i - 1) |> Binding.prefix_name (co_recN ^ "_");
765    val co_rec_def_bind = rpair [] o Binding.concealed o Thm.def_binding o co_rec_bind;
766
767    fun co_rec_spec i =
768      fold_rev (Term.absfree o Term.dest_Free) co_rec_ss (nth co_rec_specs (i - 1));
769
770    val ((co_rec_frees, (_, co_rec_def_frees)), (lthy, lthy_old)) =
771      lthy
772      |> Local_Theory.open_target |> snd
773      |> fold_map (fn i =>
774        Local_Theory.define ((co_rec_bind i, NoSyn), (co_rec_def_bind i, co_rec_spec i))) ks
775      |>> apsnd split_list o split_list
776      ||> `Local_Theory.close_target;
777
778    val phi = Proof_Context.export_morphism lthy_old lthy;
779    val co_rec_names = map (fst o dest_Const o Morphism.term phi) co_rec_frees;
780    val co_recs = @{map 2} (fn name => fn resT =>
781      Const (name, co_rec_argTs ---> resT)) co_rec_names co_rec_resTs;
782    val co_rec_defs = map (fn def =>
783      mk_unabs_def n (HOLogic.mk_obj_eq (Morphism.thm phi def))) co_rec_def_frees;
784
785    val xtor_un_fold_xtor_thms =
786      mk_xtor_un_fold_xtor_thms lthy fp (map (Term.subst_atomic_types (Xs ~~ Ts)) un_folds)
787        xtors xtor_un_fold_unique map_id0s absT_info_opts;
788
789    val co_rec_id_thms =
790      let
791        val goal = @{map 2} (fn T => fn t => HOLogic.mk_eq (t, HOLogic.id_const T)) Ts co_rec_ids
792          |> Library.foldr1 HOLogic.mk_conj |> HOLogic.mk_Trueprop;
793        val vars = Variable.add_free_names lthy goal [];
794      in
795        Goal.prove_sorry lthy vars [] goal
796          (fn {context = ctxt, prems = _} => mk_xtor_co_rec_id_tac ctxt xtor_un_fold_xtor_thms
797            xtor_un_fold_unique xtor_un_folds map_comps)
798        |> Thm.close_derivation \<^here>
799        |> split_conj_thm
800      end;
801
802    val co_rec_app_ss = map (fn co_rec => list_comb (co_rec, co_rec_ss)) co_recs;
803    val co_products = @{map 2} (fn T => mk_co_product (HOLogic.id_const T)) Ts co_rec_app_ss;
804    val co_rec_maps_rev = @{map 2} (fn Ds =>
805      mk_map_of_bnf Ds (As @ case_fp fp Ts co_rec_Xs) (As @ case_fp fp co_rec_Xs Ts)) Dss pre_bnfs;
806    fun mk_co_app f g x = case_fp fp (f $ (g $ x)) (g $ (f $ x));
807    val co_rec_expand_thms = map (fn thm => thm RS
808      case_fp fp @{thm convol_expand_snd} @{thm case_sum_expand_Inr_pointfree}) co_rec_id_thms;
809    val xtor_co_rec_thms =
810      let
811        fun mk_goal co_rec s mapx xtor x absT_info_opt =
812          let
813            val lhs = mk_co_app co_rec xtor x;
814            val rhs = mk_co_app s
815              (list_comb (mapx, ids @ co_products) |> absT_info_decode thy fp absT_info_opt) x;
816          in
817            mk_Trueprop_eq (lhs, rhs)
818          end;
819        val goals =
820          @{map 6} mk_goal co_rec_app_ss co_rec_ss co_rec_maps_rev xtors xs absT_info_opts;
821      in
822        map2 (fn goal => fn un_fold =>
823          Variable.add_free_names lthy goal []
824          |> (fn vars => Goal.prove_sorry lthy vars [] goal
825            (fn {context = ctxt, prems = _} =>
826              mk_xtor_co_rec_tac ctxt un_fold co_rec_defs co_rec_expand_thms))
827          |> Thm.close_derivation \<^here>)
828        goals xtor_un_folds
829      end;
830
831    val co_product_fs = @{map 2} (fn T => mk_co_product (HOLogic.id_const T)) Ts fs;
832    val co_rec_expand'_thms = map (fn thm =>
833      thm RS case_fp fp @{thm convol_expand_snd'} @{thm case_sum_expand_Inr'}) co_rec_id_thms;
834    val xtor_co_rec_unique_thm =
835      let
836        fun mk_prem f s mapx xtor absT_info_opt =
837          let
838            val lhs = mk_co_comp f xtor;
839            val rhs = mk_co_comp s (list_comb (mapx, ids @ co_product_fs))
840              |> absT_info_decode thy fp absT_info_opt;
841          in
842            mk_Trueprop_eq (co_swap (lhs, rhs))
843          end;
844        val prems = @{map 5} mk_prem fs co_rec_ss co_rec_maps_rev xtors absT_info_opts;
845        val concl = @{map 2} (curry HOLogic.mk_eq) fs co_rec_app_ss
846          |> Library.foldr1 HOLogic.mk_conj |> HOLogic.mk_Trueprop;
847        val goal = Logic.list_implies (prems, concl);
848        val vars = Variable.add_free_names lthy goal [];
849        fun mk_inverses NONE = []
850          | mk_inverses (SOME (src, dst)) =
851            [#type_definition dst RS @{thm type_copy_Rep_o_Abs} RS rewrite_comp_comp,
852             #type_definition src RS @{thm type_copy_Abs_o_Rep}];
853        val inverses = maps mk_inverses absT_info_opts;
854      in
855        Goal.prove_sorry lthy vars [] goal
856          (fn {context = ctxt, prems = _} => mk_xtor_co_rec_unique_tac ctxt fp co_rec_defs
857            co_rec_expand'_thms xtor_un_fold_unique map_id0s sym_map_comp0s inverses)
858        |> Thm.close_derivation \<^here>
859      end;
860
861    val xtor_co_rec_o_map_thms = if forall is_none absT_info_opts
862      then
863        mk_xtor_co_iter_o_map_thms fp true m xtor_co_rec_unique_thm
864          (map (mk_pointfree2 lthy) xtor_maps) (map (mk_pointfree2 lthy) xtor_co_rec_thms)
865          sym_map_comp0s map_cong0s
866      else
867        replicate n refl (* FIXME *);
868
869    val ABphiTs = @{map 2} mk_pred2T As Bs;
870    val XYphiTs = @{map 2} mk_pred2T Xs Ys;
871
872    val ((ABphis, XYphis), names_lthy) = names_lthy
873      |> mk_Frees "R" ABphiTs
874      ||>> mk_Frees "S" XYphiTs;
875
876    val xtor_co_rec_transfer_thms = if forall is_none absT_info_opts
877      then
878        let
879          val pre_rels =
880            @{map 2} (fn Ds => mk_rel_of_bnf Ds (As @ co_rec_Xs) (Bs @ co_rec_Ys)) Dss pre_bnfs;
881          val rels = @{map 3} (fn T => fn T' => Thm.prop_of #> HOLogic.dest_Trueprop
882              #> fst o dest_comb #> fst o dest_comb #> funpow n (snd o dest_comb)
883              #> case_fp fp (fst o dest_comb #> snd o dest_comb) (snd o dest_comb) #> head_of
884              #> force_typ names_lthy (ABphiTs ---> mk_pred2T T T'))
885            Ts Us xtor_un_fold_transfers;
886
887          fun tac {context = ctxt, prems = _} = mk_xtor_co_rec_transfer_tac ctxt fp n m co_rec_defs
888            xtor_un_fold_transfers map_transfers xtor_rels;
889
890          val mk_rel_co_product = case_fp fp mk_rel_prod mk_rel_sum;
891          val rec_phis =
892            map2 (fn rel => mk_rel_co_product (Term.list_comb (rel, ABphis))) rels XYphis;
893        in
894          mk_xtor_co_iter_transfer_thms fp pre_rels rec_phis XYphis rels ABphis
895            co_recs (map (subst_atomic_types (ABs @ XYs)) co_recs) tac lthy
896        end
897      else
898        replicate n TrueI (* FIXME *);
899
900    val notes =
901      [(case_fp fp ctor_recN dtor_corecN, xtor_co_rec_thms),
902       (case_fp fp ctor_rec_uniqueN dtor_corec_uniqueN, split_conj_thm xtor_co_rec_unique_thm),
903       (case_fp fp ctor_rec_o_mapN dtor_corec_o_mapN, xtor_co_rec_o_map_thms),
904       (case_fp fp ctor_rec_transferN dtor_corec_transferN, xtor_co_rec_transfer_thms)]
905      |> map (apsnd (map single))
906      |> maps (fn (thmN, thmss) =>
907        map2 (fn b => fn thms =>
908          ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]))
909        bs thmss);
910
911     val lthy = lthy |> Config.get lthy bnf_internals ? snd o Local_Theory.notes notes;
912  in
913    ((co_recs,
914     (xtor_co_rec_thms, xtor_co_rec_unique_thm, xtor_co_rec_o_map_thms, xtor_co_rec_transfer_thms)),
915      lthy)
916  end;
917
918fun raw_qualify extra_qualify base_b =
919  let
920    val qs = Binding.path_of base_b;
921    val n = Binding.name_of base_b;
922  in
923    Binding.prefix_name rawN
924    #> fold_rev (fn (s, mand) => Binding.qualify mand s) (qs @ [(n, true)])
925    #> extra_qualify #> Binding.concealed
926  end;
927
928fun fixpoint_bnf force_out_of_line extra_qualify construct_fp bs resBs Ds0 Xs rhsXs comp_cache0
929    lthy =
930  let
931    val time = time lthy;
932    val timer = time (Timer.startRealTimer ());
933
934    fun flatten_tyargs Ass =
935      subtract (op =) Xs (filter (fn T => exists (fn Ts => member (op =) Ts T) Ass) resBs) @ Xs;
936
937    val ((bnfs, (deadss, livess)), (comp_cache_unfold_set, lthy')) =
938      apfst (apsnd split_list o split_list)
939        (@{fold_map 2}
940          (fn b => bnf_of_typ true Smart_Inline (raw_qualify extra_qualify b) flatten_tyargs Xs Ds0)
941          bs rhsXs ((comp_cache0, empty_unfolds), lthy));
942
943    fun norm_qualify i =
944      Binding.qualify true (Binding.name_of (nth bs (Int.max (0, i - 1))))
945      #> extra_qualify #> Binding.concealed;
946
947    val Ass = map (map dest_TFree) livess;
948    val Ds' = fold (fold Term.add_tfreesT) deadss [];
949    val Ds = union (op =) Ds' Ds0;
950    val missing = resBs |> fold (subtract (op =)) (Ds' :: Ass);
951    val (dead_phantoms, live_phantoms) = List.partition (member (op =) Ds0) missing;
952    val resBs' = resBs |> fold (subtract (op =)) [dead_phantoms, Ds];
953
954    val timer = time (timer "Construction of BNFs");
955
956    val ((kill_posss, _), (bnfs', ((comp_cache', unfold_set'), lthy''))) =
957      normalize_bnfs norm_qualify Ass Ds (K (resBs' @ Xs)) bnfs (comp_cache_unfold_set, lthy');
958
959    val Dss = @{map 3} (fn lives => fn kill_posss => fn deads => deads @ map (nth lives) kill_posss)
960      livess kill_posss deadss;
961    val all_Dss = Dss |> force_out_of_line ? map (fn Ds' => union (op =) Ds' (map TFree Ds0));
962
963    fun pre_qualify b =
964      Binding.qualify false (Binding.name_of b)
965      #> extra_qualify
966      #> not (Config.get lthy'' bnf_internals) ? Binding.concealed;
967
968    val ((pre_bnfs, (deadss, absT_infos)), lthy''') = lthy''
969      |> @{fold_map 5} (fn b => seal_bnf (pre_qualify b) unfold_set' (Binding.prefix_name preN b))
970        bs (replicate (length rhsXs) (force_out_of_line orelse not (null live_phantoms))) Dss
971        all_Dss bnfs'
972      |>> split_list
973      |>> apsnd split_list;
974
975    val timer = time (timer "Normalization & sealing of BNFs");
976
977    val res = construct_fp bs resBs (map TFree dead_phantoms, deadss) pre_bnfs absT_infos lthy''';
978
979    val timer = time (timer "FP construction in total");
980  in
981    (((pre_bnfs, absT_infos), comp_cache'), res)
982  end;
983
984end;
985