src/HOL/BNF/Tools/bnf_fp.ML
changeset 51850 106afdf5806c
parent 51849 19ee0cebe76d
child 51851 7e9265a0eb01
equal deleted inserted replaced
51849:19ee0cebe76d 51850:106afdf5806c
     1 (*  Title:      HOL/BNF/Tools/bnf_fp.ML
       
     2     Author:     Dmitriy Traytel, TU Muenchen
       
     3     Author:     Jasmin Blanchette, TU Muenchen
       
     4     Copyright   2012, 2013
       
     5 
       
     6 Shared library for the datatype and codatatype constructions.
       
     7 *)
       
     8 
       
     9 signature BNF_FP =
       
    10 sig
       
    11   type fp_result =
       
    12     {bnfs: BNF_Def.bnf list,
       
    13      ctors: term list,
       
    14      dtors: term list,
       
    15      folds: term list,
       
    16      recs: term list,
       
    17      induct: thm,
       
    18      strong_induct: thm,
       
    19      dtor_ctors: thm list,
       
    20      ctor_dtors: thm list,
       
    21      ctor_injects: thm list,
       
    22      map_thms: thm list,
       
    23      set_thmss: thm list list,
       
    24      rel_thms: thm list,
       
    25      fold_thms: thm list,
       
    26      rec_thms: thm list}
       
    27 
       
    28   val fp_name_of_ctor: term -> string
       
    29   val morph_fp_result: morphism -> fp_result -> fp_result
       
    30   val eq_fp_result: fp_result * fp_result -> bool
       
    31 
       
    32   val time: Timer.real_timer -> string -> Timer.real_timer
       
    33 
       
    34   val IITN: string
       
    35   val LevN: string
       
    36   val algN: string
       
    37   val behN: string
       
    38   val bisN: string
       
    39   val carTN: string
       
    40   val caseN: string
       
    41   val coN: string
       
    42   val coinductN: string
       
    43   val corecN: string
       
    44   val ctorN: string
       
    45   val ctor_dtorN: string
       
    46   val ctor_dtor_corecN: string
       
    47   val ctor_dtor_unfoldN: string
       
    48   val ctor_exhaustN: string
       
    49   val ctor_induct2N: string
       
    50   val ctor_inductN: string
       
    51   val ctor_injectN: string
       
    52   val ctor_foldN: string
       
    53   val ctor_fold_uniqueN: string
       
    54   val ctor_mapN: string
       
    55   val ctor_map_uniqueN: string
       
    56   val ctor_recN: string
       
    57   val ctor_rec_uniqueN: string
       
    58   val ctor_relN: string
       
    59   val ctor_set_inclN: string
       
    60   val ctor_set_set_inclN: string
       
    61   val ctor_srelN: string
       
    62   val disc_unfoldN: string
       
    63   val disc_unfold_iffN: string
       
    64   val disc_corecN: string
       
    65   val disc_corec_iffN: string
       
    66   val dtorN: string
       
    67   val dtor_coinductN: string
       
    68   val dtor_corecN: string
       
    69   val dtor_corec_uniqueN: string
       
    70   val dtor_ctorN: string
       
    71   val dtor_exhaustN: string
       
    72   val dtor_injectN: string
       
    73   val dtor_mapN: string
       
    74   val dtor_map_coinductN: string
       
    75   val dtor_map_strong_coinductN: string
       
    76   val dtor_map_uniqueN: string
       
    77   val dtor_relN: string
       
    78   val dtor_set_inclN: string
       
    79   val dtor_set_set_inclN: string
       
    80   val dtor_srelN: string
       
    81   val dtor_srel_coinductN: string
       
    82   val dtor_srel_strong_coinductN: string
       
    83   val dtor_strong_coinductN: string
       
    84   val dtor_unfoldN: string
       
    85   val dtor_unfold_uniqueN: string
       
    86   val exhaustN: string
       
    87   val foldN: string
       
    88   val hsetN: string
       
    89   val hset_recN: string
       
    90   val inductN: string
       
    91   val injectN: string
       
    92   val isNodeN: string
       
    93   val lsbisN: string
       
    94   val mapN: string
       
    95   val map_uniqueN: string
       
    96   val min_algN: string
       
    97   val morN: string
       
    98   val nchotomyN: string
       
    99   val recN: string
       
   100   val rel_injectN: string
       
   101   val rel_distinctN: string
       
   102   val rvN: string
       
   103   val sel_corecN: string
       
   104   val set_inclN: string
       
   105   val set_set_inclN: string
       
   106   val sel_unfoldN: string
       
   107   val setsN: string
       
   108   val simpsN: string
       
   109   val strTN: string
       
   110   val str_initN: string
       
   111   val strong_coinductN: string
       
   112   val sum_bdN: string
       
   113   val sum_bdTN: string
       
   114   val unfoldN: string
       
   115   val uniqueN: string
       
   116 
       
   117   (* TODO: Don't index set facts. Isabelle packages traditionally generate uniform names. *)
       
   118   val mk_ctor_setN: int -> string
       
   119   val mk_dtor_setN: int -> string
       
   120   val mk_dtor_set_inductN: int -> string
       
   121   val mk_set_inductN: int -> string
       
   122 
       
   123   val mk_common_name: string list -> string
       
   124 
       
   125   val split_conj_thm: thm -> thm list
       
   126   val split_conj_prems: int -> thm -> thm
       
   127 
       
   128   val mk_sumTN: typ list -> typ
       
   129   val mk_sumTN_balanced: typ list -> typ
       
   130 
       
   131   val id_const: typ -> term
       
   132 
       
   133   val Inl_const: typ -> typ -> term
       
   134   val Inr_const: typ -> typ -> term
       
   135 
       
   136   val mk_Inl: typ -> term -> term
       
   137   val mk_Inr: typ -> term -> term
       
   138   val mk_InN: typ list -> term -> int -> term
       
   139   val mk_InN_balanced: typ -> int -> term -> int -> term
       
   140   val mk_sum_case: term * term -> term
       
   141   val mk_sum_caseN: term list -> term
       
   142   val mk_sum_caseN_balanced: term list -> term
       
   143 
       
   144   val dest_sumT: typ -> typ * typ
       
   145   val dest_sumTN: int -> typ -> typ list
       
   146   val dest_sumTN_balanced: int -> typ -> typ list
       
   147   val dest_tupleT: int -> typ -> typ list
       
   148 
       
   149   val mk_Field: term -> term
       
   150   val mk_If: term -> term -> term -> term
       
   151   val mk_union: term * term -> term
       
   152 
       
   153   val mk_sumEN: int -> thm
       
   154   val mk_sumEN_balanced: int -> thm
       
   155   val mk_sumEN_tupled_balanced: int list -> thm
       
   156   val mk_sum_casesN: int -> int -> thm
       
   157   val mk_sum_casesN_balanced: int -> int -> thm
       
   158 
       
   159   val fixpoint: ('a * 'a -> bool) -> ('a list -> 'a list) -> 'a list -> 'a list
       
   160 
       
   161   val fp_bnf: (mixfix list -> (string * sort) list option -> binding list -> binding list ->
       
   162       binding list -> binding list list -> typ list * typ list list -> BNF_Def.bnf list ->
       
   163       local_theory -> 'a) ->
       
   164     binding list -> mixfix list -> binding list -> binding list -> binding list list ->
       
   165     (string * sort) list -> ((string * sort) * typ) list -> local_theory -> BNF_Def.bnf list * 'a
       
   166 end;
       
   167 
       
   168 structure BNF_FP : BNF_FP =
       
   169 struct
       
   170 
       
   171 open BNF_Comp
       
   172 open BNF_Def
       
   173 open BNF_Util
       
   174 
       
   175 type fp_result =
       
   176   {bnfs: BNF_Def.bnf list,
       
   177    ctors: term list,
       
   178    dtors: term list,
       
   179    folds: term list,
       
   180    recs: term list,
       
   181    induct: thm,
       
   182    strong_induct: thm,
       
   183    dtor_ctors: thm list,
       
   184    ctor_dtors: thm list,
       
   185    ctor_injects: thm list,
       
   186    map_thms: thm list,
       
   187    set_thmss: thm list list,
       
   188    rel_thms: thm list,
       
   189    fold_thms: thm list,
       
   190    rec_thms: thm list};
       
   191 
       
   192 val fp_name_of_ctor = fst o dest_Type o range_type o fastype_of;
       
   193 
       
   194 fun morph_fp_result phi {bnfs, ctors, dtors, folds, recs, induct, strong_induct, dtor_ctors,
       
   195     ctor_dtors, ctor_injects, map_thms, set_thmss, rel_thms, fold_thms, rec_thms} =
       
   196   {bnfs = map (morph_bnf phi) bnfs,
       
   197    ctors = map (Morphism.term phi) ctors,
       
   198    dtors = map (Morphism.term phi) dtors,
       
   199    folds = map (Morphism.term phi) folds,
       
   200    recs = map (Morphism.term phi) recs,
       
   201    induct = Morphism.thm phi induct,
       
   202    strong_induct = Morphism.thm phi strong_induct,
       
   203    dtor_ctors = map (Morphism.thm phi) dtor_ctors,
       
   204    ctor_dtors = map (Morphism.thm phi) ctor_dtors,
       
   205    ctor_injects = map (Morphism.thm phi) ctor_injects,
       
   206    map_thms = map (Morphism.thm phi) map_thms,
       
   207    set_thmss = map (map (Morphism.thm phi)) set_thmss,
       
   208    rel_thms = map (Morphism.thm phi) rel_thms,
       
   209    fold_thms = map (Morphism.thm phi) fold_thms,
       
   210    rec_thms = map (Morphism.thm phi) rec_thms};
       
   211 
       
   212 fun eq_fp_result ({bnfs = bnfs1, ...} : fp_result, {bnfs = bnfs2, ...} : fp_result) =
       
   213   eq_list eq_bnf (bnfs1, bnfs2);
       
   214 
       
   215 val timing = true;
       
   216 fun time timer msg = (if timing
       
   217   then warning (msg ^ ": " ^ ATP_Util.string_from_time (Timer.checkRealTimer timer))
       
   218   else (); Timer.startRealTimer ());
       
   219 
       
   220 val preN = "pre_"
       
   221 val rawN = "raw_"
       
   222 
       
   223 val coN = "co"
       
   224 val unN = "un"
       
   225 val algN = "alg"
       
   226 val IITN = "IITN"
       
   227 val foldN = "fold"
       
   228 val unfoldN = unN ^ foldN
       
   229 val uniqueN = "_unique"
       
   230 val simpsN = "simps"
       
   231 val ctorN = "ctor"
       
   232 val dtorN = "dtor"
       
   233 val ctor_foldN = ctorN ^ "_" ^ foldN
       
   234 val dtor_unfoldN = dtorN ^ "_" ^ unfoldN
       
   235 val ctor_fold_uniqueN = ctor_foldN ^ uniqueN
       
   236 val dtor_unfold_uniqueN = dtor_unfoldN ^ uniqueN
       
   237 val ctor_dtor_unfoldN = ctorN ^ "_" ^ dtor_unfoldN
       
   238 val ctor_mapN = ctorN ^ "_" ^ mapN
       
   239 val dtor_mapN = dtorN ^ "_" ^ mapN
       
   240 val map_uniqueN = mapN ^ uniqueN
       
   241 val ctor_map_uniqueN = ctorN ^ "_" ^ map_uniqueN
       
   242 val dtor_map_uniqueN = dtorN ^ "_" ^ map_uniqueN
       
   243 val min_algN = "min_alg"
       
   244 val morN = "mor"
       
   245 val bisN = "bis"
       
   246 val lsbisN = "lsbis"
       
   247 val sum_bdTN = "sbdT"
       
   248 val sum_bdN = "sbd"
       
   249 val carTN = "carT"
       
   250 val strTN = "strT"
       
   251 val isNodeN = "isNode"
       
   252 val LevN = "Lev"
       
   253 val rvN = "recover"
       
   254 val behN = "beh"
       
   255 val setsN = "sets"
       
   256 val mk_ctor_setN = prefix (ctorN ^ "_") o mk_setN
       
   257 val mk_dtor_setN = prefix (dtorN ^ "_") o mk_setN
       
   258 fun mk_set_inductN i = mk_setN i ^ "_induct"
       
   259 val mk_dtor_set_inductN = prefix (dtorN ^ "_") o mk_set_inductN
       
   260 
       
   261 val str_initN = "str_init"
       
   262 val recN = "rec"
       
   263 val corecN = coN ^ recN
       
   264 val ctor_recN = ctorN ^ "_" ^ recN
       
   265 val ctor_rec_uniqueN = ctor_recN ^ uniqueN
       
   266 val dtor_corecN = dtorN ^ "_" ^ corecN
       
   267 val dtor_corec_uniqueN = dtor_corecN ^ uniqueN
       
   268 val ctor_dtor_corecN = ctorN ^ "_" ^ dtor_corecN
       
   269 
       
   270 val ctor_dtorN = ctorN ^ "_" ^ dtorN
       
   271 val dtor_ctorN = dtorN ^ "_" ^ ctorN
       
   272 val nchotomyN = "nchotomy"
       
   273 val injectN = "inject"
       
   274 val exhaustN = "exhaust"
       
   275 val ctor_injectN = ctorN ^ "_" ^ injectN
       
   276 val ctor_exhaustN = ctorN ^ "_" ^ exhaustN
       
   277 val dtor_injectN = dtorN ^ "_" ^ injectN
       
   278 val dtor_exhaustN = dtorN ^ "_" ^ exhaustN
       
   279 val ctor_relN = ctorN ^ "_" ^ relN
       
   280 val dtor_relN = dtorN ^ "_" ^ relN
       
   281 val ctor_srelN = ctorN ^ "_" ^ srelN
       
   282 val dtor_srelN = dtorN ^ "_" ^ srelN
       
   283 val inductN = "induct"
       
   284 val coinductN = coN ^ inductN
       
   285 val ctor_inductN = ctorN ^ "_" ^ inductN
       
   286 val ctor_induct2N = ctor_inductN ^ "2"
       
   287 val dtor_map_coinductN = dtor_mapN ^ "_" ^ coinductN
       
   288 val dtor_coinductN = dtorN ^ "_" ^ coinductN
       
   289 val dtor_srel_coinductN = dtor_srelN ^ "_" ^ coinductN
       
   290 val strong_coinductN = "strong_" ^ coinductN
       
   291 val dtor_map_strong_coinductN = dtor_mapN ^ "_" ^ strong_coinductN
       
   292 val dtor_strong_coinductN = dtorN ^ "_" ^ strong_coinductN
       
   293 val dtor_srel_strong_coinductN = dtor_srelN ^ "_" ^ strong_coinductN
       
   294 val hsetN = "Hset"
       
   295 val hset_recN = hsetN ^ "_rec"
       
   296 val set_inclN = "set_incl"
       
   297 val ctor_set_inclN = ctorN ^ "_" ^ set_inclN
       
   298 val dtor_set_inclN = dtorN ^ "_" ^ set_inclN
       
   299 val set_set_inclN = "set_set_incl"
       
   300 val ctor_set_set_inclN = ctorN ^ "_" ^ set_set_inclN
       
   301 val dtor_set_set_inclN = dtorN ^ "_" ^ set_set_inclN
       
   302 
       
   303 val caseN = "case"
       
   304 val discN = "disc"
       
   305 val disc_unfoldN = discN ^ "_" ^ unfoldN
       
   306 val disc_corecN = discN ^ "_" ^ corecN
       
   307 val iffN = "_iff"
       
   308 val disc_unfold_iffN = discN ^ "_" ^ unfoldN ^ iffN
       
   309 val disc_corec_iffN = discN ^ "_" ^ corecN ^ iffN
       
   310 val distinctN = "distinct"
       
   311 val rel_distinctN = relN ^ "_" ^ distinctN
       
   312 val injectN = "inject"
       
   313 val rel_injectN = relN ^ "_" ^ injectN
       
   314 val selN = "sel"
       
   315 val sel_unfoldN = selN ^ "_" ^ unfoldN
       
   316 val sel_corecN = selN ^ "_" ^ corecN
       
   317 
       
   318 val mk_common_name = space_implode "_";
       
   319 
       
   320 fun dest_sumT (Type (@{type_name sum}, [T, T'])) = (T, T');
       
   321 
       
   322 fun dest_sumTN 1 T = [T]
       
   323   | dest_sumTN n (Type (@{type_name sum}, [T, T'])) = T :: dest_sumTN (n - 1) T';
       
   324 
       
   325 val dest_sumTN_balanced = Balanced_Tree.dest dest_sumT;
       
   326 
       
   327 (* TODO: move something like this to "HOLogic"? *)
       
   328 fun dest_tupleT 0 @{typ unit} = []
       
   329   | dest_tupleT 1 T = [T]
       
   330   | dest_tupleT n (Type (@{type_name prod}, [T, T'])) = T :: dest_tupleT (n - 1) T';
       
   331 
       
   332 val mk_sumTN = Library.foldr1 mk_sumT;
       
   333 val mk_sumTN_balanced = Balanced_Tree.make mk_sumT;
       
   334 
       
   335 fun id_const T = Const (@{const_name id}, T --> T);
       
   336 
       
   337 fun Inl_const LT RT = Const (@{const_name Inl}, LT --> mk_sumT (LT, RT));
       
   338 fun mk_Inl RT t = Inl_const (fastype_of t) RT $ t;
       
   339 
       
   340 fun Inr_const LT RT = Const (@{const_name Inr}, RT --> mk_sumT (LT, RT));
       
   341 fun mk_Inr LT t = Inr_const LT (fastype_of t) $ t;
       
   342 
       
   343 fun mk_InN [_] t 1 = t
       
   344   | mk_InN (_ :: Ts) t 1 = mk_Inl (mk_sumTN Ts) t
       
   345   | mk_InN (LT :: Ts) t m = mk_Inr LT (mk_InN Ts t (m - 1))
       
   346   | mk_InN Ts t _ = raise (TYPE ("mk_InN", Ts, [t]));
       
   347 
       
   348 fun mk_InN_balanced sum_T n t k =
       
   349   let
       
   350     fun repair_types T (Const (s as @{const_name Inl}, _) $ t) = repair_inj_types T s fst t
       
   351       | repair_types T (Const (s as @{const_name Inr}, _) $ t) = repair_inj_types T s snd t
       
   352       | repair_types _ t = t
       
   353     and repair_inj_types T s get t =
       
   354       let val T' = get (dest_sumT T) in
       
   355         Const (s, T' --> T) $ repair_types T' t
       
   356       end;
       
   357   in
       
   358     Balanced_Tree.access {left = mk_Inl dummyT, right = mk_Inr dummyT, init = t} n k
       
   359     |> repair_types sum_T
       
   360   end;
       
   361 
       
   362 fun mk_sum_case (f, g) =
       
   363   let
       
   364     val fT = fastype_of f;
       
   365     val gT = fastype_of g;
       
   366   in
       
   367     Const (@{const_name sum_case},
       
   368       fT --> gT --> mk_sumT (domain_type fT, domain_type gT) --> range_type fT) $ f $ g
       
   369   end;
       
   370 
       
   371 val mk_sum_caseN = Library.foldr1 mk_sum_case;
       
   372 val mk_sum_caseN_balanced = Balanced_Tree.make mk_sum_case;
       
   373 
       
   374 fun mk_If p t f =
       
   375   let val T = fastype_of t;
       
   376   in Const (@{const_name If}, HOLogic.boolT --> T --> T --> T) $ p $ t $ f end;
       
   377 
       
   378 fun mk_Field r =
       
   379   let val T = fst (dest_relT (fastype_of r));
       
   380   in Const (@{const_name Field}, mk_relT (T, T) --> HOLogic.mk_setT T) $ r end;
       
   381 
       
   382 val mk_union = HOLogic.mk_binop @{const_name sup};
       
   383 
       
   384 (*dangerous; use with monotonic, converging functions only!*)
       
   385 fun fixpoint eq f X = if subset eq (f X, X) then X else fixpoint eq f (f X);
       
   386 
       
   387 (* stolen from "~~/src/HOL/Tools/Datatype/datatype_aux.ML" *)
       
   388 fun split_conj_thm th =
       
   389   ((th RS conjunct1) :: split_conj_thm (th RS conjunct2)) handle THM _ => [th];
       
   390 
       
   391 fun split_conj_prems limit th =
       
   392   let
       
   393     fun split n i th =
       
   394       if i = n then th else split n (i + 1) (conjI RSN (i, th)) handle THM _ => th;
       
   395   in split limit 1 th end;
       
   396 
       
   397 fun mk_sumEN 1 = @{thm one_pointE}
       
   398   | mk_sumEN 2 = @{thm sumE}
       
   399   | mk_sumEN n =
       
   400     (fold (fn i => fn thm => @{thm obj_sum_step} RSN (i, thm)) (2 upto n - 1) @{thm obj_sumE}) OF
       
   401       replicate n (impI RS allI);
       
   402 
       
   403 fun mk_obj_sumEN_balanced n =
       
   404   Balanced_Tree.make (fn (thm1, thm2) => thm1 RSN (1, thm2 RSN (2, @{thm obj_sumE_f})))
       
   405     (replicate n asm_rl);
       
   406 
       
   407 fun mk_sumEN_balanced' n all_impIs = mk_obj_sumEN_balanced n OF all_impIs RS @{thm obj_one_pointE};
       
   408 
       
   409 fun mk_sumEN_balanced 1 = @{thm one_pointE} (*optimization*)
       
   410   | mk_sumEN_balanced 2 = @{thm sumE} (*optimization*)
       
   411   | mk_sumEN_balanced n = mk_sumEN_balanced' n (replicate n (impI RS allI));
       
   412 
       
   413 fun mk_tupled_allIN 0 = @{thm unit_all_impI}
       
   414   | mk_tupled_allIN 1 = @{thm impI[THEN allI]}
       
   415   | mk_tupled_allIN 2 = @{thm prod_all_impI} (*optimization*)
       
   416   | mk_tupled_allIN n = mk_tupled_allIN (n - 1) RS @{thm prod_all_impI_step};
       
   417 
       
   418 fun mk_sumEN_tupled_balanced ms =
       
   419   let val n = length ms in
       
   420     if forall (curry (op =) 1) ms then mk_sumEN_balanced n
       
   421     else mk_sumEN_balanced' n (map mk_tupled_allIN ms)
       
   422   end;
       
   423 
       
   424 fun mk_sum_casesN 1 1 = refl
       
   425   | mk_sum_casesN _ 1 = @{thm sum.cases(1)}
       
   426   | mk_sum_casesN 2 2 = @{thm sum.cases(2)}
       
   427   | mk_sum_casesN n k = trans OF [@{thm sum_case_step(2)}, mk_sum_casesN (n - 1) (k - 1)];
       
   428 
       
   429 fun mk_sum_step base step thm =
       
   430   if Thm.eq_thm_prop (thm, refl) then base else trans OF [step, thm];
       
   431 
       
   432 fun mk_sum_casesN_balanced 1 1 = refl
       
   433   | mk_sum_casesN_balanced n k =
       
   434     Balanced_Tree.access {left = mk_sum_step @{thm sum.cases(1)} @{thm sum_case_step(1)},
       
   435       right = mk_sum_step @{thm sum.cases(2)} @{thm sum_case_step(2)}, init = refl} n k;
       
   436 
       
   437 (* FIXME: because of "@ lhss", the output could contain type variables that are not in the input;
       
   438    also, "fp_sort" should put the "resBs" first and in the order in which they appear *)
       
   439 fun fp_sort lhss NONE Ass = Library.sort (Term_Ord.typ_ord o pairself TFree)
       
   440     (subtract (op =) lhss (fold (fold (insert (op =))) Ass [])) @ lhss
       
   441   | fp_sort lhss (SOME resBs) Ass =
       
   442     (subtract (op =) lhss (filter (fn T => exists (fn Ts => member (op =) Ts T) Ass) resBs)) @ lhss;
       
   443 
       
   444 fun mk_fp_bnf timer construct_fp resBs bs map_bs rel_bs set_bss sort lhss bnfs deadss livess
       
   445     unfold_set lthy =
       
   446   let
       
   447     val name = mk_common_name (map Binding.name_of bs);
       
   448     fun qualify i =
       
   449       let val namei = name ^ nonzero_string_of_int i;
       
   450       in Binding.qualify true namei end;
       
   451 
       
   452     val Ass = map (map dest_TFree) livess;
       
   453     val resDs = (case resBs of NONE => [] | SOME Ts => fold (subtract (op =)) Ass Ts);
       
   454     val Ds = fold (fold Term.add_tfreesT) deadss [];
       
   455 
       
   456     val _ = (case Library.inter (op =) Ds lhss of [] => ()
       
   457       | A :: _ => error ("Inadmissible type recursion (cannot take fixed point of dead type \
       
   458         \variable " ^ quote (Syntax.string_of_typ lthy (TFree A)) ^ ")"));
       
   459 
       
   460     val timer = time (timer "Construction of BNFs");
       
   461 
       
   462     val ((kill_poss, _), (bnfs', (unfold_set', lthy'))) =
       
   463       normalize_bnfs qualify Ass Ds sort bnfs unfold_set lthy;
       
   464 
       
   465     val Dss = map3 (append oo map o nth) livess kill_poss deadss;
       
   466 
       
   467     val ((bnfs'', deadss), lthy'') =
       
   468       fold_map3 (seal_bnf unfold_set') (map (Binding.prefix_name preN) bs) Dss bnfs' lthy'
       
   469       |>> split_list;
       
   470 
       
   471     val timer = time (timer "Normalization & sealing of BNFs");
       
   472 
       
   473     val res = construct_fp resBs bs map_bs rel_bs set_bss (map TFree resDs, deadss) bnfs'' lthy'';
       
   474 
       
   475     val timer = time (timer "FP construction in total");
       
   476   in
       
   477     timer; (bnfs'', res)
       
   478   end;
       
   479 
       
   480 fun fp_bnf construct_fp bs mixfixes map_bs rel_bs set_bss resBs eqs lthy =
       
   481   let
       
   482     val timer = time (Timer.startRealTimer ());
       
   483     val (lhss, rhss) = split_list eqs;
       
   484     val sort = fp_sort lhss (SOME resBs);
       
   485     fun qualify b = Binding.qualify true (Binding.name_of (Binding.prefix_name rawN b));
       
   486     val ((bnfs, (Dss, Ass)), (unfold_set, lthy')) = apfst (apsnd split_list o split_list)
       
   487       (fold_map2 (fn b => bnf_of_typ Smart_Inline (qualify b) sort) bs rhss
       
   488         (empty_unfolds, lthy));
       
   489   in
       
   490     mk_fp_bnf timer (construct_fp mixfixes) (SOME resBs) bs map_bs rel_bs set_bss sort lhss bnfs Dss
       
   491       Ass unfold_set lthy'
       
   492   end;
       
   493 
       
   494 end;