src/HOL/BNF/Tools/bnf_fp_n2m_sugar.ML
author blanchet
Wed, 06 Nov 2013 22:42:54 +0100
changeset 55735 6f0a49ed1bb1
parent 55734 32b5c4821d9d
child 55736 0b53378080d9
permissions -rw-r--r--
fourth attempt at generalizing N2M types (to leverage caching)
blanchet@54440
     1
(*  Title:      HOL/BNF/Tools/bnf_fp_n2m_sugar.ML
blanchet@54440
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@54440
     3
    Copyright   2013
blanchet@54440
     4
blanchet@54440
     5
Suggared flattening of nested to mutual (co)recursion.
blanchet@54440
     6
*)
blanchet@54440
     7
blanchet@54440
     8
signature BNF_FP_N2M_SUGAR =
blanchet@54440
     9
sig
blanchet@55695
    10
  val unfold_let: term -> term
blanchet@55695
    11
  val dest_map: Proof.context -> string -> term -> term * term list
blanchet@55695
    12
blanchet@55146
    13
  val mutualize_fp_sugars: bool -> BNF_FP_Util.fp_kind -> binding list -> typ list ->
blanchet@54440
    14
    (term -> int list) -> term list list list list -> BNF_FP_Def_Sugar.fp_sugar list ->
blanchet@54883
    15
    local_theory ->
blanchet@54883
    16
    (BNF_FP_Def_Sugar.fp_sugar list
blanchet@54883
    17
     * (BNF_FP_Def_Sugar.lfp_sugar_thms option * BNF_FP_Def_Sugar.gfp_sugar_thms option))
blanchet@54883
    18
    * local_theory
blanchet@55719
    19
  val indexify_callsss: BNF_FP_Def_Sugar.fp_sugar -> (term * term list list) list ->
blanchet@55719
    20
    term list list list
blanchet@55146
    21
  val nested_to_mutual_fps: BNF_FP_Util.fp_kind -> binding list -> typ list -> (term -> int list) ->
blanchet@55146
    22
    (term * term list list) list list -> local_theory ->
blanchet@54883
    23
    (typ list * int list * BNF_FP_Def_Sugar.fp_sugar list
blanchet@54883
    24
     * (BNF_FP_Def_Sugar.lfp_sugar_thms option * BNF_FP_Def_Sugar.gfp_sugar_thms option))
blanchet@54883
    25
    * local_theory
blanchet@54440
    26
end;
blanchet@54440
    27
blanchet@54440
    28
structure BNF_FP_N2M_Sugar : BNF_FP_N2M_SUGAR =
blanchet@54440
    29
struct
blanchet@54440
    30
blanchet@55143
    31
open Ctr_Sugar
blanchet@54440
    32
open BNF_Util
blanchet@54440
    33
open BNF_Def
blanchet@54440
    34
open BNF_FP_Util
blanchet@54440
    35
open BNF_FP_Def_Sugar
blanchet@54440
    36
open BNF_FP_N2M
blanchet@54440
    37
blanchet@54440
    38
val n2mN = "n2m_"
blanchet@54440
    39
blanchet@55708
    40
type n2m_sugar = fp_sugar list * (lfp_sugar_thms option * gfp_sugar_thms option);
blanchet@55708
    41
blanchet@55708
    42
structure Data = Generic_Data
blanchet@55708
    43
(
blanchet@55708
    44
  type T = n2m_sugar Typtab.table;
blanchet@55708
    45
  val empty = Typtab.empty;
blanchet@55708
    46
  val extend = I;
blanchet@55708
    47
  val merge = Typtab.merge (eq_fst (eq_list eq_fp_sugar));
blanchet@55708
    48
);
blanchet@55708
    49
blanchet@55708
    50
fun morph_n2m_sugar phi (fp_sugars, (lfp_sugar_thms_opt, gfp_sugar_thms_opt)) =
blanchet@55708
    51
  (map (morph_fp_sugar phi) fp_sugars,
blanchet@55708
    52
   (Option.map (morph_lfp_sugar_thms phi) lfp_sugar_thms_opt,
blanchet@55708
    53
    Option.map (morph_gfp_sugar_thms phi) gfp_sugar_thms_opt));
blanchet@55708
    54
blanchet@55708
    55
val transfer_n2m_sugar =
blanchet@55708
    56
  morph_n2m_sugar o Morphism.thm_morphism o Thm.transfer o Proof_Context.theory_of;
blanchet@55708
    57
blanchet@55708
    58
fun n2m_sugar_of ctxt =
blanchet@55708
    59
  Typtab.lookup (Data.get (Context.Proof ctxt))
blanchet@55708
    60
  #> Option.map (transfer_n2m_sugar ctxt);
blanchet@55708
    61
blanchet@55708
    62
fun register_n2m_sugar key n2m_sugar =
blanchet@55708
    63
  Local_Theory.declaration {syntax = false, pervasive = false}
blanchet@55708
    64
    (fn phi => Data.map (Typtab.default (key, morph_n2m_sugar phi n2m_sugar)));
blanchet@55708
    65
blanchet@55695
    66
fun unfold_let (Const (@{const_name Let}, _) $ arg1 $ arg2) = unfold_let (betapply (arg2, arg1))
blanchet@55695
    67
  | unfold_let (Const (@{const_name prod_case}, _) $ t) =
blanchet@55695
    68
    (case unfold_let t of
blanchet@55695
    69
      t' as Abs (s1, T1, Abs (s2, T2, _)) =>
blanchet@55695
    70
      let
blanchet@55695
    71
        val x = (s1 ^ s2, Term.maxidx_of_term t + 1);
blanchet@55695
    72
        val v = Var (x, HOLogic.mk_prodT (T1, T2));
blanchet@55695
    73
      in
blanchet@55695
    74
        lambda v (unfold_let (betapplys (t', [HOLogic.mk_fst v, HOLogic.mk_snd v])))
blanchet@55695
    75
      end
blanchet@55695
    76
    | _ => t)
blanchet@55695
    77
  | unfold_let (t $ u) = betapply (unfold_let t, unfold_let u)
blanchet@55695
    78
  | unfold_let (Abs (s, T, t)) = Abs (s, T, unfold_let t)
blanchet@55695
    79
  | unfold_let t = t;
blanchet@55695
    80
blanchet@55695
    81
fun mk_map_pattern ctxt s =
blanchet@55695
    82
  let
blanchet@55695
    83
    val bnf = the (bnf_of ctxt s);
blanchet@55695
    84
    val mapx = map_of_bnf bnf;
blanchet@55695
    85
    val live = live_of_bnf bnf;
blanchet@55695
    86
    val (f_Ts, _) = strip_typeN live (fastype_of mapx);
blanchet@55717
    87
    val fs = map_index (fn (i, T) => Var (("?f", i), T)) f_Ts;
blanchet@55695
    88
  in
blanchet@55695
    89
    (mapx, betapplys (mapx, fs))
blanchet@55695
    90
  end;
blanchet@55695
    91
blanchet@55695
    92
fun dest_map ctxt s call =
blanchet@55695
    93
  let
blanchet@55695
    94
    val (map0, pat) = mk_map_pattern ctxt s;
blanchet@55695
    95
    val (_, tenv) = fo_match ctxt call pat;
blanchet@55695
    96
  in
blanchet@55695
    97
    (map0, Vartab.fold_rev (fn (_, (_, f)) => cons f) tenv [])
blanchet@55695
    98
  end;
blanchet@55695
    99
blanchet@55728
   100
fun dest_abs_or_applied_map _ _ (Abs (_, _, t)) = (Term.dummy, [t])
blanchet@55728
   101
  | dest_abs_or_applied_map ctxt s (t1 $ _) = dest_map ctxt s t1;
blanchet@55691
   102
blanchet@55697
   103
fun map_partition f xs =
blanchet@55697
   104
  fold_rev (fn x => fn (ys, (good, bad)) =>
blanchet@55697
   105
      case f x of SOME y => (y :: ys, (x :: good, bad)) | NONE => (ys, (good, x :: bad)))
blanchet@55697
   106
    xs ([], ([], []));
blanchet@55697
   107
blanchet@55708
   108
fun key_of_fp_eqs fp fpTs fp_eqs =
blanchet@55719
   109
  Type (fp_case fp "l" "g", fpTs @ maps (fn (x, T) => [TFree x, T]) fp_eqs);
blanchet@55708
   110
blanchet@54440
   111
(* TODO: test with sort constraints on As *)
blanchet@54440
   112
(* TODO: use right sorting order for "fp_sort" w.r.t. original BNFs (?) -- treat new variables
blanchet@54440
   113
   as deads? *)
blanchet@55697
   114
fun mutualize_fp_sugars has_nested fp bs fpTs get_indices callssss fp_sugars0 no_defs_lthy0 =
blanchet@55705
   115
  if has_nested then
blanchet@54440
   116
    let
blanchet@54440
   117
      val thy = Proof_Context.theory_of no_defs_lthy0;
blanchet@54440
   118
blanchet@54440
   119
      val qsotm = quote o Syntax.string_of_term no_defs_lthy0;
blanchet@54440
   120
blanchet@55697
   121
      fun incompatible_calls t1 t2 =
blanchet@55697
   122
        error ("Incompatible " ^ co_prefix fp ^ "recursive calls: " ^ qsotm t1 ^ " vs. " ^
blanchet@55697
   123
          qsotm t2);
blanchet@55697
   124
blanchet@54440
   125
      val b_names = map Binding.name_of bs;
blanchet@54440
   126
      val fp_b_names = map base_name_of_typ fpTs;
blanchet@54440
   127
blanchet@54440
   128
      val nn = length fpTs;
blanchet@54440
   129
wenzelm@55111
   130
      fun target_ctr_sugar_of_fp_sugar fpT ({T, index, ctr_sugars, ...} : fp_sugar) =
blanchet@54440
   131
        let
blanchet@54440
   132
          val rho = Vartab.fold (cons o apsnd snd) (Sign.typ_match thy (T, fpT) Vartab.empty) [];
blanchet@54440
   133
          val phi = Morphism.term_morphism (Term.subst_TVars rho);
blanchet@54440
   134
        in
blanchet@54440
   135
          morph_ctr_sugar phi (nth ctr_sugars index)
blanchet@54440
   136
        end;
blanchet@54440
   137
blanchet@54440
   138
      val ctr_defss = map (of_fp_sugar #ctr_defss) fp_sugars0;
blanchet@54613
   139
      val mapss = map (of_fp_sugar #mapss) fp_sugars0;
blanchet@54440
   140
      val ctr_sugars0 = map2 target_ctr_sugar_of_fp_sugar fpTs fp_sugars0;
blanchet@54440
   141
blanchet@54440
   142
      val ctrss = map #ctrs ctr_sugars0;
blanchet@54440
   143
      val ctr_Tss = map (map fastype_of) ctrss;
blanchet@54440
   144
blanchet@54440
   145
      val As' = fold (fold Term.add_tfreesT) ctr_Tss [];
blanchet@54440
   146
      val As = map TFree As';
blanchet@54440
   147
blanchet@54440
   148
      val ((Cs, Xs), no_defs_lthy) =
blanchet@54440
   149
        no_defs_lthy0
blanchet@54440
   150
        |> fold Variable.declare_typ As
blanchet@54440
   151
        |> mk_TFrees nn
blanchet@54440
   152
        ||>> variant_tfrees fp_b_names;
blanchet@54440
   153
blanchet@55697
   154
      fun check_call_dead live_call call =
blanchet@55697
   155
        if null (get_indices call) then () else incompatible_calls live_call call;
blanchet@55697
   156
blanchet@55707
   157
      fun freeze_fpTs_simple (T as Type (s, Ts)) =
blanchet@55705
   158
          (case find_index (curry (op =) T) fpTs of
blanchet@55707
   159
            ~1 => Type (s, map freeze_fpTs_simple Ts)
blanchet@55705
   160
          | kk => nth Xs kk)
blanchet@55707
   161
        | freeze_fpTs_simple T = T;
blanchet@55705
   162
blanchet@55707
   163
      fun freeze_fpTs_map (callss, (live_call :: _, dead_calls)) s Ts =
blanchet@55697
   164
        (List.app (check_call_dead live_call) dead_calls;
blanchet@55707
   165
         Type (s, map2 freeze_fpTs (flatten_type_args_of_bnf (the (bnf_of no_defs_lthy s)) []
blanchet@55697
   166
           (transpose callss)) Ts))
blanchet@55707
   167
      and freeze_fpTs calls (T as Type (s, Ts)) =
blanchet@55697
   168
          (case map_partition (try (snd o dest_map no_defs_lthy s)) calls of
blanchet@55697
   169
            ([], _) =>
blanchet@55728
   170
            (case map_partition (try (snd o dest_abs_or_applied_map no_defs_lthy s)) calls of
blanchet@55707
   171
              ([], _) => freeze_fpTs_simple T
blanchet@55707
   172
            | callsp => freeze_fpTs_map callsp s Ts)
blanchet@55707
   173
          | callsp => freeze_fpTs_map callsp s Ts)
blanchet@55707
   174
        | freeze_fpTs _ T = T;
blanchet@54440
   175
blanchet@54440
   176
      val ctr_Tsss = map (map binder_types) ctr_Tss;
blanchet@55707
   177
      val ctrXs_Tsss = map2 (map2 (map2 freeze_fpTs)) callssss ctr_Tsss;
blanchet@54440
   178
      val ctrXs_sum_prod_Ts = map (mk_sumTN_balanced o map HOLogic.mk_tupleT) ctrXs_Tsss;
blanchet@54440
   179
      val Ts = map (body_type o hd) ctr_Tss;
blanchet@54440
   180
blanchet@54440
   181
      val ns = map length ctr_Tsss;
blanchet@54440
   182
      val kss = map (fn n => 1 upto n) ns;
blanchet@54440
   183
      val mss = map (map length) ctr_Tsss;
blanchet@54440
   184
blanchet@54440
   185
      val fp_eqs = map dest_TFree Xs ~~ ctrXs_sum_prod_Ts;
blanchet@55708
   186
      val key = key_of_fp_eqs fp fpTs fp_eqs;
blanchet@55708
   187
    in
blanchet@55708
   188
      (case n2m_sugar_of no_defs_lthy key of
blanchet@55708
   189
        SOME n2m_sugar => (n2m_sugar, no_defs_lthy)
blanchet@55708
   190
      | NONE =>
blanchet@55708
   191
        let
blanchet@55708
   192
          val base_fp_names = Name.variant_list [] fp_b_names;
blanchet@55708
   193
          val fp_bs = map2 (fn b_name => fn base_fp_name =>
blanchet@55708
   194
              Binding.qualify true b_name (Binding.name (n2mN ^ base_fp_name)))
blanchet@55708
   195
            b_names base_fp_names;
blanchet@54440
   196
blanchet@55708
   197
          val (pre_bnfs, (fp_res as {xtor_co_iterss = xtor_co_iterss0, xtor_co_induct,
blanchet@55708
   198
                 dtor_injects, dtor_ctors, xtor_co_iter_thmss, ...}, lthy)) =
blanchet@55708
   199
            fp_bnf (construct_mutualized_fp fp fpTs fp_sugars0) fp_bs As' fp_eqs no_defs_lthy;
blanchet@54440
   200
blanchet@55708
   201
          val nesting_bnfs = nesty_bnfs lthy ctrXs_Tsss As;
blanchet@55708
   202
          val nested_bnfs = nesty_bnfs lthy ctrXs_Tsss Xs;
blanchet@54440
   203
blanchet@55708
   204
          val ((xtor_co_iterss, iters_args_types, coiters_args_types), _) =
blanchet@55708
   205
            mk_co_iters_prelims fp ctr_Tsss fpTs Cs ns mss xtor_co_iterss0 lthy;
blanchet@54440
   206
blanchet@55708
   207
          fun mk_binding b suf = Binding.suffix_name ("_" ^ suf) b;
blanchet@54440
   208
blanchet@55708
   209
          val ((co_iterss, co_iter_defss), lthy) =
blanchet@55708
   210
            fold_map2 (fn b =>
blanchet@55708
   211
              (if fp = Least_FP then define_iters [foldN, recN] (the iters_args_types)
blanchet@55708
   212
               else define_coiters [unfoldN, corecN] (the coiters_args_types))
blanchet@55708
   213
                (mk_binding b) fpTs Cs) fp_bs xtor_co_iterss lthy
blanchet@55708
   214
            |>> split_list;
blanchet@54440
   215
blanchet@55708
   216
          val rho = tvar_subst thy Ts fpTs;
blanchet@55708
   217
          val ctr_sugar_phi =
blanchet@55708
   218
            Morphism.compose (Morphism.typ_morphism (Term.typ_subst_TVars rho))
blanchet@55708
   219
              (Morphism.term_morphism (Term.subst_TVars rho));
blanchet@55708
   220
          val inst_ctr_sugar = morph_ctr_sugar ctr_sugar_phi;
blanchet@54440
   221
blanchet@55708
   222
          val ctr_sugars = map inst_ctr_sugar ctr_sugars0;
blanchet@54440
   223
blanchet@55708
   224
          val ((co_inducts, un_fold_thmss, co_rec_thmss, disc_unfold_thmss, disc_corec_thmss,
blanchet@55708
   225
                sel_unfold_thmsss, sel_corec_thmsss), fp_sugar_thms) =
blanchet@55708
   226
            if fp = Least_FP then
blanchet@55708
   227
              derive_induct_iters_thms_for_types pre_bnfs (the iters_args_types) xtor_co_induct
blanchet@55708
   228
                xtor_co_iter_thmss nesting_bnfs nested_bnfs fpTs Cs Xs ctrXs_Tsss ctrss ctr_defss
blanchet@55708
   229
                co_iterss co_iter_defss lthy
blanchet@55708
   230
              |> `(fn ((_, induct, _), (fold_thmss, rec_thmss, _)) =>
blanchet@55708
   231
                ([induct], fold_thmss, rec_thmss, [], [], [], []))
blanchet@55708
   232
              ||> (fn info => (SOME info, NONE))
blanchet@55708
   233
            else
blanchet@55708
   234
              derive_coinduct_coiters_thms_for_types pre_bnfs (the coiters_args_types)
blanchet@55708
   235
                xtor_co_induct dtor_injects dtor_ctors xtor_co_iter_thmss nesting_bnfs fpTs Cs Xs
blanchet@55708
   236
                ctrXs_Tsss kss mss ns ctr_defss ctr_sugars co_iterss co_iter_defss
blanchet@55708
   237
                (Proof_Context.export lthy no_defs_lthy) lthy
blanchet@55708
   238
              |> `(fn ((coinduct_thms_pairs, _), (unfold_thmss, corec_thmss, _),
blanchet@55708
   239
                      (disc_unfold_thmss, disc_corec_thmss, _), _,
blanchet@55708
   240
                      (sel_unfold_thmsss, sel_corec_thmsss, _)) =>
blanchet@55708
   241
                (map snd coinduct_thms_pairs, unfold_thmss, corec_thmss, disc_unfold_thmss,
blanchet@55708
   242
                 disc_corec_thmss, sel_unfold_thmsss, sel_corec_thmsss))
blanchet@55708
   243
              ||> (fn info => (NONE, SOME info));
blanchet@54440
   244
blanchet@55708
   245
          val phi = Proof_Context.export_morphism no_defs_lthy no_defs_lthy0;
blanchet@54440
   246
blanchet@55708
   247
          fun mk_target_fp_sugar (kk, T) =
blanchet@55708
   248
            {T = T, fp = fp, index = kk, pre_bnfs = pre_bnfs, nested_bnfs = nested_bnfs,
blanchet@55708
   249
             nesting_bnfs = nesting_bnfs, fp_res = fp_res, ctr_defss = ctr_defss,
blanchet@55708
   250
             ctr_sugars = ctr_sugars, co_iterss = co_iterss, mapss = mapss, co_inducts = co_inducts,
blanchet@55708
   251
             co_iter_thmsss = transpose [un_fold_thmss, co_rec_thmss],
blanchet@55708
   252
             disc_co_itersss = transpose [disc_unfold_thmss, disc_corec_thmss],
blanchet@55708
   253
             sel_co_iterssss = transpose [sel_unfold_thmsss, sel_corec_thmsss]}
blanchet@55708
   254
            |> morph_fp_sugar phi;
blanchet@54440
   255
blanchet@55708
   256
          val n2m_sugar = (map_index mk_target_fp_sugar fpTs, fp_sugar_thms);
blanchet@55708
   257
        in
blanchet@55708
   258
          (n2m_sugar, lthy |> register_n2m_sugar key n2m_sugar)
blanchet@55708
   259
        end)
blanchet@54440
   260
    end
blanchet@54440
   261
  else
blanchet@54883
   262
    ((fp_sugars0, (NONE, NONE)), no_defs_lthy0);
blanchet@54440
   263
blanchet@54440
   264
fun indexify_callsss fp_sugar callsss =
blanchet@54440
   265
  let
blanchet@54440
   266
    val {ctrs, ...} = of_fp_sugar #ctr_sugars fp_sugar;
blanchet@55735
   267
    fun indexify_ctr ctr =
blanchet@54440
   268
      (case AList.lookup Term.aconv_untyped callsss ctr of
blanchet@54440
   269
        NONE => replicate (num_binder_types (fastype_of ctr)) []
blanchet@55695
   270
      | SOME callss => map (map (Envir.beta_eta_contract o unfold_let)) callss);
blanchet@54440
   271
  in
blanchet@55735
   272
    map indexify_ctr ctrs
blanchet@54440
   273
  end;
blanchet@54440
   274
blanchet@55735
   275
fun retypargs tyargs (Type (s, _)) = Type (s, tyargs);
blanchet@55735
   276
blanchet@55735
   277
fun fold_subtype_pairs f (T as Type (s, Ts), U as Type (s', Us)) =
blanchet@55735
   278
    f (T, U) #> (if s = s' then fold (fold_subtype_pairs f) (Ts ~~ Us) else I)
blanchet@55735
   279
  | fold_subtype_pairs f TU = f TU;
blanchet@55735
   280
blanchet@55146
   281
fun nested_to_mutual_fps fp actual_bs actual_Ts get_indices actual_callssss0 lthy =
blanchet@54440
   282
  let
blanchet@54440
   283
    val qsoty = quote o Syntax.string_of_typ lthy;
blanchet@54440
   284
    val qsotys = space_implode " or " o map qsoty;
blanchet@54440
   285
blanchet@55705
   286
    fun duplicate_datatype T = error (qsoty T ^ " is not mutually recursive with itself");
blanchet@54440
   287
    fun not_co_datatype0 T = error (qsoty T ^ " is not a " ^ co_prefix fp ^ "datatype");
blanchet@54440
   288
    fun not_co_datatype (T as Type (s, _)) =
blanchet@54440
   289
        if fp = Least_FP andalso
blanchet@54440
   290
           is_some (Datatype_Data.get_info (Proof_Context.theory_of lthy) s) then
blanchet@54440
   291
          error (qsoty T ^ " is not a new-style datatype (cf. \"datatype_new\")")
blanchet@54440
   292
        else
blanchet@54440
   293
          not_co_datatype0 T
blanchet@54440
   294
      | not_co_datatype T = not_co_datatype0 T;
blanchet@54440
   295
blanchet@55705
   296
    val _ = (case Library.duplicates (op =) actual_Ts of [] => () | T :: _ => duplicate_datatype T);
blanchet@55705
   297
blanchet@55725
   298
    val perm_actual_Ts as Type (_, tyargs0) :: _ =
blanchet@55718
   299
      sort (prod_ord int_ord Term_Ord.typ_ord o pairself (`Term.size_of_typ)) actual_Ts;
blanchet@54440
   300
blanchet@55735
   301
    fun the_ctrs_of (Type (s, Ts)) = map (mk_ctr Ts) (#ctrs (the (ctr_sugar_of lthy s)));
blanchet@55735
   302
blanchet@55734
   303
    fun the_fp_sugar_of (T as Type (T_name, _)) =
blanchet@55734
   304
      (case fp_sugar_of lthy T_name of
blanchet@55734
   305
        SOME (fp_sugar as {fp = fp', ...}) => if fp = fp' then fp_sugar else not_co_datatype T
blanchet@55734
   306
      | NONE => not_co_datatype T);
blanchet@55734
   307
blanchet@55735
   308
    fun gen_rhss_in gen_Ts rho subTs =
blanchet@55735
   309
      let
blanchet@55735
   310
        fun maybe_insert (T, Type (_, gen_tyargs)) =
blanchet@55735
   311
            if member (op =) subTs T then insert (op =) gen_tyargs else I
blanchet@55735
   312
          | maybe_insert _ = I;
blanchet@55735
   313
blanchet@55735
   314
        val ctrs = maps the_ctrs_of gen_Ts;
blanchet@55735
   315
        val gen_ctr_Ts = maps (binder_types o fastype_of) ctrs;
blanchet@55735
   316
        val ctr_Ts = map (Term.typ_subst_atomic rho) gen_ctr_Ts;
blanchet@55735
   317
      in
blanchet@55735
   318
        fold (fold_subtype_pairs maybe_insert) (ctr_Ts ~~ gen_ctr_Ts) []
blanchet@55735
   319
      end;
blanchet@55735
   320
blanchet@55735
   321
    fun check_enrich_with_mutuals _ _ seen gen_seen [] = (seen, gen_seen)
blanchet@55735
   322
      | check_enrich_with_mutuals lthy rho seen gen_seen ((T as Type (_, tyargs)) :: Ts) =
blanchet@55734
   323
        let
blanchet@55735
   324
          val {fp_res = {Ts = mutual_Ts0, ...}, ...} = the_fp_sugar_of T;
blanchet@55735
   325
          val mutual_Ts = map (retypargs tyargs) mutual_Ts0;
blanchet@55735
   326
blanchet@55735
   327
          fun fresh_tyargs () =
blanchet@55735
   328
            let
blanchet@55735
   329
              (* The name "'z" is unlikely to clash with the context, yielding more cache hits. *)
blanchet@55735
   330
              val (gen_tyargs, lthy') =
blanchet@55735
   331
                variant_tfrees (replicate (length tyargs) "z") lthy
blanchet@55735
   332
                |>> map Logic.varifyT_global;
blanchet@55735
   333
              val rho' = (gen_tyargs ~~ tyargs) @ rho;
blanchet@55735
   334
            in
blanchet@55735
   335
              (rho', gen_tyargs, gen_seen, lthy')
blanchet@55735
   336
            end;
blanchet@55735
   337
blanchet@55735
   338
          val (rho', gen_tyargs, gen_seen', lthy') =
blanchet@55735
   339
            if exists (exists_subtype_in seen) mutual_Ts then
blanchet@55735
   340
              (case gen_rhss_in gen_seen rho mutual_Ts of
blanchet@55735
   341
                [] => fresh_tyargs ()
blanchet@55735
   342
              | [gen_tyargs] => (rho, gen_tyargs, gen_seen, lthy)
blanchet@55735
   343
              | gen_tyargss as gen_tyargs :: gen_tyargss_tl =>
blanchet@55735
   344
                let
blanchet@55735
   345
                  val unify_pairs = split_list (maps (curry (op ~~) gen_tyargs) gen_tyargss_tl);
blanchet@55735
   346
                  val mgu = Type.raw_unifys unify_pairs Vartab.empty;
blanchet@55735
   347
                  val gen_tyargs' = map (Envir.subst_type mgu) gen_tyargs;
blanchet@55735
   348
                  val gen_seen' = map (Envir.subst_type mgu) gen_seen;
blanchet@55735
   349
                in
blanchet@55735
   350
                  (rho, gen_tyargs', gen_seen', lthy)
blanchet@55735
   351
                end)
blanchet@55735
   352
            else
blanchet@55735
   353
              fresh_tyargs ();
blanchet@55735
   354
blanchet@55735
   355
          val gen_mutual_Ts = map (retypargs gen_tyargs) mutual_Ts0;
blanchet@55735
   356
          val Ts' = filter_out (member (op =) mutual_Ts) Ts;
blanchet@55734
   357
        in
blanchet@55735
   358
          check_enrich_with_mutuals lthy' rho' (seen @ mutual_Ts) (gen_seen' @ gen_mutual_Ts) Ts'
blanchet@55734
   359
        end
blanchet@55735
   360
      | check_enrich_with_mutuals _ _ _ _ (T :: _) = not_co_datatype T;
blanchet@54440
   361
blanchet@55735
   362
    val (perm_Ts, perm_gen_Ts) = check_enrich_with_mutuals lthy [] [] [] perm_actual_Ts;
blanchet@55735
   363
    val perm_frozen_gen_Ts = map Logic.unvarifyT_global perm_gen_Ts;
blanchet@55735
   364
blanchet@54440
   365
    val missing_Ts = perm_Ts |> subtract (op =) actual_Ts;
blanchet@54440
   366
    val Ts = actual_Ts @ missing_Ts;
blanchet@54440
   367
blanchet@54440
   368
    val nn = length Ts;
blanchet@54440
   369
    val kks = 0 upto nn - 1;
blanchet@54440
   370
blanchet@55719
   371
    val callssss0 = pad_list [] nn actual_callssss0;
blanchet@55719
   372
blanchet@54440
   373
    val common_name = mk_common_name (map Binding.name_of actual_bs);
blanchet@54440
   374
    val bs = pad_list (Binding.name common_name) nn actual_bs;
blanchet@54440
   375
blanchet@54440
   376
    fun permute xs = permute_like (op =) Ts perm_Ts xs;
blanchet@54440
   377
    fun unpermute perm_xs = permute_like (op =) perm_Ts Ts perm_xs;
blanchet@54440
   378
blanchet@54440
   379
    val perm_bs = permute bs;
blanchet@54440
   380
    val perm_kks = permute kks;
blanchet@55719
   381
    val perm_callssss0 = permute callssss0;
blanchet@54440
   382
    val perm_fp_sugars0 = map (the o fp_sugar_of lthy o fst o dest_Type) perm_Ts;
blanchet@54440
   383
blanchet@55725
   384
    val has_nested = exists (fn Type (_, tyargs) => tyargs <> tyargs0) Ts;
blanchet@55719
   385
    val perm_callssss = map2 indexify_callsss perm_fp_sugars0 perm_callssss0;
blanchet@54440
   386
blanchet@54440
   387
    val get_perm_indices = map (fn kk => find_index (curry (op =) kk) perm_kks) o get_indices;
blanchet@54440
   388
blanchet@54883
   389
    val ((perm_fp_sugars, fp_sugar_thms), lthy) =
blanchet@55735
   390
      mutualize_fp_sugars has_nested fp perm_bs perm_frozen_gen_Ts get_perm_indices perm_callssss
blanchet@54440
   391
        perm_fp_sugars0 lthy;
blanchet@54440
   392
blanchet@54440
   393
    val fp_sugars = unpermute perm_fp_sugars;
blanchet@54440
   394
  in
blanchet@54883
   395
    ((missing_Ts, perm_kks, fp_sugars, fp_sugar_thms), lthy)
blanchet@54440
   396
  end;
blanchet@54440
   397
blanchet@54440
   398
end;