| author | wenzelm | 
| Sun, 19 Mar 2017 14:43:54 +0100 | |
| changeset 65326 | cb7cb57c7ce1 | 
| parent 64637 | a15785625f7c | 
| child 69593 | 3dda49e08b9d | 
| permissions | -rw-r--r-- | 
| 55061 | 1 | (* Title: HOL/Tools/BNF/bnf_fp_n2m_sugar.ML | 
| 53303 
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 | 
| 55343 | 10 | val unfold_lets_splits: term -> term | 
| 57895 | 11 | val unfold_splits_lets: term -> term | 
| 54243 | 12 | val dest_map: Proof.context -> string -> term -> term * term list | 
| 62326 
3cf7a067599c
allow predicator instead of map function in 'primrec'
 blanchet parents: 
62321diff
changeset | 13 | val dest_pred: Proof.context -> string -> term -> term * term list | 
| 54243 | 14 | |
| 58335 
a5a3b576fcfb
generate 'code' attribute only if 'code' plugin is enabled
 blanchet parents: 
58315diff
changeset | 15 | val mutualize_fp_sugars: (string -> bool) -> BNF_Util.fp_kind -> int list -> binding list -> | 
| 
a5a3b576fcfb
generate 'code' attribute only if 'code' plugin is enabled
 blanchet parents: 
58315diff
changeset | 16 | typ list -> term list -> term list list list list -> BNF_FP_Def_Sugar.fp_sugar list -> | 
| 
a5a3b576fcfb
generate 'code' attribute only if 'code' plugin is enabled
 blanchet parents: 
58315diff
changeset | 17 | local_theory -> | 
| 58180 | 18 | (BNF_FP_Def_Sugar.fp_sugar list | 
| 53746 
bd038e48526d
have "datatype_new_compat" register induction and recursion theorems in nested case
 blanchet parents: 
53678diff
changeset | 19 | * (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: 
53678diff
changeset | 20 | * local_theory | 
| 58335 
a5a3b576fcfb
generate 'code' attribute only if 'code' plugin is enabled
 blanchet parents: 
58315diff
changeset | 21 | val nested_to_mutual_fps: (string -> bool) -> BNF_Util.fp_kind -> binding list -> typ list -> | 
| 
a5a3b576fcfb
generate 'code' attribute only if 'code' plugin is enabled
 blanchet parents: 
58315diff
changeset | 22 | term list -> (term * term list list) list list -> local_theory -> | 
| 58180 | 23 | (typ list * int list * BNF_FP_Def_Sugar.fp_sugar list | 
| 53746 
bd038e48526d
have "datatype_new_compat" register induction and recursion theorems in nested case
 blanchet parents: 
53678diff
changeset | 24 | * (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: 
53678diff
changeset | 25 | * local_theory | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 26 | end; | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 27 | |
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 28 | 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 | 29 | struct | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 30 | |
| 54006 | 31 | open Ctr_Sugar | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 32 | open BNF_Util | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 33 | open BNF_Def | 
| 62720 | 34 | open BNF_Comp | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 35 | open BNF_FP_Util | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 36 | 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 | 37 | open BNF_FP_N2M | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 38 | |
| 59662 | 39 | val n2mN = "n2m_"; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 40 | |
| 54256 | 41 | type n2m_sugar = fp_sugar list * (lfp_sugar_thms option * gfp_sugar_thms option); | 
| 42 | ||
| 43 | structure Data = Generic_Data | |
| 44 | ( | |
| 45 | type T = n2m_sugar Typtab.table; | |
| 46 | val empty = Typtab.empty; | |
| 47 | val extend = I; | |
| 55394 
d5ebe055b599
more liberal merging of BNFs and constructor sugar
 blanchet parents: 
55343diff
changeset | 48 | fun merge data : T = Typtab.merge (K true) data; | 
| 54256 | 49 | ); | 
| 50 | ||
| 51 | fun morph_n2m_sugar phi (fp_sugars, (lfp_sugar_thms_opt, gfp_sugar_thms_opt)) = | |
| 52 | (map (morph_fp_sugar phi) fp_sugars, | |
| 53 | (Option.map (morph_lfp_sugar_thms phi) lfp_sugar_thms_opt, | |
| 54 | Option.map (morph_gfp_sugar_thms phi) gfp_sugar_thms_opt)); | |
| 55 | ||
| 58115 | 56 | val transfer_n2m_sugar = morph_n2m_sugar o Morphism.transfer_morphism; | 
| 54256 | 57 | |
| 58 | fun n2m_sugar_of ctxt = | |
| 59 | Typtab.lookup (Data.get (Context.Proof ctxt)) | |
| 58115 | 60 | #> Option.map (transfer_n2m_sugar (Proof_Context.theory_of ctxt)); | 
| 54256 | 61 | |
| 62 | fun register_n2m_sugar key n2m_sugar = | |
| 63 |   Local_Theory.declaration {syntax = false, pervasive = false}
 | |
| 55444 
ec73f81e49e7
iteration n in the 'default' vs. 'update_new' vs. 'update' saga -- 'update' makes sense now that we honor the canonical order on 'merge' (as opposed to raising 'DUP')
 blanchet parents: 
55414diff
changeset | 64 | (fn phi => Data.map (Typtab.update (key, morph_n2m_sugar phi n2m_sugar))); | 
| 54256 | 65 | |
| 57895 | 66 | fun unfold_lets_splits (Const (@{const_name Let}, _) $ t $ u) =
 | 
| 57897 | 67 | unfold_lets_splits (betapply (unfold_splits_lets u, t)) | 
| 57895 | 68 | | unfold_lets_splits (t $ u) = betapply (unfold_lets_splits t, unfold_lets_splits u) | 
| 69 | | unfold_lets_splits (Abs (s, T, t)) = Abs (s, T, unfold_lets_splits t) | |
| 70 | | unfold_lets_splits t = t | |
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61334diff
changeset | 71 | and unfold_splits_lets ((t as Const (@{const_name case_prod}, _)) $ u) =
 | 
| 57895 | 72 | (case unfold_splits_lets u of | 
| 55336 | 73 | u' as Abs (s1, T1, Abs (s2, T2, _)) => | 
| 74 | let val v = Var ((s1 ^ s2, Term.maxidx_of_term u' + 1), HOLogic.mk_prodT (T1, T2)) in | |
| 75 | lambda v (incr_boundvars 1 (betapplys (u', [HOLogic.mk_fst v, HOLogic.mk_snd v]))) | |
| 54243 | 76 | end | 
| 57897 | 77 | | _ => t $ unfold_lets_splits u) | 
| 78 |   | unfold_splits_lets (t as Const (@{const_name Let}, _) $ _ $ _) = unfold_lets_splits t
 | |
| 79 | | unfold_splits_lets (t $ u) = betapply (unfold_splits_lets t, unfold_lets_splits u) | |
| 57895 | 80 | | unfold_splits_lets (Abs (s, T, t)) = Abs (s, T, unfold_splits_lets t) | 
| 81 | | unfold_splits_lets t = unfold_lets_splits t; | |
| 54243 | 82 | |
| 62326 
3cf7a067599c
allow predicator instead of map function in 'primrec'
 blanchet parents: 
62321diff
changeset | 83 | fun dest_either_map_or_pred map_or_pred_of_bnf ctxt T_name call = | 
| 54243 | 84 | let | 
| 62326 
3cf7a067599c
allow predicator instead of map function in 'primrec'
 blanchet parents: 
62321diff
changeset | 85 | val bnf = the (bnf_of ctxt T_name); | 
| 
3cf7a067599c
allow predicator instead of map function in 'primrec'
 blanchet parents: 
62321diff
changeset | 86 | val const0 = map_or_pred_of_bnf bnf; | 
| 54243 | 87 | val live = live_of_bnf bnf; | 
| 62326 
3cf7a067599c
allow predicator instead of map function in 'primrec'
 blanchet parents: 
62321diff
changeset | 88 | val (f_Ts, _) = strip_typeN live (fastype_of const0); | 
| 
3cf7a067599c
allow predicator instead of map function in 'primrec'
 blanchet parents: 
62321diff
changeset | 89 |     val fs = map_index (fn (i, T) => Var (("f", i), T)) f_Ts;
 | 
| 
3cf7a067599c
allow predicator instead of map function in 'primrec'
 blanchet parents: 
62321diff
changeset | 90 | val pat = betapplys (const0, fs); | 
| 
3cf7a067599c
allow predicator instead of map function in 'primrec'
 blanchet parents: 
62321diff
changeset | 91 | val (_, tenv) = fo_match ctxt call pat; | 
| 54243 | 92 | in | 
| 62326 
3cf7a067599c
allow predicator instead of map function in 'primrec'
 blanchet parents: 
62321diff
changeset | 93 | (const0, Vartab.fold_rev (fn (_, (_, f)) => cons f) tenv []) | 
| 54243 | 94 | end; | 
| 95 | ||
| 62326 
3cf7a067599c
allow predicator instead of map function in 'primrec'
 blanchet parents: 
62321diff
changeset | 96 | val dest_map = dest_either_map_or_pred map_of_bnf; | 
| 
3cf7a067599c
allow predicator instead of map function in 'primrec'
 blanchet parents: 
62321diff
changeset | 97 | val dest_pred = dest_either_map_or_pred pred_of_bnf; | 
| 54243 | 98 | |
| 62326 
3cf7a067599c
allow predicator instead of map function in 'primrec'
 blanchet parents: 
62321diff
changeset | 99 | fun dest_map_or_pred ctxt T_name call = | 
| 
3cf7a067599c
allow predicator instead of map function in 'primrec'
 blanchet parents: 
62321diff
changeset | 100 | (case try (dest_map ctxt T_name) call of | 
| 
3cf7a067599c
allow predicator instead of map function in 'primrec'
 blanchet parents: 
62321diff
changeset | 101 | SOME res => res | 
| 
3cf7a067599c
allow predicator instead of map function in 'primrec'
 blanchet parents: 
62321diff
changeset | 102 | | NONE => dest_pred ctxt T_name call); | 
| 
3cf7a067599c
allow predicator instead of map function in 'primrec'
 blanchet parents: 
62321diff
changeset | 103 | |
| 
3cf7a067599c
allow predicator instead of map function in 'primrec'
 blanchet parents: 
62321diff
changeset | 104 | fun dest_abs_or_applied_map_or_pred _ _ (Abs (_, _, t)) = (Term.dummy, [t]) | 
| 
3cf7a067599c
allow predicator instead of map function in 'primrec'
 blanchet parents: 
62321diff
changeset | 105 | | dest_abs_or_applied_map_or_pred ctxt s (t1 $ _) = dest_map_or_pred ctxt s t1; | 
| 54239 | 106 | |
| 54245 | 107 | fun map_partition f xs = | 
| 108 | fold_rev (fn x => fn (ys, (good, bad)) => | |
| 109 | case f x of SOME y => (y :: ys, (x :: good, bad)) | NONE => (ys, (good, x :: bad))) | |
| 110 | xs ([], ([], [])); | |
| 111 | ||
| 63798 | 112 | fun key_of_fp_eqs fp As fpTs Xs ctrXs_repTs = | 
| 113 | Type (case_fp fp "l" "g", fpTs @ Xs @ ctrXs_repTs) | |
| 58222 | 114 | |> Term.map_atyps (fn T as TFree (_, S) => | 
| 115 | (case find_index (curry (op =) T) As of | |
| 116 | ~1 => T | |
| 117 |     | j => TFree ("'" ^ string_of_int j, S)));
 | |
| 54256 | 118 | |
| 55772 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 119 | fun get_indices callers t = | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 120 | callers | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 121 | |> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE) | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 122 | |> map_filter I; | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 123 | |
| 58340 | 124 | fun mutualize_fp_sugars plugins fp mutual_cliques bs fpTs callers callssss fp_sugars0 no_defs_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: 
54284diff
changeset | 125 | let | 
| 58280 
2ec3e2de34c3
more canonical (and correct) local theory threading
 blanchet parents: 
58236diff
changeset | 126 | val thy = Proof_Context.theory_of no_defs_lthy; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 127 | |
| 58280 
2ec3e2de34c3
more canonical (and correct) local theory threading
 blanchet parents: 
58236diff
changeset | 128 | val qsotm = quote o Syntax.string_of_term no_defs_lthy; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 129 | |
| 55479 
ece4910c3ea0
improved 'datatype_new_compat': generate more fixpoint equations for types like 'datatype_new x = C (x list) (x list)' (here, one equation for each x list instead of a single for both), for higher compatibility + code generation attributes on the recursor
 blanchet parents: 
55444diff
changeset | 130 | fun incompatible_calls ts = | 
| 
ece4910c3ea0
improved 'datatype_new_compat': generate more fixpoint equations for types like 'datatype_new x = C (x list) (x list)' (here, one equation for each x list instead of a single for both), for higher compatibility + code generation attributes on the recursor
 blanchet parents: 
55444diff
changeset | 131 |       error ("Incompatible " ^ co_prefix fp ^ "recursive calls: " ^ commas (map qsotm ts));
 | 
| 55772 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 132 | fun mutual_self_call caller t = | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 133 |       error ("Unsupported mutual self-call " ^ qsotm t ^ " from " ^ qsotm caller);
 | 
| 54301 
e0943a2ee400
added check to avoid odd situations the N2M code cannot handle
 blanchet parents: 
54300diff
changeset | 134 | fun nested_self_call t = | 
| 
e0943a2ee400
added check to avoid odd situations the N2M code cannot handle
 blanchet parents: 
54300diff
changeset | 135 |       error ("Unsupported nested self-call " ^ qsotm t);
 | 
| 54245 | 136 | |
| 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: 
54284diff
changeset | 137 | 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: 
54284diff
changeset | 138 | 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 | 139 | |
| 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: 
54284diff
changeset | 140 | val nn = length fpTs; | 
| 55539 
0819931d652d
simplified data structure by reducing the incidence of clumsy indices
 blanchet parents: 
55486diff
changeset | 141 | val kks = 0 upto nn - 1; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 142 | |
| 58460 | 143 |     fun target_ctr_sugar_of_fp_sugar fpT ({T, fp_ctr_sugar = {ctr_sugar, ...}, ...} : fp_sugar) =
 | 
| 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: 
54284diff
changeset | 144 | 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: 
54284diff
changeset | 145 | val rho = Vartab.fold (cons o apsnd snd) (Sign.typ_match thy (T, fpT) Vartab.empty) []; | 
| 54740 | 146 | 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: 
54284diff
changeset | 147 | in | 
| 55539 
0819931d652d
simplified data structure by reducing the incidence of clumsy indices
 blanchet parents: 
55486diff
changeset | 148 | morph_ctr_sugar phi ctr_sugar | 
| 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: 
54284diff
changeset | 149 | end; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 150 | |
| 62321 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 blanchet parents: 
61551diff
changeset | 151 | val ctor_iff_dtors = map (#ctor_iff_dtor o #fp_ctr_sugar) fp_sugars0; | 
| 58460 | 152 | val ctr_defss = map (#ctr_defs o #fp_ctr_sugar) fp_sugars0; | 
| 58462 | 153 | val mapss = map (#map_thms o #fp_bnf_sugar) fp_sugars0; | 
| 54302 | 154 | 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 | 155 | |
| 54302 | 156 | 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: 
54284diff
changeset | 157 | 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 | 158 | |
| 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: 
54284diff
changeset | 159 | 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: 
54284diff
changeset | 160 | 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 | 161 | |
| 61334 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 traytel parents: 
61125diff
changeset | 162 | val ((Cs, Xs), _) = | 
| 58280 
2ec3e2de34c3
more canonical (and correct) local theory threading
 blanchet parents: 
58236diff
changeset | 163 | no_defs_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: 
54284diff
changeset | 164 | |> 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: 
54284diff
changeset | 165 | |> 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: 
54284diff
changeset | 166 | ||>> 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 | 167 | |
| 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: 
54284diff
changeset | 168 | fun check_call_dead live_call call = | 
| 55772 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 169 | if null (get_indices callers call) then () else incompatible_calls [live_call, call]; | 
| 54245 | 170 | |
| 55480 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55479diff
changeset | 171 | fun freeze_fpTs_type_based_default (T as Type (s, Ts)) = | 
| 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55479diff
changeset | 172 | (case filter (curry (op =) T o snd) (map_index I fpTs) of | 
| 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55479diff
changeset | 173 | [(kk, _)] => nth Xs kk | 
| 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55479diff
changeset | 174 | | _ => Type (s, map freeze_fpTs_type_based_default Ts)) | 
| 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55479diff
changeset | 175 | | freeze_fpTs_type_based_default T = T; | 
| 55479 
ece4910c3ea0
improved 'datatype_new_compat': generate more fixpoint equations for types like 'datatype_new x = C (x list) (x list)' (here, one equation for each x list instead of a single for both), for higher compatibility + code generation attributes on the recursor
 blanchet parents: 
55444diff
changeset | 176 | |
| 55772 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 177 | fun freeze_fpTs_mutual_call kk fpT calls T = | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 178 | (case fold (union (op =)) (map (get_indices callers) calls) [] of | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 179 | [] => if T = fpT then nth Xs kk else freeze_fpTs_type_based_default T | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 180 | | [kk'] => | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 181 | if T = fpT andalso kk' <> kk then | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 182 | mutual_self_call (nth callers kk) | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 183 | (the (find_first (not o null o get_indices callers) calls)) | 
| 58290 
14e186d2dd3a
proper checks -- the calls data structure may contain spurious entries
 blanchet parents: 
58280diff
changeset | 184 | else if T = nth fpTs kk' then | 
| 
14e186d2dd3a
proper checks -- the calls data structure may contain spurious entries
 blanchet parents: 
58280diff
changeset | 185 | nth Xs kk' | 
| 55772 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 186 | else | 
| 58290 
14e186d2dd3a
proper checks -- the calls data structure may contain spurious entries
 blanchet parents: 
58280diff
changeset | 187 | freeze_fpTs_type_based_default T | 
| 55480 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55479diff
changeset | 188 | | _ => incompatible_calls calls); | 
| 54253 | 189 | |
| 55772 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 190 | fun freeze_fpTs_map kk (fpT as Type (_, Ts')) (callss, (live_call :: _, dead_calls)) | 
| 55480 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55479diff
changeset | 191 | (Type (s, Ts)) = | 
| 54301 
e0943a2ee400
added check to avoid odd situations the N2M code cannot handle
 blanchet parents: 
54300diff
changeset | 192 | if Ts' = Ts then | 
| 
e0943a2ee400
added check to avoid odd situations the N2M code cannot handle
 blanchet parents: 
54300diff
changeset | 193 | nested_self_call live_call | 
| 
e0943a2ee400
added check to avoid odd situations the N2M code cannot handle
 blanchet parents: 
54300diff
changeset | 194 | else | 
| 
e0943a2ee400
added check to avoid odd situations the N2M code cannot handle
 blanchet parents: 
54300diff
changeset | 195 | (List.app (check_call_dead live_call) dead_calls; | 
| 55772 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 196 | Type (s, map2 (freeze_fpTs_call kk fpT) | 
| 55480 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55479diff
changeset | 197 | (flatten_type_args_of_bnf (the (bnf_of no_defs_lthy s)) [] (transpose callss)) Ts)) | 
| 55772 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 198 | and freeze_fpTs_call kk 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: 
54284diff
changeset | 199 | (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: 
54284diff
changeset | 200 | ([], _) => | 
| 62326 
3cf7a067599c
allow predicator instead of map function in 'primrec'
 blanchet parents: 
62321diff
changeset | 201 | (case map_partition (try (snd o dest_abs_or_applied_map_or_pred no_defs_lthy s)) calls of | 
| 55772 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 202 | ([], _) => freeze_fpTs_mutual_call kk fpT calls T | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 203 | | callsp => freeze_fpTs_map kk fpT callsp T) | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 204 | | callsp => freeze_fpTs_map kk fpT callsp T) | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 205 | | freeze_fpTs_call _ _ _ T = T; | 
| 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: 
54284diff
changeset | 207 | val ctr_Tsss = map (map binder_types) ctr_Tss; | 
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58577diff
changeset | 208 |     val ctrXs_Tsss = @{map 4} (map2 o map2 oo freeze_fpTs_call) kks fpTs callssss ctr_Tsss;
 | 
| 55966 | 209 | val ctrXs_repTs = map mk_sumprodT_balanced ctrXs_Tsss; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 210 | |
| 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: 
54284diff
changeset | 211 | 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: 
54284diff
changeset | 212 | 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: 
54284diff
changeset | 213 | 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 | 214 | |
| 63798 | 215 | val key = key_of_fp_eqs fp As fpTs Xs ctrXs_repTs; | 
| 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: 
54284diff
changeset | 216 | 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: 
54284diff
changeset | 217 | (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: 
54284diff
changeset | 218 | 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: 
54284diff
changeset | 219 | | 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: 
54284diff
changeset | 220 | 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: 
54284diff
changeset | 221 | 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: 
54284diff
changeset | 222 | 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: 
54284diff
changeset | 223 | 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: 
54284diff
changeset | 224 | 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 | 225 | |
| 55702 
63c80031d8dd
improved accounting for dead variables when naming set functions (refines d71c2737ee21)
 blanchet parents: 
55701diff
changeset | 226 | val Type (s, Us) :: _ = fpTs; | 
| 
63c80031d8dd
improved accounting for dead variables when naming set functions (refines d71c2737ee21)
 blanchet parents: 
55701diff
changeset | 227 | val killed_As' = | 
| 
63c80031d8dd
improved accounting for dead variables when naming set functions (refines d71c2737ee21)
 blanchet parents: 
55701diff
changeset | 228 | (case bnf_of no_defs_lthy s of | 
| 63551 | 229 | NONE => As' | 
| 55702 
63c80031d8dd
improved accounting for dead variables when naming set functions (refines d71c2737ee21)
 blanchet parents: 
55701diff
changeset | 230 | | SOME bnf => | 
| 
63c80031d8dd
improved accounting for dead variables when naming set functions (refines d71c2737ee21)
 blanchet parents: 
55701diff
changeset | 231 | let | 
| 
63c80031d8dd
improved accounting for dead variables when naming set functions (refines d71c2737ee21)
 blanchet parents: 
55701diff
changeset | 232 | val Type (_, Ts) = T_of_bnf bnf; | 
| 
63c80031d8dd
improved accounting for dead variables when naming set functions (refines d71c2737ee21)
 blanchet parents: 
55701diff
changeset | 233 | val deads = deads_of_bnf bnf; | 
| 
63c80031d8dd
improved accounting for dead variables when naming set functions (refines d71c2737ee21)
 blanchet parents: 
55701diff
changeset | 234 | val dead_Us = | 
| 
63c80031d8dd
improved accounting for dead variables when naming set functions (refines d71c2737ee21)
 blanchet parents: 
55701diff
changeset | 235 | map_filter (fn (T, U) => if member (op =) deads T then SOME U else NONE) (Ts ~~ Us); | 
| 
63c80031d8dd
improved accounting for dead variables when naming set functions (refines d71c2737ee21)
 blanchet parents: 
55701diff
changeset | 236 | in fold Term.add_tfreesT dead_Us [] end); | 
| 
63c80031d8dd
improved accounting for dead variables when naming set functions (refines d71c2737ee21)
 blanchet parents: 
55701diff
changeset | 237 | |
| 55803 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 traytel parents: 
55772diff
changeset | 238 | val fp_absT_infos = map #absT_info fp_sugars0; | 
| 62719 | 239 | val indexed_fp_ress = map (apsnd #fp_res o `(#fp_res_index)) fp_sugars0; | 
| 62684 
cb20e8828196
document that n2m does not depend on most things in fp_sugar in its type
 traytel parents: 
62335diff
changeset | 240 | |
| 62720 | 241 |         val (((pre_bnfs, absT_infos), _), (fp_res as {xtor_co_recs = xtor_co_recs0, xtor_co_induct,
 | 
| 58211 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 242 | dtor_injects, dtor_ctors, xtor_co_rec_thms, ...}, lthy)) = | 
| 63802 | 243 | fixpoint_bnf false I (construct_mutualized_fp fp mutual_cliques fpTs indexed_fp_ress) | 
| 63798 | 244 | fp_bs As' killed_As' (map dest_TFree Xs) ctrXs_repTs empty_comp_cache no_defs_lthy; | 
| 55803 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 traytel parents: 
55772diff
changeset | 245 | |
| 62685 | 246 | val time = time lthy; | 
| 247 | val timer = time (Timer.startRealTimer ()); | |
| 248 | ||
| 58211 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 249 | val fp_abs_inverses = map #abs_inverse fp_absT_infos; | 
| 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 250 | val fp_type_definitions = map #type_definition fp_absT_infos; | 
| 55803 
74d3fe9031d8
joint work with blanchet: intermediate typedef for the input to fp-operations
 traytel parents: 
55772diff
changeset | 251 | |
| 58211 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 252 | val abss = map #abs absT_infos; | 
| 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 253 | val reps = map #rep absT_infos; | 
| 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 254 | val absTs = map #absT absT_infos; | 
| 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 255 | val repTs = map #repT absT_infos; | 
| 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 256 | val abs_inverses = map #abs_inverse absT_infos; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 257 | |
| 58211 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 258 | val fp_nesting_bnfs = nesting_bnfs lthy ctrXs_Tsss Xs; | 
| 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 259 | val live_nesting_bnfs = nesting_bnfs lthy ctrXs_Tsss As; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 260 | |
| 61551 | 261 | val (xtor_co_recs, recs_args_types, corecs_args_types) = | 
| 262 | mk_co_recs_prelims lthy fp ctr_Tsss fpTs Cs absTs repTs ns mss xtor_co_recs0; | |
| 54256 | 263 | |
| 58211 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 264 | fun mk_binding b pre = Binding.prefix_name (pre ^ "_") b; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 265 | |
| 58211 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 266 | val ((co_recs, co_rec_defs), lthy) = | 
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58577diff
changeset | 267 |           @{fold_map 2} (fn b =>
 | 
| 58211 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 268 | if fp = Least_FP then define_rec (the recs_args_types) (mk_binding b) fpTs Cs reps | 
| 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 269 | else define_corec (the corecs_args_types) (mk_binding b) fpTs Cs abss) | 
| 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 270 | fp_bs xtor_co_recs lthy | 
| 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 271 | |>> split_list; | 
| 58131 
1abeda3c3bc2
drop hopeless feature -- unfolding of BNF datatype info without a prior 'datatype_compat'
 blanchet parents: 
58117diff
changeset | 272 | |
| 62685 | 273 |         val timer = time (timer ("High-level " ^ co_prefix fp ^ "recursors"));
 | 
| 274 | ||
| 58211 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 275 | val ((common_co_inducts, co_inductss', co_rec_thmss, co_rec_disc_thmss, co_rec_sel_thmsss), | 
| 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 276 | fp_sugar_thms) = | 
| 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 277 | if fp = Least_FP then | 
| 58335 
a5a3b576fcfb
generate 'code' attribute only if 'code' plugin is enabled
 blanchet parents: 
58315diff
changeset | 278 | derive_induct_recs_thms_for_types plugins pre_bnfs recs_args_types xtor_co_induct | 
| 58211 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 279 | xtor_co_rec_thms live_nesting_bnfs fp_nesting_bnfs fpTs Cs Xs ctrXs_Tsss | 
| 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 280 | fp_abs_inverses fp_type_definitions abs_inverses ctrss ctr_defss co_recs co_rec_defs | 
| 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 281 | lthy | 
| 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 282 | |> `(fn ((inducts, induct, _), (rec_thmss, _)) => | 
| 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 283 | ([induct], [inducts], rec_thmss, replicate nn [], replicate nn [])) | 
| 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 284 | ||> (fn info => (SOME info, NONE)) | 
| 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 285 | else | 
| 64637 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 blanchet parents: 
63859diff
changeset | 286 | derive_coinduct_corecs_thms_for_types lthy pre_bnfs (the corecs_args_types) | 
| 58211 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 287 | xtor_co_induct dtor_injects dtor_ctors xtor_co_rec_thms live_nesting_bnfs fpTs Cs Xs | 
| 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 288 | ctrXs_Tsss kss mss ns fp_abs_inverses abs_inverses | 
| 64637 
a15785625f7c
export ML functions (towards nonuniform codatatypes) + signature tuning
 blanchet parents: 
63859diff
changeset | 289 |               (fn thm => thm RS @{thm vimage2p_refl}) ctr_defss ctr_sugars co_recs co_rec_defs
 | 
| 58211 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 290 | |> `(fn ((coinduct_thms_pairs, _), corec_thmss, corec_disc_thmss, _, | 
| 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 291 | (corec_sel_thmsss, _)) => | 
| 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 292 | (map snd coinduct_thms_pairs, map fst coinduct_thms_pairs, corec_thmss, | 
| 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 293 | corec_disc_thmss, corec_sel_thmsss)) | 
| 
c1f3fa32d322
extended 'datatype_compat' to generate the expected, old-style recursor in the presence of recursion through functions
 blanchet parents: 
58180diff
changeset | 294 | ||> (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 | 295 | |
| 62685 | 296 |         val timer = time (timer ("High-level " ^ co_prefix fp ^ "induction principles"));
 | 
| 297 | ||
| 61334 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 traytel parents: 
61125diff
changeset | 298 | val names_lthy = lthy |> fold Variable.declare_typ (As @ Cs @ Xs); | 
| 
8d40ddaa427f
collect the names from goals in favor of fragile exports
 traytel parents: 
61125diff
changeset | 299 | val phi = Proof_Context.export_morphism names_lthy lthy; | 
| 54256 | 300 | |
| 62321 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 blanchet parents: 
61551diff
changeset | 301 | fun mk_target_fp_sugar T X kk pre_bnf absT_info ctrXs_Tss ctor_iff_dtor ctr_defs ctr_sugar | 
| 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 blanchet parents: 
61551diff
changeset | 302 | co_rec co_rec_def map_thms co_inducts co_rec_thms co_rec_disc_thms co_rec_sel_thmss | 
| 58676 | 303 |             ({fp_ctr_sugar = {ctr_transfers, case_transfers, disc_transfers, sel_transfers, ...},
 | 
| 58672 
3f0d612ebd8e
preserve the structure of 'map_sel' theorem in ML
 desharna parents: 
58671diff
changeset | 304 |               fp_bnf_sugar = {map_disc_iffs, map_selss, rel_injects, rel_distincts, rel_sels,
 | 
| 62335 | 305 | rel_intros, rel_cases, pred_injects, set_thms, set_selssss, set_introssss, | 
| 306 | set_cases, ...}, | |
| 63859 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 blanchet parents: 
63802diff
changeset | 307 |               fp_co_induct_sugar = SOME {co_rec_disc_iffs, co_rec_codes, co_rec_transfers,
 | 
| 58734 | 308 | co_rec_o_maps, common_rel_co_inducts, rel_co_inducts, common_set_inducts, | 
| 58732 | 309 | set_inducts, ...}, | 
| 58568 | 310 | ...} : fp_sugar) = | 
| 58117 
9608028d8f43
more compatibility between old- and new-style datatypes
 blanchet parents: 
58115diff
changeset | 311 |           {T = T, BT = T (*wrong but harmless*), X = X, fp = fp, fp_res = fp_res, fp_res_index = kk,
 | 
| 58674 | 312 | pre_bnf = pre_bnf, fp_bnf = nth (#bnfs fp_res) kk, absT_info = absT_info, | 
| 313 | fp_nesting_bnfs = fp_nesting_bnfs, live_nesting_bnfs = live_nesting_bnfs, | |
| 58460 | 314 | fp_ctr_sugar = | 
| 315 |              {ctrXs_Tss = ctrXs_Tss,
 | |
| 62321 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 blanchet parents: 
61551diff
changeset | 316 | ctor_iff_dtor = ctor_iff_dtor, | 
| 58460 | 317 | ctr_defs = ctr_defs, | 
| 58569 | 318 | ctr_sugar = ctr_sugar, | 
| 58570 | 319 | ctr_transfers = ctr_transfers, | 
| 58571 | 320 | case_transfers = case_transfers, | 
| 58676 | 321 | disc_transfers = disc_transfers, | 
| 322 | sel_transfers = sel_transfers}, | |
| 58459 | 323 | fp_bnf_sugar = | 
| 58462 | 324 |              {map_thms = map_thms,
 | 
| 58560 | 325 | map_disc_iffs = map_disc_iffs, | 
| 58672 
3f0d612ebd8e
preserve the structure of 'map_sel' theorem in ML
 desharna parents: 
58671diff
changeset | 326 | map_selss = map_selss, | 
| 58462 | 327 | rel_injects = rel_injects, | 
| 58562 | 328 | rel_distincts = rel_distincts, | 
| 58563 | 329 | rel_sels = rel_sels, | 
| 58564 | 330 | rel_intros = rel_intros, | 
| 58565 | 331 | rel_cases = rel_cases, | 
| 62335 | 332 | pred_injects = pred_injects, | 
| 58566 | 333 | set_thms = set_thms, | 
| 58671 
4cc24f1e52d5
preserve the structure of 'set_sel' theorem in ML
 desharna parents: 
58634diff
changeset | 334 | set_selssss = set_selssss, | 
| 58673 
add1a78da098
preserve the structure of 'set_intros' theorem in ML
 desharna parents: 
58672diff
changeset | 335 | set_introssss = set_introssss, | 
| 58568 | 336 | set_cases = set_cases}, | 
| 63859 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 blanchet parents: 
63802diff
changeset | 337 | fp_co_induct_sugar = SOME | 
| 58461 | 338 |              {co_rec = co_rec,
 | 
| 339 | common_co_inducts = common_co_inducts, | |
| 340 | co_inducts = co_inducts, | |
| 58674 | 341 | co_rec_def = co_rec_def, | 
| 58459 | 342 | co_rec_thms = co_rec_thms, | 
| 343 | co_rec_discs = co_rec_disc_thms, | |
| 58572 | 344 | co_rec_disc_iffs = co_rec_disc_iffs, | 
| 58573 | 345 | co_rec_selss = co_rec_sel_thmss, | 
| 346 | co_rec_codes = co_rec_codes, | |
| 58574 | 347 | co_rec_transfers = co_rec_transfers, | 
| 58734 | 348 | co_rec_o_maps = co_rec_o_maps, | 
| 58575 | 349 | common_rel_co_inducts = common_rel_co_inducts, | 
| 58576 | 350 | rel_co_inducts = rel_co_inducts, | 
| 58577 | 351 | common_set_inducts = common_set_inducts, | 
| 352 | set_inducts = set_inducts}} | |
| 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: 
54284diff
changeset | 353 | |> morph_fp_sugar phi; | 
| 54256 | 354 | |
| 55539 
0819931d652d
simplified data structure by reducing the incidence of clumsy indices
 blanchet parents: 
55486diff
changeset | 355 | val target_fp_sugars = | 
| 62321 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 blanchet parents: 
61551diff
changeset | 356 |           @{map 17} mk_target_fp_sugar fpTs Xs kks pre_bnfs absT_infos ctrXs_Tsss ctor_iff_dtors
 | 
| 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 blanchet parents: 
61551diff
changeset | 357 | ctr_defss ctr_sugars co_recs co_rec_defs mapss (transpose co_inductss') co_rec_thmss | 
| 58634 
9f10d82e8188
added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
 wenzelm parents: 
58577diff
changeset | 358 | co_rec_disc_thmss co_rec_sel_thmsss fp_sugars0; | 
| 55539 
0819931d652d
simplified data structure by reducing the incidence of clumsy indices
 blanchet parents: 
55486diff
changeset | 359 | |
| 
0819931d652d
simplified data structure by reducing the incidence of clumsy indices
 blanchet parents: 
55486diff
changeset | 360 | val n2m_sugar = (target_fp_sugars, fp_sugar_thms); | 
| 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: 
54284diff
changeset | 361 | 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: 
54284diff
changeset | 362 | (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: 
54284diff
changeset | 363 | 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: 
54284diff
changeset | 364 | end; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 365 | |
| 55567 | 366 | fun indexify_callsss ctrs callsss = | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 367 | let | 
| 54283 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 368 | 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 | 369 | (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 | 370 | NONE => replicate (num_binder_types (fastype_of ctr)) [] | 
| 55343 | 371 | | SOME callss => map (map (Envir.beta_eta_contract o unfold_lets_splits)) callss); | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 372 | in | 
| 54283 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 373 | 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 | 374 | end; | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 375 | |
| 54283 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 376 | fun retypargs tyargs (Type (s, _)) = Type (s, tyargs); | 
| 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 377 | |
| 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 378 | 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: 
54282diff
changeset | 379 | 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: 
54282diff
changeset | 380 | | fold_subtype_pairs f TU = f TU; | 
| 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 381 | |
| 55772 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 382 | val impossible_caller = Bound ~1; | 
| 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 383 | |
| 58335 
a5a3b576fcfb
generate 'code' attribute only if 'code' plugin is enabled
 blanchet parents: 
58315diff
changeset | 384 | fun nested_to_mutual_fps plugins fp actual_bs actual_Ts actual_callers actual_callssss0 lthy = | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 385 | let | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 386 | val qsoty = quote o Syntax.string_of_typ lthy; | 
| 62780 
9cfc2b365a8b
reintroduced check that may guard some tactic failures
 blanchet parents: 
62722diff
changeset | 387 | val qsotys = space_implode " or " o map qsoty; | 
| 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 | 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 | 390 | 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 | 391 | if fp = Least_FP andalso | 
| 58112 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 blanchet parents: 
57983diff
changeset | 392 | is_some (Old_Datatype_Data.get_info (Proof_Context.theory_of lthy) s) then | 
| 58315 | 393 | error (qsoty T ^ " is an old-style datatype") | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 394 | else | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 395 | not_co_datatype0 T | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 396 | | not_co_datatype T = not_co_datatype0 T; | 
| 62780 
9cfc2b365a8b
reintroduced check that may guard some tactic failures
 blanchet parents: 
62722diff
changeset | 397 | fun not_mutually_nested_rec Ts1 Ts2 = | 
| 
9cfc2b365a8b
reintroduced check that may guard some tactic failures
 blanchet parents: 
62722diff
changeset | 398 | error (qsotys Ts1 ^ " is neither mutually " ^ co_prefix fp ^ "recursive with " ^ qsotys Ts2 ^ | 
| 
9cfc2b365a8b
reintroduced check that may guard some tactic failures
 blanchet parents: 
62722diff
changeset | 399 | " nor nested " ^ co_prefix fp ^ "recursive through " ^ | 
| 
9cfc2b365a8b
reintroduced check that may guard some tactic failures
 blanchet parents: 
62722diff
changeset | 400 | (if Ts1 = Ts2 andalso length Ts1 = 1 then "itself" else qsotys Ts2)); | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 401 | |
| 55480 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55479diff
changeset | 402 | val sorted_actual_Ts = | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
59040diff
changeset | 403 | sort (prod_ord int_ord Term_Ord.typ_ord o apply2 (`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 | 404 | |
| 54283 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 405 | 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: 
54282diff
changeset | 406 | |
| 62688 | 407 | fun gen_rhss_in gen_Ts rho (subTs as Type (_, sub_tyargs) :: _) = | 
| 408 | let | |
| 409 | fun maybe_insert (T, Type (_, gen_tyargs)) = | |
| 410 | member (op =) subTs T ? insert (op =) gen_tyargs | |
| 411 | | maybe_insert _ = I; | |
| 412 | ||
| 413 | val gen_ctrs = maps the_ctrs_of gen_Ts; | |
| 414 | val gen_ctr_Ts = maps (binder_types o fastype_of) gen_ctrs; | |
| 415 | val ctr_Ts = map (Term.typ_subst_atomic rho) gen_ctr_Ts; | |
| 416 | in | |
| 417 | (case fold (fold_subtype_pairs maybe_insert) (ctr_Ts ~~ gen_ctr_Ts) [] of | |
| 418 | [] => [map (typ_subst_nonatomic (map swap rho)) sub_tyargs] | |
| 419 | | gen_tyargss => gen_tyargss) | |
| 420 | end; | |
| 421 | ||
| 54282 | 422 | fun the_fp_sugar_of (T as Type (T_name, _)) = | 
| 423 | (case fp_sugar_of lthy T_name of | |
| 63859 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 blanchet parents: 
63802diff
changeset | 424 |         SOME (fp_sugar as {fp = fp', fp_co_induct_sugar = SOME _, ...}) =>
 | 
| 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 blanchet parents: 
63802diff
changeset | 425 | if fp = fp' then fp_sugar else not_co_datatype T | 
| 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 blanchet parents: 
63802diff
changeset | 426 | | _ => not_co_datatype T); | 
| 54282 | 427 | |
| 56485 | 428 | fun gather_types _ _ rev_seens gen_seen [] = (rev rev_seens, gen_seen) | 
| 429 | | gather_types lthy rho rev_seens gen_seen ((T as Type (_, tyargs)) :: Ts) = | |
| 54282 | 430 | let | 
| 58234 | 431 |           val {T = Type (_, tyargs0), fp_res = {Ts = mutual_Ts0, ...}, ...} = the_fp_sugar_of T;
 | 
| 54283 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 432 | val mutual_Ts = map (retypargs tyargs) mutual_Ts0; | 
| 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 433 | |
| 62780 
9cfc2b365a8b
reintroduced check that may guard some tactic failures
 blanchet parents: 
62722diff
changeset | 434 | val rev_seen = flat rev_seens; | 
| 
9cfc2b365a8b
reintroduced check that may guard some tactic failures
 blanchet parents: 
62722diff
changeset | 435 | val _ = null rev_seens orelse exists (exists_strict_subtype_in rev_seen) mutual_Ts orelse | 
| 
9cfc2b365a8b
reintroduced check that may guard some tactic failures
 blanchet parents: 
62722diff
changeset | 436 | not_mutually_nested_rec mutual_Ts rev_seen; | 
| 
9cfc2b365a8b
reintroduced check that may guard some tactic failures
 blanchet parents: 
62722diff
changeset | 437 | |
| 54283 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 438 | fun fresh_tyargs () = | 
| 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 439 | let | 
| 58234 | 440 | val (unsorted_gen_tyargs, lthy') = | 
| 58280 
2ec3e2de34c3
more canonical (and correct) local theory threading
 blanchet parents: 
58236diff
changeset | 441 | variant_tfrees (replicate (length tyargs) "z") lthy | 
| 54283 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 442 | |>> map Logic.varifyT_global; | 
| 58234 | 443 | val gen_tyargs = | 
| 444 | map2 (resort_tfree_or_tvar o snd o dest_TFree_or_TVar) tyargs0 unsorted_gen_tyargs; | |
| 54283 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 445 | val rho' = (gen_tyargs ~~ tyargs) @ rho; | 
| 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 446 | in | 
| 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 447 | (rho', gen_tyargs, gen_seen, lthy') | 
| 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 448 | end; | 
| 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 449 | |
| 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 450 | val (rho', gen_tyargs, gen_seen', lthy') = | 
| 62688 | 451 | if exists (exists_subtype_in (flat rev_seens)) mutual_Ts then | 
| 54283 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 452 | (case gen_rhss_in gen_seen rho mutual_Ts of | 
| 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 453 | [] => fresh_tyargs () | 
| 54899 
7a01387c47d5
added tactic to prove 'disc_iff' properties in 'primcorec'
 blanchet parents: 
54740diff
changeset | 454 | | gen_tyargs :: gen_tyargss_tl => | 
| 54283 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 455 | let | 
| 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 456 | 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: 
54282diff
changeset | 457 | val mgu = Type.raw_unifys unify_pairs Vartab.empty; | 
| 58948 
f6ecc0fb2066
proper Envir.norm_type for result of Type.raw_unifys;
 wenzelm parents: 
58734diff
changeset | 458 | val gen_tyargs' = map (Envir.norm_type mgu) gen_tyargs; | 
| 
f6ecc0fb2066
proper Envir.norm_type for result of Type.raw_unifys;
 wenzelm parents: 
58734diff
changeset | 459 | val gen_seen' = map (Envir.norm_type mgu) gen_seen; | 
| 54283 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 460 | in | 
| 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 461 | (rho, gen_tyargs', gen_seen', lthy) | 
| 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 462 | end) | 
| 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 463 | else | 
| 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 464 | fresh_tyargs (); | 
| 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 465 | |
| 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 466 | val gen_mutual_Ts = map (retypargs gen_tyargs) mutual_Ts0; | 
| 55480 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55479diff
changeset | 467 | val other_mutual_Ts = remove1 (op =) T mutual_Ts; | 
| 
59cc4a8bc28a
allow different functions to recurse on the same type, like in the old package
 blanchet parents: 
55479diff
changeset | 468 | val Ts' = fold (remove1 (op =)) other_mutual_Ts Ts; | 
| 54282 | 469 | in | 
| 56485 | 470 | gather_types lthy' rho' (mutual_Ts :: rev_seens) (gen_seen' @ gen_mutual_Ts) Ts' | 
| 54282 | 471 | end | 
| 56485 | 472 | | 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 | 473 | |
| 56485 | 474 | val (perm_Tss, perm_gen_Ts) = gather_types lthy [] [] [] sorted_actual_Ts; | 
| 58340 | 475 | val (perm_mutual_cliques, perm_Ts) = | 
| 56485 | 476 | split_list (flat (map_index (fn (i, perm_Ts) => map (pair i) perm_Ts) perm_Tss)); | 
| 477 | ||
| 54283 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 478 | val perm_frozen_gen_Ts = map Logic.unvarifyT_global perm_gen_Ts; | 
| 
6f0a49ed1bb1
fourth attempt at generalizing N2M types (to leverage caching)
 blanchet parents: 
54282diff
changeset | 479 | |
| 56484 | 480 | val missing_Ts = fold (remove1 (op =)) actual_Ts perm_Ts; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 481 | 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 | 482 | |
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 483 | val nn = length Ts; | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 484 | 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 | 485 | |
| 54267 | 486 | val callssss0 = pad_list [] nn actual_callssss0; | 
| 487 | ||
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 488 | 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 | 489 | val bs = pad_list (Binding.name common_name) nn actual_bs; | 
| 55772 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 490 | val callers = pad_list impossible_caller nn actual_callers; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 491 | |
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 492 | 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 | 493 | 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 | 494 | |
| 58340 | 495 | (* The assignment of callers to mutual cliques is incorrect in general. This code would need to | 
| 496 | inspect the actual calls to discover the correct cliques in the presence of type duplicates. | |
| 497 | However, the naive scheme implemented here supports "prim(co)rec" specifications following | |
| 498 | reasonable ordering of the duplicates (e.g., keeping the cliques together). *) | |
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 499 | val perm_bs = permute bs; | 
| 55772 
367ec44763fd
correct most general type for mutual recursion when several identical types are involved
 blanchet parents: 
55702diff
changeset | 500 | val perm_callers = permute callers; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 501 | val perm_kks = permute kks; | 
| 54267 | 502 | 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 | 503 | 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 | 504 | |
| 58460 | 505 | val perm_callssss = | 
| 506 | map2 (indexify_callsss o #ctrs o #ctr_sugar o #fp_ctr_sugar) 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 | 507 | |
| 53746 
bd038e48526d
have "datatype_new_compat" register induction and recursion theorems in nested case
 blanchet parents: 
53678diff
changeset | 508 | val ((perm_fp_sugars, fp_sugar_thms), lthy) = | 
| 58117 
9608028d8f43
more compatibility between old- and new-style datatypes
 blanchet parents: 
58115diff
changeset | 509 | if length perm_Tss = 1 then | 
| 56485 | 510 | ((perm_fp_sugars0, (NONE, NONE)), lthy) | 
| 511 | else | |
| 58340 | 512 | mutualize_fp_sugars plugins fp perm_mutual_cliques perm_bs perm_frozen_gen_Ts perm_callers | 
| 58335 
a5a3b576fcfb
generate 'code' attribute only if 'code' plugin is enabled
 blanchet parents: 
58315diff
changeset | 513 | perm_callssss perm_fp_sugars0 lthy; | 
| 53303 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 514 | |
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 515 | 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 | 516 | in | 
| 53746 
bd038e48526d
have "datatype_new_compat" register induction and recursion theorems in nested case
 blanchet parents: 
53678diff
changeset | 517 | ((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 | 518 | end; | 
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 519 | |
| 
ae49b835ca01
moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
 blanchet parents: diff
changeset | 520 | end; |