src/HOL/Tools/BNF/Tools/bnf_fp_n2m_sugar.ML
author blanchet
Mon, 20 Jan 2014 18:24:56 +0100
changeset 55058 4e700eb471d4
parent 54957 src/HOL/BNF/Tools/bnf_fp_n2m_sugar.ML@99eebac5fcb3
permissions -rw-r--r--
moved BNF files to 'HOL'
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/BNF/Tools/bnf_fp_n2m_sugar.ML
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     3
    Copyright   2013
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     4
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     5
Suggared flattening of nested to mutual (co)recursion.
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     6
*)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     7
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     8
signature BNF_FP_N2M_SUGAR =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     9
sig
54243
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    10
  val unfold_let: term -> term
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    11
  val dest_map: Proof.context -> string -> term -> term * term list
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    12
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
    13
  val mutualize_fp_sugars: BNF_FP_Util.fp_kind -> binding list -> typ list -> (term -> int list) ->
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
    14
    term list list list list -> BNF_FP_Def_Sugar.fp_sugar list -> local_theory ->
53746
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53678
diff changeset
    15
    (BNF_FP_Def_Sugar.fp_sugar list
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53678
diff changeset
    16
     * (BNF_FP_Def_Sugar.lfp_sugar_thms option * BNF_FP_Def_Sugar.gfp_sugar_thms option))
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53678
diff changeset
    17
    * local_theory
54267
78e8a178b690 use right permutation in 'map2'
blanchet
parents: 54266
diff changeset
    18
  val indexify_callsss: BNF_FP_Def_Sugar.fp_sugar -> (term * term list list) list ->
78e8a178b690 use right permutation in 'map2'
blanchet
parents: 54266
diff changeset
    19
    term list list list
54009
f138452e8265 got rid of dead feature
blanchet
parents: 54006
diff changeset
    20
  val nested_to_mutual_fps: BNF_FP_Util.fp_kind -> binding list -> typ list -> (term -> int list) ->
f138452e8265 got rid of dead feature
blanchet
parents: 54006
diff changeset
    21
    (term * term list list) list list -> local_theory ->
53746
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53678
diff changeset
    22
    (typ list * int list * BNF_FP_Def_Sugar.fp_sugar list
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53678
diff changeset
    23
     * (BNF_FP_Def_Sugar.lfp_sugar_thms option * BNF_FP_Def_Sugar.gfp_sugar_thms option))
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53678
diff changeset
    24
    * local_theory
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    25
end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    26
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    27
structure BNF_FP_N2M_Sugar : BNF_FP_N2M_SUGAR =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    28
struct
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    29
54006
9fe1bd54d437 renamed theory file
blanchet
parents: 53974
diff changeset
    30
open Ctr_Sugar
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    31
open BNF_Util
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    32
open BNF_Def
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    33
open BNF_FP_Util
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    34
open BNF_FP_Def_Sugar
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    35
open BNF_FP_N2M
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    36
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    37
val n2mN = "n2m_"
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    38
54256
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    39
type n2m_sugar = fp_sugar list * (lfp_sugar_thms option * gfp_sugar_thms option);
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    40
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    41
structure Data = Generic_Data
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    42
(
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    43
  type T = n2m_sugar Typtab.table;
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    44
  val empty = Typtab.empty;
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    45
  val extend = I;
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    46
  val merge = Typtab.merge (eq_fst (eq_list eq_fp_sugar));
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    47
);
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    48
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    49
fun morph_n2m_sugar phi (fp_sugars, (lfp_sugar_thms_opt, gfp_sugar_thms_opt)) =
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    50
  (map (morph_fp_sugar phi) fp_sugars,
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    51
   (Option.map (morph_lfp_sugar_thms phi) lfp_sugar_thms_opt,
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    52
    Option.map (morph_gfp_sugar_thms phi) gfp_sugar_thms_opt));
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    53
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    54
val transfer_n2m_sugar =
54740
91f54d386680 maintain morphism names for diagnostic purposes;
wenzelm
parents: 54303
diff changeset
    55
  morph_n2m_sugar o Morphism.transfer_morphism o Proof_Context.theory_of;
54256
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    56
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    57
fun n2m_sugar_of ctxt =
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    58
  Typtab.lookup (Data.get (Context.Proof ctxt))
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    59
  #> Option.map (transfer_n2m_sugar ctxt);
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    60
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    61
fun register_n2m_sugar key n2m_sugar =
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    62
  Local_Theory.declaration {syntax = false, pervasive = false}
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    63
    (fn phi => Data.map (Typtab.default (key, morph_n2m_sugar phi n2m_sugar)));
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
    64
54243
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    65
fun unfold_let (Const (@{const_name Let}, _) $ arg1 $ arg2) = unfold_let (betapply (arg2, arg1))
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    66
  | unfold_let (Const (@{const_name prod_case}, _) $ t) =
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    67
    (case unfold_let t of
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    68
      t' as Abs (s1, T1, Abs (s2, T2, _)) =>
54957
99eebac5fcb3 fixed de Bruijn bug in 'unfold_lets'
blanchet
parents: 54899
diff changeset
    69
      let val v = Var ((s1 ^ s2, Term.maxidx_of_term t' + 1), HOLogic.mk_prodT (T1, T2)) in
99eebac5fcb3 fixed de Bruijn bug in 'unfold_lets'
blanchet
parents: 54899
diff changeset
    70
        lambda v (incr_boundvars 1 (betapplys (t', [HOLogic.mk_fst v, HOLogic.mk_snd v])))
54243
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    71
      end
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    72
    | _ => t)
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    73
  | unfold_let (t $ u) = betapply (unfold_let t, unfold_let u)
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    74
  | unfold_let (Abs (s, T, t)) = Abs (s, T, unfold_let t)
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    75
  | unfold_let t = t;
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    76
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    77
fun mk_map_pattern ctxt s =
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    78
  let
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    79
    val bnf = the (bnf_of ctxt s);
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    80
    val mapx = map_of_bnf bnf;
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    81
    val live = live_of_bnf bnf;
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    82
    val (f_Ts, _) = strip_typeN live (fastype_of mapx);
54265
3e1d230f1c00 make local theory operations non-pervasive (makes more intuitive sense)
blanchet
parents: 54256
diff changeset
    83
    val fs = map_index (fn (i, T) => Var (("?f", i), T)) f_Ts;
54243
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    84
  in
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    85
    (mapx, betapplys (mapx, fs))
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    86
  end;
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    87
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    88
fun dest_map ctxt s call =
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    89
  let
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    90
    val (map0, pat) = mk_map_pattern ctxt s;
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    91
    val (_, tenv) = fo_match ctxt call pat;
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    92
  in
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    93
    (map0, Vartab.fold_rev (fn (_, (_, f)) => cons f) tenv [])
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    94
  end;
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
    95
54276
d26b6b935a6f took out loopy code
blanchet
parents: 54275
diff changeset
    96
fun dest_abs_or_applied_map _ _ (Abs (_, _, t)) = (Term.dummy, [t])
d26b6b935a6f took out loopy code
blanchet
parents: 54275
diff changeset
    97
  | dest_abs_or_applied_map ctxt s (t1 $ _) = dest_map ctxt s t1;
54239
9bd91d5d8a7b handle constructor syntax in n2m primcorec
blanchet
parents: 54235
diff changeset
    98
54245
f91022745c85 better error handling
blanchet
parents: 54243
diff changeset
    99
fun map_partition f xs =
f91022745c85 better error handling
blanchet
parents: 54243
diff changeset
   100
  fold_rev (fn x => fn (ys, (good, bad)) =>
f91022745c85 better error handling
blanchet
parents: 54243
diff changeset
   101
      case f x of SOME y => (y :: ys, (x :: good, bad)) | NONE => (ys, (good, x :: bad)))
f91022745c85 better error handling
blanchet
parents: 54243
diff changeset
   102
    xs ([], ([], []));
f91022745c85 better error handling
blanchet
parents: 54243
diff changeset
   103
54256
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
   104
fun key_of_fp_eqs fp fpTs fp_eqs =
54267
78e8a178b690 use right permutation in 'map2'
blanchet
parents: 54266
diff changeset
   105
  Type (fp_case fp "l" "g", fpTs @ maps (fn (x, T) => [TFree x, T]) fp_eqs);
54256
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
   106
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   107
(* TODO: test with sort constraints on As *)
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   108
fun mutualize_fp_sugars fp bs fpTs get_indices callssss fp_sugars0 no_defs_lthy0 =
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   109
  let
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   110
    val thy = Proof_Context.theory_of no_defs_lthy0;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   111
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   112
    val qsotm = quote o Syntax.string_of_term no_defs_lthy0;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   113
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   114
    fun incompatible_calls t1 t2 =
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   115
      error ("Incompatible " ^ co_prefix fp ^ "recursive calls: " ^ qsotm t1 ^ " vs. " ^ qsotm t2);
54301
e0943a2ee400 added check to avoid odd situations the N2M code cannot handle
blanchet
parents: 54300
diff changeset
   116
    fun nested_self_call t =
e0943a2ee400 added check to avoid odd situations the N2M code cannot handle
blanchet
parents: 54300
diff changeset
   117
      error ("Unsupported nested self-call " ^ qsotm t);
54245
f91022745c85 better error handling
blanchet
parents: 54243
diff changeset
   118
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   119
    val b_names = map Binding.name_of bs;
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   120
    val fp_b_names = map base_name_of_typ fpTs;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   121
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   122
    val nn = length fpTs;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   123
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   124
    fun target_ctr_sugar_of_fp_sugar fpT ({T, index, ctr_sugars, ...} : fp_sugar) =
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   125
      let
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   126
        val rho = Vartab.fold (cons o apsnd snd) (Sign.typ_match thy (T, fpT) Vartab.empty) [];
54740
91f54d386680 maintain morphism names for diagnostic purposes;
wenzelm
parents: 54303
diff changeset
   127
        val phi = Morphism.term_morphism "BNF" (Term.subst_TVars rho);
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   128
      in
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   129
        morph_ctr_sugar phi (nth ctr_sugars index)
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   130
      end;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   131
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   132
    val ctr_defss = map (of_fp_sugar #ctr_defss) fp_sugars0;
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   133
    val mapss = map (of_fp_sugar #mapss) fp_sugars0;
54302
880e5417b800 removed dead code
blanchet
parents: 54301
diff changeset
   134
    val ctr_sugars = map2 target_ctr_sugar_of_fp_sugar fpTs fp_sugars0;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   135
54302
880e5417b800 removed dead code
blanchet
parents: 54301
diff changeset
   136
    val ctrss = map #ctrs ctr_sugars;
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   137
    val ctr_Tss = map (map fastype_of) ctrss;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   138
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   139
    val As' = fold (fold Term.add_tfreesT) ctr_Tss [];
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   140
    val As = map TFree As';
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   141
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   142
    val ((Cs, Xs), no_defs_lthy) =
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   143
      no_defs_lthy0
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   144
      |> fold Variable.declare_typ As
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   145
      |> mk_TFrees nn
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   146
      ||>> variant_tfrees fp_b_names;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   147
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   148
    fun check_call_dead live_call call =
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   149
      if null (get_indices call) then () else incompatible_calls live_call call;
54245
f91022745c85 better error handling
blanchet
parents: 54243
diff changeset
   150
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   151
    fun freeze_fpTs_simple (T as Type (s, Ts)) =
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   152
        (case find_index (curry (op =) T) fpTs of
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   153
          ~1 => Type (s, map freeze_fpTs_simple Ts)
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   154
        | kk => nth Xs kk)
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   155
      | freeze_fpTs_simple T = T;
54253
04cd231e2b9e nicer error message in case of duplicates
blanchet
parents: 54245
diff changeset
   156
54301
e0943a2ee400 added check to avoid odd situations the N2M code cannot handle
blanchet
parents: 54300
diff changeset
   157
    fun freeze_fpTs_map (fpT as Type (_, Ts')) (callss, (live_call :: _, dead_calls))
e0943a2ee400 added check to avoid odd situations the N2M code cannot handle
blanchet
parents: 54300
diff changeset
   158
        (T as Type (s, Ts)) =
e0943a2ee400 added check to avoid odd situations the N2M code cannot handle
blanchet
parents: 54300
diff changeset
   159
      if Ts' = Ts then
e0943a2ee400 added check to avoid odd situations the N2M code cannot handle
blanchet
parents: 54300
diff changeset
   160
        nested_self_call live_call
e0943a2ee400 added check to avoid odd situations the N2M code cannot handle
blanchet
parents: 54300
diff changeset
   161
      else
e0943a2ee400 added check to avoid odd situations the N2M code cannot handle
blanchet
parents: 54300
diff changeset
   162
        (List.app (check_call_dead live_call) dead_calls;
e0943a2ee400 added check to avoid odd situations the N2M code cannot handle
blanchet
parents: 54300
diff changeset
   163
         Type (s, map2 (freeze_fpTs fpT) (flatten_type_args_of_bnf (the (bnf_of no_defs_lthy s)) []
e0943a2ee400 added check to avoid odd situations the N2M code cannot handle
blanchet
parents: 54300
diff changeset
   164
           (transpose callss)) Ts))
e0943a2ee400 added check to avoid odd situations the N2M code cannot handle
blanchet
parents: 54300
diff changeset
   165
    and freeze_fpTs fpT calls (T as Type (s, _)) =
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   166
        (case map_partition (try (snd o dest_map no_defs_lthy s)) calls of
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   167
          ([], _) =>
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   168
          (case map_partition (try (snd o dest_abs_or_applied_map no_defs_lthy s)) calls of
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   169
            ([], _) => freeze_fpTs_simple T
54301
e0943a2ee400 added check to avoid odd situations the N2M code cannot handle
blanchet
parents: 54300
diff changeset
   170
          | callsp => freeze_fpTs_map fpT callsp T)
e0943a2ee400 added check to avoid odd situations the N2M code cannot handle
blanchet
parents: 54300
diff changeset
   171
        | callsp => freeze_fpTs_map fpT callsp T)
e0943a2ee400 added check to avoid odd situations the N2M code cannot handle
blanchet
parents: 54300
diff changeset
   172
      | freeze_fpTs _ _ T = T;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   173
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   174
    val ctr_Tsss = map (map binder_types) ctr_Tss;
54301
e0943a2ee400 added check to avoid odd situations the N2M code cannot handle
blanchet
parents: 54300
diff changeset
   175
    val ctrXs_Tsss = map3 (map2 o map2 o freeze_fpTs) fpTs callssss ctr_Tsss;
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   176
    val ctrXs_sum_prod_Ts = map (mk_sumTN_balanced o map HOLogic.mk_tupleT) ctrXs_Tsss;
54301
e0943a2ee400 added check to avoid odd situations the N2M code cannot handle
blanchet
parents: 54300
diff changeset
   177
    val ctr_Ts = map (body_type o hd) ctr_Tss;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   178
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   179
    val ns = map length ctr_Tsss;
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   180
    val kss = map (fn n => 1 upto n) ns;
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   181
    val mss = map (map length) ctr_Tsss;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   182
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   183
    val fp_eqs = map dest_TFree Xs ~~ ctrXs_sum_prod_Ts;
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   184
    val key = key_of_fp_eqs fp fpTs fp_eqs;
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   185
  in
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   186
    (case n2m_sugar_of no_defs_lthy key of
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   187
      SOME n2m_sugar => (n2m_sugar, no_defs_lthy)
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   188
    | NONE =>
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   189
      let
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   190
        val base_fp_names = Name.variant_list [] fp_b_names;
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   191
        val fp_bs = map2 (fn b_name => fn base_fp_name =>
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   192
            Binding.qualify true b_name (Binding.name (n2mN ^ base_fp_name)))
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   193
          b_names base_fp_names;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   194
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   195
        val (pre_bnfs, (fp_res as {xtor_co_iterss = xtor_co_iterss0, xtor_co_induct, dtor_injects,
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   196
               dtor_ctors, xtor_co_iter_thmss, ...}, lthy)) =
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   197
          fp_bnf (construct_mutualized_fp fp fpTs fp_sugars0) fp_bs As' fp_eqs no_defs_lthy;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   198
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   199
        val nesting_bnfs = nesty_bnfs lthy ctrXs_Tsss As;
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   200
        val nested_bnfs = nesty_bnfs lthy ctrXs_Tsss Xs;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   201
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   202
        val ((xtor_co_iterss, iters_args_types, coiters_args_types), _) =
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   203
          mk_co_iters_prelims fp ctr_Tsss fpTs Cs ns mss xtor_co_iterss0 lthy;
54256
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
   204
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   205
        fun mk_binding b suf = Binding.suffix_name ("_" ^ suf) b;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   206
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   207
        val ((co_iterss, co_iter_defss), lthy) =
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   208
          fold_map2 (fn b =>
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   209
            (if fp = Least_FP then define_iters [foldN, recN] (the iters_args_types)
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   210
             else define_coiters [unfoldN, corecN] (the coiters_args_types))
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   211
              (mk_binding b) fpTs Cs) fp_bs xtor_co_iterss lthy
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   212
          |>> split_list;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   213
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   214
        val ((co_inducts, un_fold_thmss, co_rec_thmss, disc_unfold_thmss, disc_corec_thmss,
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   215
              sel_unfold_thmsss, sel_corec_thmsss), fp_sugar_thms) =
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   216
          if fp = Least_FP then
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   217
            derive_induct_iters_thms_for_types pre_bnfs (the iters_args_types) xtor_co_induct
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   218
              xtor_co_iter_thmss nesting_bnfs nested_bnfs fpTs Cs Xs ctrXs_Tsss ctrss ctr_defss
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   219
              co_iterss co_iter_defss lthy
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   220
            |> `(fn ((_, induct, _), (fold_thmss, rec_thmss, _)) =>
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   221
              ([induct], fold_thmss, rec_thmss, [], [], [], []))
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   222
            ||> (fn info => (SOME info, NONE))
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   223
          else
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   224
            derive_coinduct_coiters_thms_for_types pre_bnfs (the coiters_args_types) xtor_co_induct
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   225
              dtor_injects dtor_ctors xtor_co_iter_thmss nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   226
              ns ctr_defss ctr_sugars co_iterss co_iter_defss
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   227
              (Proof_Context.export lthy no_defs_lthy) lthy
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   228
            |> `(fn ((coinduct_thms_pairs, _), (unfold_thmss, corec_thmss, _),
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   229
                    (disc_unfold_thmss, disc_corec_thmss, _), _,
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   230
                    (sel_unfold_thmsss, sel_corec_thmsss, _)) =>
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   231
              (map snd coinduct_thms_pairs, unfold_thmss, corec_thmss, disc_unfold_thmss,
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   232
               disc_corec_thmss, sel_unfold_thmsss, sel_corec_thmsss))
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   233
            ||> (fn info => (NONE, SOME info));
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   234
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   235
        val phi = Proof_Context.export_morphism no_defs_lthy no_defs_lthy0;
54256
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
   236
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   237
        fun mk_target_fp_sugar (kk, T) =
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   238
          {T = T, fp = fp, index = kk, pre_bnfs = pre_bnfs, nested_bnfs = nested_bnfs,
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   239
           nesting_bnfs = nesting_bnfs, fp_res = fp_res, ctr_defss = ctr_defss,
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   240
           ctr_sugars = ctr_sugars, co_iterss = co_iterss, mapss = mapss, co_inducts = co_inducts,
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   241
           co_iter_thmsss = transpose [un_fold_thmss, co_rec_thmss],
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   242
           disc_co_itersss = transpose [disc_unfold_thmss, disc_corec_thmss],
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   243
           sel_co_iterssss = transpose [sel_unfold_thmsss, sel_corec_thmsss]}
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   244
          |> morph_fp_sugar phi;
54256
4843082be7ef added some N2M caching
blanchet
parents: 54255
diff changeset
   245
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   246
        val n2m_sugar = (map_index mk_target_fp_sugar fpTs, fp_sugar_thms);
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   247
      in
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   248
        (n2m_sugar, lthy |> register_n2m_sugar key n2m_sugar)
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   249
      end)
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   250
  end;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   251
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   252
fun indexify_callsss fp_sugar callsss =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   253
  let
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   254
    val {ctrs, ...} = of_fp_sugar #ctr_sugars fp_sugar;
54283
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   255
    fun indexify_ctr ctr =
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   256
      (case AList.lookup Term.aconv_untyped callsss ctr of
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   257
        NONE => replicate (num_binder_types (fastype_of ctr)) []
54243
a596292be9a8 more robust n2m w.r.t. 'let's
blanchet
parents: 54240
diff changeset
   258
      | SOME callss => map (map (Envir.beta_eta_contract o unfold_let)) callss);
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   259
  in
54283
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   260
    map indexify_ctr ctrs
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   261
  end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   262
54283
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   263
fun retypargs tyargs (Type (s, _)) = Type (s, tyargs);
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   264
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   265
fun fold_subtype_pairs f (T as Type (s, Ts), U as Type (s', Us)) =
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   266
    f (T, U) #> (if s = s' then fold (fold_subtype_pairs f) (Ts ~~ Us) else I)
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   267
  | fold_subtype_pairs f TU = f TU;
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   268
54009
f138452e8265 got rid of dead feature
blanchet
parents: 54006
diff changeset
   269
fun nested_to_mutual_fps fp actual_bs actual_Ts get_indices actual_callssss0 lthy =
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   270
  let
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   271
    val qsoty = quote o Syntax.string_of_typ lthy;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   272
    val qsotys = space_implode " or " o map qsoty;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   273
54253
04cd231e2b9e nicer error message in case of duplicates
blanchet
parents: 54245
diff changeset
   274
    fun duplicate_datatype T = error (qsoty T ^ " is not mutually recursive with itself");
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   275
    fun not_co_datatype0 T = error (qsoty T ^ " is not a " ^ co_prefix fp ^ "datatype");
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   276
    fun not_co_datatype (T as Type (s, _)) =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   277
        if fp = Least_FP andalso
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   278
           is_some (Datatype_Data.get_info (Proof_Context.theory_of lthy) s) then
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   279
          error (qsoty T ^ " is not a new-style datatype (cf. \"datatype_new\")")
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   280
        else
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   281
          not_co_datatype0 T
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   282
      | not_co_datatype T = not_co_datatype0 T;
54303
4f55054d197c reintroduced check from bd36da55d825 after all, due to hard-to-solve issues outside the N2M code
blanchet
parents: 54302
diff changeset
   283
    fun not_mutually_nested_rec Ts1 Ts2 =
4f55054d197c reintroduced check from bd36da55d825 after all, due to hard-to-solve issues outside the N2M code
blanchet
parents: 54302
diff changeset
   284
      error (qsotys Ts1 ^ " is neither mutually recursive with " ^ qsotys Ts2 ^
4f55054d197c reintroduced check from bd36da55d825 after all, due to hard-to-solve issues outside the N2M code
blanchet
parents: 54302
diff changeset
   285
        " nor nested recursive via " ^ qsotys Ts2);
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   286
54253
04cd231e2b9e nicer error message in case of duplicates
blanchet
parents: 54245
diff changeset
   287
    val _ = (case Library.duplicates (op =) actual_Ts of [] => () | T :: _ => duplicate_datatype T);
04cd231e2b9e nicer error message in case of duplicates
blanchet
parents: 54245
diff changeset
   288
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   289
    val perm_actual_Ts =
54266
4e0738356ac9 stronger normalization, to increase n2m cache effectiveness
blanchet
parents: 54265
diff changeset
   290
      sort (prod_ord int_ord Term_Ord.typ_ord o pairself (`Term.size_of_typ)) actual_Ts;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   291
54283
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   292
    fun the_ctrs_of (Type (s, Ts)) = map (mk_ctr Ts) (#ctrs (the (ctr_sugar_of lthy s)));
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   293
54282
blanchet
parents: 54280
diff changeset
   294
    fun the_fp_sugar_of (T as Type (T_name, _)) =
blanchet
parents: 54280
diff changeset
   295
      (case fp_sugar_of lthy T_name of
blanchet
parents: 54280
diff changeset
   296
        SOME (fp_sugar as {fp = fp', ...}) => if fp = fp' then fp_sugar else not_co_datatype T
blanchet
parents: 54280
diff changeset
   297
      | NONE => not_co_datatype T);
blanchet
parents: 54280
diff changeset
   298
54283
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   299
    fun gen_rhss_in gen_Ts rho subTs =
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   300
      let
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   301
        fun maybe_insert (T, Type (_, gen_tyargs)) =
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   302
            if member (op =) subTs T then insert (op =) gen_tyargs else I
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   303
          | maybe_insert _ = I;
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   304
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   305
        val ctrs = maps the_ctrs_of gen_Ts;
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   306
        val gen_ctr_Ts = maps (binder_types o fastype_of) ctrs;
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   307
        val ctr_Ts = map (Term.typ_subst_atomic rho) gen_ctr_Ts;
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   308
      in
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   309
        fold (fold_subtype_pairs maybe_insert) (ctr_Ts ~~ gen_ctr_Ts) []
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   310
      end;
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   311
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   312
    fun gather_types _ _ num_groups seen gen_seen [] = (num_groups, seen, gen_seen)
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   313
      | gather_types lthy rho num_groups seen gen_seen ((T as Type (_, tyargs)) :: Ts) =
54282
blanchet
parents: 54280
diff changeset
   314
        let
54283
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   315
          val {fp_res = {Ts = mutual_Ts0, ...}, ...} = the_fp_sugar_of T;
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   316
          val mutual_Ts = map (retypargs tyargs) mutual_Ts0;
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   317
54303
4f55054d197c reintroduced check from bd36da55d825 after all, due to hard-to-solve issues outside the N2M code
blanchet
parents: 54302
diff changeset
   318
          val _ = seen = [] orelse exists (exists_subtype_in seen) mutual_Ts orelse
4f55054d197c reintroduced check from bd36da55d825 after all, due to hard-to-solve issues outside the N2M code
blanchet
parents: 54302
diff changeset
   319
            not_mutually_nested_rec mutual_Ts seen;
4f55054d197c reintroduced check from bd36da55d825 after all, due to hard-to-solve issues outside the N2M code
blanchet
parents: 54302
diff changeset
   320
54283
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   321
          fun fresh_tyargs () =
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   322
            let
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   323
              (* The name "'z" is unlikely to clash with the context, yielding more cache hits. *)
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   324
              val (gen_tyargs, lthy') =
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   325
                variant_tfrees (replicate (length tyargs) "z") lthy
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   326
                |>> map Logic.varifyT_global;
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   327
              val rho' = (gen_tyargs ~~ tyargs) @ rho;
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   328
            in
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   329
              (rho', gen_tyargs, gen_seen, lthy')
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   330
            end;
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   331
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   332
          val (rho', gen_tyargs, gen_seen', lthy') =
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   333
            if exists (exists_subtype_in seen) mutual_Ts then
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   334
              (case gen_rhss_in gen_seen rho mutual_Ts of
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   335
                [] => fresh_tyargs ()
54899
7a01387c47d5 added tactic to prove 'disc_iff' properties in 'primcorec'
blanchet
parents: 54740
diff changeset
   336
              | gen_tyargs :: gen_tyargss_tl =>
54283
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   337
                let
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   338
                  val unify_pairs = split_list (maps (curry (op ~~) gen_tyargs) gen_tyargss_tl);
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   339
                  val mgu = Type.raw_unifys unify_pairs Vartab.empty;
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   340
                  val gen_tyargs' = map (Envir.subst_type mgu) gen_tyargs;
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   341
                  val gen_seen' = map (Envir.subst_type mgu) gen_seen;
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   342
                in
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   343
                  (rho, gen_tyargs', gen_seen', lthy)
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   344
                end)
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   345
            else
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   346
              fresh_tyargs ();
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   347
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   348
          val gen_mutual_Ts = map (retypargs gen_tyargs) mutual_Ts0;
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   349
          val Ts' = filter_out (member (op =) mutual_Ts) Ts;
54282
blanchet
parents: 54280
diff changeset
   350
        in
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   351
          gather_types lthy' rho' (num_groups + 1) (seen @ mutual_Ts) (gen_seen' @ gen_mutual_Ts)
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   352
            Ts'
54282
blanchet
parents: 54280
diff changeset
   353
        end
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   354
      | gather_types _ _ _ _ _ (T :: _) = not_co_datatype T;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   355
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   356
    val (num_groups, perm_Ts, perm_gen_Ts) = gather_types lthy [] 0 [] [] perm_actual_Ts;
54283
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   357
    val perm_frozen_gen_Ts = map Logic.unvarifyT_global perm_gen_Ts;
6f0a49ed1bb1 fourth attempt at generalizing N2M types (to leverage caching)
blanchet
parents: 54282
diff changeset
   358
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   359
    val missing_Ts = perm_Ts |> subtract (op =) actual_Ts;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   360
    val Ts = actual_Ts @ missing_Ts;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   361
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   362
    val nn = length Ts;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   363
    val kks = 0 upto nn - 1;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   364
54267
78e8a178b690 use right permutation in 'map2'
blanchet
parents: 54266
diff changeset
   365
    val callssss0 = pad_list [] nn actual_callssss0;
78e8a178b690 use right permutation in 'map2'
blanchet
parents: 54266
diff changeset
   366
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   367
    val common_name = mk_common_name (map Binding.name_of actual_bs);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   368
    val bs = pad_list (Binding.name common_name) nn actual_bs;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   369
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   370
    fun permute xs = permute_like (op =) Ts perm_Ts xs;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   371
    fun unpermute perm_xs = permute_like (op =) perm_Ts Ts perm_xs;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   372
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   373
    val perm_bs = permute bs;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   374
    val perm_kks = permute kks;
54267
78e8a178b690 use right permutation in 'map2'
blanchet
parents: 54266
diff changeset
   375
    val perm_callssss0 = permute callssss0;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   376
    val perm_fp_sugars0 = map (the o fp_sugar_of lthy o fst o dest_Type) perm_Ts;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   377
54267
78e8a178b690 use right permutation in 'map2'
blanchet
parents: 54266
diff changeset
   378
    val perm_callssss = map2 indexify_callsss perm_fp_sugars0 perm_callssss0;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   379
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   380
    val get_perm_indices = map (fn kk => find_index (curry (op =) kk) perm_kks) o get_indices;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   381
53746
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53678
diff changeset
   382
    val ((perm_fp_sugars, fp_sugar_thms), lthy) =
54286
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   383
      if num_groups > 1 then
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   384
        mutualize_fp_sugars fp perm_bs perm_frozen_gen_Ts get_perm_indices perm_callssss
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   385
          perm_fp_sugars0 lthy
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   386
      else
22616f65d4ea properly detect when to perform n2m -- e.g. handle the case of two independent functions on irrelevant types being defined in parallel
blanchet
parents: 54284
diff changeset
   387
        ((perm_fp_sugars0, (NONE, NONE)), lthy);
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   388
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   389
    val fp_sugars = unpermute perm_fp_sugars;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   390
  in
53746
bd038e48526d have "datatype_new_compat" register induction and recursion theorems in nested case
blanchet
parents: 53678
diff changeset
   391
    ((missing_Ts, perm_kks, fp_sugars, fp_sugar_thms), lthy)
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   392
  end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   393
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   394
end;