src/HOL/Tools/BNF/bnf_fp_n2m.ML
author traytel
Tue, 25 Feb 2014 18:14:26 +0100
changeset 55803 74d3fe9031d8
parent 55575 a5e33e18fb5c
child 55810 63d63d854fae
permissions -rw-r--r--
joint work with blanchet: intermediate typedef for the input to fp-operations
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
55061
a0adf838e2d1 adjusted comments
blanchet
parents: 55060
diff changeset
     1
(*  Title:      HOL/Tools/BNF/bnf_fp_n2m.ML
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     2
    Author:     Dmitriy Traytel, 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
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 =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
     9
sig
55575
a5e33e18fb5c moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
blanchet
parents: 55566
diff changeset
    10
  val construct_mutualized_fp: BNF_Util.fp_kind -> typ list -> BNF_FP_Def_Sugar.fp_sugar list ->
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    11
    binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list ->
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    12
    BNF_Comp.absT_info list -> local_theory -> BNF_FP_Util.fp_result * local_theory
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    13
end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    14
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    15
structure BNF_FP_N2M : BNF_FP_N2M =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    16
struct
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    17
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    18
open BNF_Def
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    19
open BNF_Util
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    20
open BNF_Comp
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    21
open BNF_FP_Util
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    22
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
    23
open BNF_Tactics
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    24
open BNF_FP_N2M_Tactics
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    25
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    26
fun force_typ ctxt T =
54233
6d64669184ae conceal definition
blanchet
parents: 53753
diff changeset
    27
  map_types Type_Infer.paramify_vars
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    28
  #> Type.constraint T
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    29
  #> Syntax.check_term ctxt
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    30
  #> singleton (Variable.polymorphic ctxt);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    31
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    32
fun mk_prod_map f g =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    33
  let
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    34
    val ((fAT, fBT), fT) = `dest_funT (fastype_of f);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    35
    val ((gAT, gBT), gT) = `dest_funT (fastype_of g);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    36
  in
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    37
    Const (@{const_name map_pair},
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    38
      fT --> gT --> HOLogic.mk_prodT (fAT, gAT) --> HOLogic.mk_prodT (fBT, gBT)) $ f $ g
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    39
  end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    40
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    41
fun mk_sum_map f g =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    42
  let
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    43
    val ((fAT, fBT), fT) = `dest_funT (fastype_of f);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    44
    val ((gAT, gBT), gT) = `dest_funT (fastype_of g);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    45
  in
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    46
    Const (@{const_name sum_map}, fT --> gT --> mk_sumT (fAT, gAT) --> mk_sumT (fBT, gBT)) $ f $ g
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    47
  end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    48
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    49
fun construct_mutualized_fp fp fpTs (fp_sugars : fp_sugar list) bs resBs (resDs, Dss) bnfs
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    50
    absT_infos lthy =
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    51
  let
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    52
    fun of_fp_res get =
55539
0819931d652d simplified data structure by reducing the incidence of clumsy indices
blanchet
parents: 55531
diff changeset
    53
      map (fn {fp_res, fp_res_index, ...} => nth (get fp_res) fp_res_index) fp_sugars;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    54
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    55
    fun mk_co_algT T U = fp_case fp (T --> U) (U --> T);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    56
    fun co_swap pair = fp_case fp I swap pair;
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    57
    val mk_co_comp = HOLogic.mk_comp o co_swap;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    58
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    59
    val dest_co_algT = co_swap o dest_funT;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    60
    val co_alg_argT = fp_case fp range_type domain_type;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    61
    val co_alg_funT = fp_case fp domain_type range_type;
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55394
diff changeset
    62
    val mk_co_product = curry (fp_case fp mk_convol mk_case_sum);
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    63
    val mk_map_co_product = fp_case fp mk_prod_map mk_sum_map;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    64
    val co_proj1_const = fp_case fp (fst_const o fst) (uncurry Inl_const o dest_sumT o snd);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    65
    val mk_co_productT = curry (fp_case fp HOLogic.mk_prodT mk_sumT);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
    66
    val dest_co_productT = fp_case fp HOLogic.dest_prodT dest_sumT;
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    67
    val rewrite_comp_comp = fp_case fp @{thm rewriteL_comp_comp} @{thm rewriteR_comp_comp};
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    68
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    69
    val fp_absT_infos = map #absT_info fp_sugars;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    70
    val fp_bnfs = of_fp_res #bnfs;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    71
    val pre_bnfs = map #pre_bnf fp_sugars;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    72
    val nesty_bnfss = map (fn sugar => #nested_bnfs sugar @ #nesting_bnfs sugar) fp_sugars;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    73
    val fp_nesty_bnfss = fp_bnfs :: nesty_bnfss;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    74
    val fp_nesty_bnfs = distinct (op = o pairself T_of_bnf) (flat fp_nesty_bnfss);
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    75
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    76
    val fp_absTs = map #absT fp_absT_infos;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    77
    val fp_repTs = map #repT fp_absT_infos;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    78
    val fp_abss = map #abs fp_absT_infos;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    79
    val fp_reps = map #rep fp_absT_infos;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    80
    val fp_type_definitions = map #type_definition fp_absT_infos;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    81
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    82
    val absTs = map #absT absT_infos;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    83
    val repTs = map #repT absT_infos;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    84
    val absTs' = map (Logic.type_map (singleton (Variable.polymorphic lthy))) absTs;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    85
    val repTs' = map (Logic.type_map (singleton (Variable.polymorphic lthy))) repTs;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    86
    val abss = map #abs absT_infos;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    87
    val reps = map #rep absT_infos;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    88
    val abs_inverses = map #abs_inverse absT_infos;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    89
    val type_definitions = map #type_definition absT_infos;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    90
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    91
    val n = length bnfs;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    92
    val deads = fold (union (op =)) Dss resDs;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    93
    val As = subtract (op =) deads (map TFree resBs);
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    94
    val names_lthy = fold Variable.declare_typ (As @ deads) lthy;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    95
    val m = length As;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    96
    val live = m + n;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    97
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    98
    val ((Xs, Bs), names_lthy) = names_lthy
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
    99
      |> mk_TFrees n
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   100
      ||>> mk_TFrees m;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   101
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   102
    val allAs = As @ Xs;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   103
    val allBs = Bs @ Xs;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   104
    val phiTs = map2 mk_pred2T As Bs;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   105
    val thetaBs = As ~~ Bs;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   106
    val fpTs' = map (Term.typ_subst_atomic thetaBs) fpTs;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   107
    val fold_thetaAs = Xs ~~ fpTs;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   108
    val fold_thetaBs = Xs ~~ fpTs';
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   109
    val rec_theta = Xs ~~ map2 mk_co_productT fpTs Xs;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   110
    val pre_phiTs = map2 mk_pred2T fpTs fpTs';
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   111
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   112
    val ((ctors, dtors), (xtor's, xtors)) =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   113
      let
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   114
        val ctors = map2 (force_typ names_lthy o (fn T => dummyT --> T)) fpTs (of_fp_res #ctors);
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   115
        val dtors = map2 (force_typ names_lthy o (fn T => T --> dummyT)) fpTs (of_fp_res #dtors);
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   116
      in
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   117
        ((ctors, dtors), `(map (Term.subst_atomic_types thetaBs)) (fp_case fp ctors dtors))
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   118
      end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   119
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   120
    val absATs = map (domain_type o fastype_of) ctors;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   121
    val absBTs = map (Term.typ_subst_atomic thetaBs) absATs;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   122
    val xTs = map (domain_type o fastype_of) xtors;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   123
    val yTs = map (domain_type o fastype_of) xtor's;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   124
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   125
    fun abs_of allAs Ds bnf = mk_abs (mk_T_of_bnf Ds allAs bnf) o #abs;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   126
    fun rep_of absAT = mk_rep absAT o #rep;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   127
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   128
    val absAs = map3 (abs_of allAs) Dss bnfs absT_infos;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   129
    val absBs = map3 (abs_of allBs) Dss bnfs absT_infos;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   130
    val fp_repAs = map2 rep_of absATs fp_absT_infos;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   131
    val fp_repBs = map2 rep_of absBTs fp_absT_infos;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   132
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   133
    val (((((phis, phis'), pre_phis), xs), ys), names_lthy) = names_lthy
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   134
      |> mk_Frees' "R" phiTs
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   135
      ||>> mk_Frees "S" pre_phiTs
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   136
      ||>> mk_Frees "x" xTs
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   137
      ||>> mk_Frees "y" yTs;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   138
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   139
    val rels =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   140
      let
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   141
        fun find_rel T As Bs = fp_nesty_bnfss
55394
d5ebe055b599 more liberal merging of BNFs and constructor sugar
blanchet
parents: 55067
diff changeset
   142
          |> map (filter_out (curry (op = o pairself name_of_bnf) BNF_Comp.DEADID_bnf))
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   143
          |> get_first (find_first (fn bnf => Type.could_unify (T_of_bnf bnf, T)))
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   144
          |> Option.map (fn bnf =>
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   145
            let val live = live_of_bnf bnf;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   146
            in (mk_rel live As Bs (rel_of_bnf bnf), live) end)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   147
          |> the_default (HOLogic.eq_const T, 0);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   148
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   149
        fun mk_rel (T as Type (_, Ts)) (Type (_, Us)) =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   150
              let
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   151
                val (rel, live) = find_rel T Ts Us;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   152
                val (Ts', Us') = fastype_of rel |> strip_typeN live |> fst |> map_split dest_pred2T;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   153
                val rels = map2 mk_rel Ts' Us';
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   154
              in
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   155
                Term.list_comb (rel, rels)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   156
              end
54244
0753e8866ac8 more robust treatment of dead variables in n2m
traytel
parents: 54242
diff changeset
   157
          | mk_rel (T as TFree _) _ = (nth phis (find_index (curry op = T) As)
0753e8866ac8 more robust treatment of dead variables in n2m
traytel
parents: 54242
diff changeset
   158
              handle General.Subscript => HOLogic.eq_const T)
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   159
          | mk_rel _ _ = raise Fail "fpTs contains schematic type variables";
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   160
      in
54298
347c3b0cab44 handle independent functions defined in parallel in N2M (in presence of type variables, see ce58fb149ff6)
traytel
parents: 54244
diff changeset
   161
        map2 (fold_rev Term.absfree phis' oo mk_rel) fpTs fpTs'
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   162
      end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   163
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   164
    val pre_rels = map2 (fn Ds => mk_rel_of_bnf Ds (As @ fpTs) (Bs @ fpTs')) Dss bnfs;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   165
55539
0819931d652d simplified data structure by reducing the incidence of clumsy indices
blanchet
parents: 55531
diff changeset
   166
    val rel_unfolds = maps (no_refl o single o rel_def_of_bnf) pre_bnfs;
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   167
    val rel_xtor_co_inducts = of_fp_res (split_conj_thm o #rel_xtor_co_induct_thm)
55539
0819931d652d simplified data structure by reducing the incidence of clumsy indices
blanchet
parents: 55531
diff changeset
   168
      |> map (unfold_thms lthy (id_apply :: rel_unfolds));
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   169
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   170
    val rel_defs = map rel_def_of_bnf bnfs;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   171
    val rel_monos = map rel_mono_of_bnf bnfs;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   172
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   173
    fun cast castA castB pre_rel =
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   174
      let
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   175
        val castAB = mk_vimage2p (Term.subst_atomic_types fold_thetaAs castA)
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   176
          (Term.subst_atomic_types fold_thetaBs castB);
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   177
      in
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   178
        fold_rev (fold_rev Term.absdummy) [phiTs, pre_phiTs]
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   179
          (castAB $ Term.list_comb (pre_rel, map Bound (live - 1 downto 0)))
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   180
      end;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   181
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   182
    val castAs = map2 (curry HOLogic.mk_comp) absAs fp_repAs;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   183
    val castBs = map2 (curry HOLogic.mk_comp) absBs fp_repBs;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   184
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   185
    val rel_xtor_co_induct_thm =
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   186
      mk_rel_xtor_co_induct_thm fp (map3 cast castAs castBs pre_rels) pre_phis rels phis xs ys xtors
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   187
        xtor's (mk_rel_xtor_co_induct_tactic fp abs_inverses rel_xtor_co_inducts rel_defs rel_monos)
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   188
        lthy;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   189
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   190
    val rel_eqs = no_refl (map rel_eq_of_bnf fp_nesty_bnfs);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   191
    val map_id0s = no_refl (map map_id0_of_bnf bnfs);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   192
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   193
    val xtor_co_induct_thm =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   194
      (case fp of
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   195
        Least_FP =>
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   196
          let
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   197
            val (Ps, names_lthy) = names_lthy
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   198
              |> mk_Frees "P" (map (fn T => T --> HOLogic.boolT) fpTs);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   199
            fun mk_Grp_id P =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   200
              let val T = domain_type (fastype_of P);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   201
              in mk_Grp (HOLogic.Collect_const T $ P) (HOLogic.id_const T) end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   202
            val cts = map (SOME o certify lthy) (map HOLogic.eq_const As @ map mk_Grp_id Ps);
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   203
            fun mk_fp_type_copy_thms thm = map (curry op RS thm)
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   204
              @{thms type_copy_Abs_o_Rep type_copy_vimage2p_Grp_Rep};
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   205
            fun mk_type_copy_thms thm = map (curry op RS thm)
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   206
              @{thms type_copy_Rep_o_Abs type_copy_vimage2p_Grp_Abs};
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   207
          in
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   208
            cterm_instantiate_pos cts rel_xtor_co_induct_thm
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   209
            |> singleton (Proof_Context.export names_lthy lthy)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   210
            |> unfold_thms lthy (@{thms eq_le_Grp_id_iff all_simps(1,2)[symmetric]} @ rel_eqs)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   211
            |> funpow n (fn thm => thm RS spec)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   212
            |> unfold_thms lthy (@{thm eq_alt} :: map rel_Grp_of_bnf bnfs @ map_id0s)
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   213
            |> unfold_thms lthy (@{thms vimage2p_comp comp_apply comp_id
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   214
               Grp_id_mono_subst eqTrueI[OF subset_UNIV] simp_thms(22)} @
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   215
               maps mk_fp_type_copy_thms fp_type_definitions @
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   216
               maps mk_type_copy_thms type_definitions)
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   217
            |> unfold_thms lthy @{thms subset_iff mem_Collect_eq
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   218
               atomize_conjL[symmetric] atomize_all[symmetric] atomize_imp[symmetric]}
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   219
          end
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   220
      | Greatest_FP =>
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   221
          let
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   222
            val cts = NONE :: map (SOME o certify lthy) (map HOLogic.eq_const As);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   223
          in
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   224
            cterm_instantiate_pos cts rel_xtor_co_induct_thm
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   225
            |> unfold_thms lthy (@{thms le_fun_def le_bool_def all_simps(1,2)[symmetric]} @ rel_eqs)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   226
            |> funpow (2 * n) (fn thm => thm RS spec)
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 54740
diff changeset
   227
            |> Conv.fconv_rule (Object_Logic.atomize lthy)
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   228
            |> funpow n (fn thm => thm RS mp)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   229
          end);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   230
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   231
    val fold_preTs = map2 (fn Ds => mk_T_of_bnf Ds allAs) Dss bnfs;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   232
    val rec_preTs = map (Term.typ_subst_atomic rec_theta) fold_preTs;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   233
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   234
    val fold_strTs = map2 mk_co_algT fold_preTs Xs;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   235
    val rec_strTs = map2 mk_co_algT rec_preTs Xs;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   236
    val resTs = map2 mk_co_algT fpTs Xs;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   237
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   238
    val (((fold_strs, fold_strs'), (rec_strs, rec_strs')), names_lthy) = names_lthy
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   239
      |> mk_Frees' "s" fold_strTs
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   240
      ||>> mk_Frees' "s" rec_strTs;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   241
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   242
    val co_iters = of_fp_res #xtor_co_iterss;
55539
0819931d652d simplified data structure by reducing the incidence of clumsy indices
blanchet
parents: 55531
diff changeset
   243
    val ns = map (length o #Ts o #fp_res) fp_sugars;
0819931d652d simplified data structure by reducing the incidence of clumsy indices
blanchet
parents: 55531
diff changeset
   244
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   245
    fun substT rho (Type (@{type_name "fun"}, [T, U])) = substT rho T --> substT rho U
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   246
      | substT rho (Type (s, Ts)) = Type (s, map (typ_subst_nonatomic rho) Ts)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   247
      | substT _ T = T;
55539
0819931d652d simplified data structure by reducing the incidence of clumsy indices
blanchet
parents: 55531
diff changeset
   248
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   249
    val typ_subst_nonatomic_sorted = fold_rev (typ_subst_nonatomic o single);
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   250
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   251
    fun force_iter is_rec i TU TU_rec raw_iters =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   252
      let
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   253
        val thy = Proof_Context.theory_of lthy;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   254
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   255
        val approx_fold = un_fold_of raw_iters
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   256
          |> force_typ names_lthy
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   257
            (replicate (nth ns i) dummyT ---> (if is_rec then TU_rec else TU));
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   258
        val subst = Term.typ_subst_atomic fold_thetaAs;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   259
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   260
        fun mk_fp_absT_repT fp_repT fp_absT = mk_absT thy fp_repT fp_absT ooo mk_repT;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   261
        val mk_fp_absT_repTs = map5 mk_fp_absT_repT fp_repTs fp_absTs absTs repTs;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   262
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   263
        val fold_preTs' = mk_fp_absT_repTs (map subst fold_preTs);
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   264
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   265
        val fold_pre_deads_only_Ts =
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   266
          map (typ_subst_nonatomic_sorted (map (rpair dummyT) (As @ fpTs))) fold_preTs';
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   267
55478
3a6efda01da4 made N2M more robust w.r.t. identical nested types
traytel
parents: 55414
diff changeset
   268
        val TUs = map_split dest_co_algT (binder_fun_types (fastype_of approx_fold))
3a6efda01da4 made N2M more robust w.r.t. identical nested types
traytel
parents: 55414
diff changeset
   269
          |>> map subst
3a6efda01da4 made N2M more robust w.r.t. identical nested types
traytel
parents: 55414
diff changeset
   270
          |> uncurry (map2 mk_co_algT);
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   271
        val cands = map2 mk_co_algT fold_preTs' Xs;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   272
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   273
        val js = find_indices Type.could_unify TUs cands;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   274
        val Tpats = map (fn j => mk_co_algT (nth fold_pre_deads_only_Ts j) (nth Xs j)) js;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   275
        val iter = raw_iters |> (if is_rec then co_rec_of else un_fold_of);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   276
      in
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   277
        force_typ names_lthy (Tpats ---> TU) iter
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   278
      end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   279
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   280
    fun mk_co_comp_abs_rep fp_absT absT fp_abs fp_rep abs rep t =
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   281
      fp_case fp (HOLogic.mk_comp (HOLogic.mk_comp (t, mk_abs absT abs), mk_rep fp_absT fp_rep))
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   282
        (HOLogic.mk_comp (mk_abs fp_absT fp_abs, HOLogic.mk_comp (mk_rep absT rep, t)));
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   283
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   284
    fun mk_iter b_opt is_rec iters lthy TU =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   285
      let
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   286
        val thy = Proof_Context.theory_of lthy;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   287
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   288
        val x = co_alg_argT TU;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   289
        val i = find_index (fn T => x = T) Xs;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   290
        val TUiter =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   291
          (case find_first (fn f => body_fun_type (fastype_of f) = TU) iters of
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   292
            NONE => nth co_iters i
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   293
              |> force_iter is_rec i
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   294
                (TU |> (is_none b_opt andalso not is_rec) ? substT (fpTs ~~ Xs))
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   295
                (TU |> (is_none b_opt) ? substT (map2 mk_co_productT fpTs Xs ~~ Xs))
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   296
          | SOME f => f);
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   297
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   298
        val TUs = binder_fun_types (fastype_of TUiter);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   299
        val iter_preTs = if is_rec then rec_preTs else fold_preTs;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   300
        val iter_strs = if is_rec then rec_strs else fold_strs;
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   301
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   302
        fun mk_s TU' =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   303
          let
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   304
            fun mk_absT_fp_repT repT absT = mk_absT thy repT absT ooo mk_repT;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   305
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   306
            val i = find_index (fn T => co_alg_argT TU' = T) Xs;
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   307
            val fp_abs = nth fp_abss i;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   308
            val fp_rep = nth fp_reps i;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   309
            val abs = nth abss i;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   310
            val rep = nth reps i;
54233
6d64669184ae conceal definition
blanchet
parents: 53753
diff changeset
   311
            val sF = co_alg_funT TU';
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   312
            val sF' =
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   313
              mk_absT_fp_repT (nth repTs' i) (nth absTs' i) (nth fp_absTs i) (nth fp_repTs i) sF
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   314
                handle Term.TYPE _ => sF;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   315
            val F = nth iter_preTs i;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   316
            val s = nth iter_strs i;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   317
          in
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   318
            if sF = F then s
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   319
            else if sF' = F then mk_co_comp_abs_rep sF sF' fp_abs fp_rep abs rep s
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   320
            else
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   321
              let
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   322
                val smapT = replicate live dummyT ---> mk_co_algT sF' F;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   323
                fun hidden_to_unit t =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   324
                  Term.subst_TVars (map (rpair HOLogic.unitT) (Term.add_tvar_names t [])) t;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   325
                val smap = map_of_bnf (nth bnfs i)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   326
                  |> force_typ names_lthy smapT
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   327
                  |> hidden_to_unit;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   328
                val smap_argTs = strip_typeN live (fastype_of smap) |> fst;
54233
6d64669184ae conceal definition
blanchet
parents: 53753
diff changeset
   329
                fun mk_smap_arg TU =
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   330
                  (if domain_type TU = range_type TU then
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   331
                    HOLogic.id_const (domain_type TU)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   332
                  else if is_rec then
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   333
                    let
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   334
                      val (TY, (U, X)) = TU |> dest_co_algT ||> dest_co_productT;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   335
                      val T = mk_co_algT TY U;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   336
                    in
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   337
                      (case try (force_typ lthy T o build_map lthy co_proj1_const o dest_funT) T of
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   338
                        SOME f => mk_co_product f
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   339
                          (fst (fst (mk_iter NONE is_rec iters lthy (mk_co_algT TY X))))
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   340
                      | NONE => mk_map_co_product
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   341
                          (build_map lthy co_proj1_const
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   342
                            (dest_funT (mk_co_algT (dest_co_productT TY |> fst) U)))
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   343
                          (HOLogic.id_const X))
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   344
                    end
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   345
                  else
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   346
                    fst (fst (mk_iter NONE is_rec iters lthy TU)))
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   347
                val smap_args = map mk_smap_arg smap_argTs;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   348
              in
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   349
                mk_co_comp_abs_rep sF sF' fp_abs fp_rep abs rep
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   350
                  (mk_co_comp (s, Term.list_comb (smap, smap_args)))
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   351
              end
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   352
          end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   353
        val t = Term.list_comb (TUiter, map mk_s TUs);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   354
      in
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   355
        (case b_opt of
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   356
          NONE => ((t, Drule.dummy_thm), lthy)
54233
6d64669184ae conceal definition
blanchet
parents: 53753
diff changeset
   357
        | SOME b => Local_Theory.define ((b, NoSyn), ((Binding.conceal (Thm.def_binding b), []),
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   358
            fold_rev Term.absfree (if is_rec then rec_strs' else fold_strs') t)) lthy |>> apsnd snd)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   359
      end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   360
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   361
    fun mk_iters is_rec name lthy =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   362
      fold2 (fn TU => fn b => fn ((iters, defs), lthy) =>
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   363
        mk_iter (SOME b) is_rec iters lthy TU |>> (fn (f, d) => (f :: iters, d :: defs)))
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   364
      resTs (map (Binding.suffix_name ("_" ^ name)) bs) (([], []), lthy)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   365
      |>> apfst rev o apsnd rev;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   366
    val foldN = fp_case fp ctor_foldN dtor_unfoldN;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   367
    val recN = fp_case fp ctor_recN dtor_corecN;
53331
20440c789759 prove theorem in the right context (that knows about local variables)
traytel
parents: 53309
diff changeset
   368
    val (((raw_un_folds, raw_un_fold_defs), (raw_co_recs, raw_co_rec_defs)), (lthy, raw_lthy)) =
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   369
      lthy
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   370
      |> mk_iters false foldN
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   371
      ||>> mk_iters true recN
53331
20440c789759 prove theorem in the right context (that knows about local variables)
traytel
parents: 53309
diff changeset
   372
      ||> `Local_Theory.restore;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   373
53331
20440c789759 prove theorem in the right context (that knows about local variables)
traytel
parents: 53309
diff changeset
   374
    val phi = Proof_Context.export_morphism raw_lthy lthy;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   375
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   376
    val un_folds = map (Morphism.term phi) raw_un_folds;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   377
    val co_recs = map (Morphism.term phi) raw_co_recs;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   378
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   379
    val (xtor_un_fold_thms, xtor_co_rec_thms) =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   380
      let
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   381
        val folds = map (fn f => Term.list_comb (f, fold_strs)) raw_un_folds;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   382
        val recs = map (fn r => Term.list_comb (r, rec_strs)) raw_co_recs;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   383
        val fold_mapTs = co_swap (As @ fpTs, As @ Xs);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   384
        val rec_mapTs = co_swap (As @ fpTs, As @ map2 mk_co_productT fpTs Xs);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   385
        val pre_fold_maps =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   386
          map2 (fn Ds => fn bnf =>
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   387
            Term.list_comb (uncurry (mk_map_of_bnf Ds) fold_mapTs bnf,
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   388
              map HOLogic.id_const As @ folds))
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   389
          Dss bnfs;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   390
        val pre_rec_maps =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   391
          map2 (fn Ds => fn bnf =>
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   392
            Term.list_comb (uncurry (mk_map_of_bnf Ds) rec_mapTs bnf,
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   393
              map HOLogic.id_const As @ map2 (mk_co_product o HOLogic.id_const) fpTs recs))
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   394
          Dss bnfs;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   395
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   396
        fun mk_goals f xtor s smap fp_abs fp_rep abs rep =
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   397
          let
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   398
            val lhs = mk_co_comp (f, xtor);
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   399
            val rhs = mk_co_comp (s, smap);
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   400
          in
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   401
            HOLogic.mk_eq (lhs,
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   402
              mk_co_comp_abs_rep (co_alg_funT (fastype_of lhs)) (co_alg_funT (fastype_of rhs))
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   403
                fp_abs fp_rep abs rep rhs)
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   404
          end;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   405
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   406
        val fold_goals =
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   407
          map8 mk_goals folds xtors fold_strs pre_fold_maps fp_abss fp_reps abss reps;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   408
        val rec_goals = map8 mk_goals recs xtors rec_strs pre_rec_maps fp_abss fp_reps abss reps;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   409
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   410
        fun mk_thms ss goals tac =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   411
          Library.foldr1 HOLogic.mk_conj goals
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   412
          |> HOLogic.mk_Trueprop
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   413
          |> fold_rev Logic.all ss
53331
20440c789759 prove theorem in the right context (that knows about local variables)
traytel
parents: 53309
diff changeset
   414
          |> (fn goal => Goal.prove_sorry raw_lthy [] [] goal tac)
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   415
          |> Thm.close_derivation
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   416
          |> Morphism.thm phi
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   417
          |> split_conj_thm
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   418
          |> map (fn thm => thm RS @{thm comp_eq_dest});
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   419
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   420
        val pre_map_defs = no_refl (map map_def_of_bnf bnfs);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   421
        val fp_pre_map_defs = no_refl (map map_def_of_bnf pre_bnfs);
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   422
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   423
        val unfold_map = map (unfold_thms lthy (id_apply :: pre_map_defs));
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   424
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   425
        val fp_xtor_co_iterss = of_fp_res #xtor_co_iter_thmss;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   426
        val fp_xtor_un_folds = map (mk_pointfree lthy o un_fold_of) fp_xtor_co_iterss;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   427
        val fp_xtor_co_recs = map (mk_pointfree lthy o co_rec_of) fp_xtor_co_iterss;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   428
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   429
        val fp_co_iter_o_mapss = of_fp_res #xtor_co_iter_o_map_thmss;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   430
        val fp_fold_o_maps = map un_fold_of fp_co_iter_o_mapss;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   431
        val fp_rec_o_maps = map co_rec_of fp_co_iter_o_mapss;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   432
        val fold_thms = fp_case fp @{thm comp_assoc} @{thm comp_assoc[symmetric]} ::
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   433
          map (fn thm => thm RS rewrite_comp_comp) @{thms map_pair.comp sum_map.comp} @
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   434
          @{thms id_apply comp_id id_comp map_pair.comp map_pair.id sum_map.comp sum_map.id};
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   435
        val rec_thms = fold_thms @ fp_case fp
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   436
          @{thms fst_convol map_pair_o_convol convol_o}
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 55394
diff changeset
   437
          @{thms case_sum_o_inj(1) case_sum_o_sum_map o_case_sum};
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   438
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   439
        val eq_thm_prop_untyped = op Term.aconv_untyped o pairself Thm.full_prop_of;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   440
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   441
        val map_thms = no_refl (maps (fn bnf =>
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   442
           let val map_comp0 = map_comp0_of_bnf bnf RS sym
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   443
           in [map_comp0, map_comp0 RS rewrite_comp_comp, map_id0_of_bnf bnf] end) fp_nesty_bnfs) @
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   444
          remove eq_thm_prop_untyped (fp_case fp @{thm comp_assoc[symmetric]} @{thm comp_assoc})
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   445
          (map2 (fn thm => fn bnf =>
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   446
            @{thm type_copy_map_comp0_undo} OF
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   447
              (replicate 3 thm @ unfold_map [map_comp0_of_bnf bnf]) RS
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   448
              rewrite_comp_comp)
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   449
          type_definitions bnfs);
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   450
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   451
        fun mk_Rep_o_Abs thm = thm RS @{thm type_copy_Rep_o_Abs} RS rewrite_comp_comp;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   452
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   453
        val fp_Rep_o_Abss = map mk_Rep_o_Abs fp_type_definitions;
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   454
        val Rep_o_Abss = map mk_Rep_o_Abs type_definitions;
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   455
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   456
        fun mk_tac defs o_map_thms xtor_thms thms {context = ctxt, prems = _} =
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   457
          unfold_thms_tac ctxt (flat [thms, defs, pre_map_defs, fp_pre_map_defs,
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   458
            xtor_thms, o_map_thms, map_thms, fp_Rep_o_Abss, Rep_o_Abss]) THEN
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   459
          CONJ_WRAP (K (HEADGOAL (rtac refl))) bnfs;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   460
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   461
        val fold_tac = mk_tac raw_un_fold_defs fp_fold_o_maps fp_xtor_un_folds fold_thms;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   462
        val rec_tac = mk_tac raw_co_rec_defs fp_rec_o_maps fp_xtor_co_recs rec_thms;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   463
      in
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   464
        (mk_thms fold_strs fold_goals fold_tac, mk_thms rec_strs rec_goals rec_tac)
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   465
      end;
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   466
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   467
    (* These results are half broken. This is deliberate. We care only about those fields that are
55531
601ca8efa000 renamed 'datatype_new_compat' to 'datatype_compat'
blanchet
parents: 55530
diff changeset
   468
       used by "primrec", "primcorecursive", and "datatype_compat". *)
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   469
    val fp_res =
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   470
      ({Ts = fpTs,
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   471
        bnfs = of_fp_res #bnfs,
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   472
        dtors = dtors,
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   473
        ctors = ctors,
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   474
        xtor_co_iterss = transpose [un_folds, co_recs],
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   475
        xtor_co_induct = xtor_co_induct_thm,
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   476
        dtor_ctors = of_fp_res #dtor_ctors (*too general types*),
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   477
        ctor_dtors = of_fp_res #ctor_dtors (*too general types*),
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   478
        ctor_injects = of_fp_res #ctor_injects (*too general types*),
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   479
        dtor_injects = of_fp_res #dtor_injects (*too general types*),
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   480
        xtor_map_thms = of_fp_res #xtor_map_thms (*too general types and terms*),
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   481
        xtor_set_thmss = of_fp_res #xtor_set_thmss (*too general types and terms*),
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   482
        xtor_rel_thms = of_fp_res #xtor_rel_thms (*too general types and terms*),
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   483
        xtor_co_iter_thmss = transpose [xtor_un_fold_thms, xtor_co_rec_thms],
55803
74d3fe9031d8 joint work with blanchet: intermediate typedef for the input to fp-operations
traytel
parents: 55575
diff changeset
   484
        xtor_co_iter_o_map_thmss = of_fp_res #xtor_co_iter_o_map_thmss
55539
0819931d652d simplified data structure by reducing the incidence of clumsy indices
blanchet
parents: 55531
diff changeset
   485
          (*theorem about old constant*),
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   486
        rel_xtor_co_induct_thm = rel_xtor_co_induct_thm}
54740
91f54d386680 maintain morphism names for diagnostic purposes;
wenzelm
parents: 54298
diff changeset
   487
       |> morph_fp_result (Morphism.term_morphism "BNF" (singleton (Variable.polymorphic lthy))));
53303
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   488
  in
ae49b835ca01 moved files related to "primrec_new", "primcorec", and "datatype_compat" from bitbucket co-rec repository
blanchet
parents:
diff changeset
   489
    (fp_res, lthy)
54242
blanchet
parents: 54233
diff changeset
   490
  end;
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
end;