src/HOL/Tools/BNF/bnf_lfp_basic_sugar.ML
author desharna
Mon, 06 Oct 2014 13:33:45 +0200
changeset 58566 62aa83edad7e
parent 58565 97cefa5ef0be
child 58567 f0d09e17edba
permissions -rw-r--r--
add 'set_sels' to 'fp_sugar'
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
58352
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/Tools/BNF/bnf_lfp_basic_sugar.ML
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
     3
    Copyright   2014
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
     4
58353
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
     5
Registration of basic types as BNF least fixpoints (datatypes).
58352
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
     6
*)
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
     7
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
     8
structure BNF_LFP_Basic_Sugar : sig end =
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
     9
struct
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    10
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    11
open Ctr_Sugar
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    12
open BNF_Util
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    13
open BNF_Def
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    14
open BNF_Comp
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    15
open BNF_FP_Rec_Sugar_Util
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    16
open BNF_FP_Util
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    17
open BNF_FP_Def_Sugar
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    18
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    19
fun trivial_absT_info_of fpT =
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    20
  {absT = fpT,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    21
   repT = fpT,
58353
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
    22
   abs = Const (@{const_name id_bnf}, fpT --> fpT),
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
    23
   rep = Const (@{const_name id_bnf}, fpT --> fpT),
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
    24
   abs_inject = @{thm type_definition.Abs_inject[OF type_definition_id_bnf_UNIV UNIV_I UNIV_I]},
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
    25
   abs_inverse = @{thm type_definition.Abs_inverse[OF type_definition_id_bnf_UNIV UNIV_I]},
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
    26
   type_definition = @{thm type_definition_id_bnf_UNIV}};
58352
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    27
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    28
fun the_frozen_ctr_sugar_of ctxt fpT_name =
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    29
  the (ctr_sugar_of ctxt fpT_name)
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    30
  |> morph_ctr_sugar (Morphism.typ_morphism "BNF" Logic.unvarifyT_global
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    31
    $> Morphism.term_morphism "BNF" (Term.map_types Logic.unvarifyT_global));
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    32
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    33
fun trivial_fp_result_of fp_bnf fpT C xtor_map xtor_sets xtor_rel ctor_rec_o_map xtor_rel_induct =
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    34
  {Ts = [fpT],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    35
   bnfs = [fp_bnf],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    36
   ctors = [Const (@{const_name xtor}, fpT --> fpT)],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    37
   dtors = [Const (@{const_name xtor}, fpT --> fpT)],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    38
   xtor_co_recs = [Const (@{const_name ctor_rec}, (fpT --> C) --> (fpT --> C))],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    39
   xtor_co_induct = @{thm xtor_induct},
58377
c6f93b8d2d8e moved old 'size' generator together with 'old_datatype'
blanchet
parents: 58353
diff changeset
    40
   dtor_ctors = @{thms xtor_xtor},
c6f93b8d2d8e moved old 'size' generator together with 'old_datatype'
blanchet
parents: 58353
diff changeset
    41
   ctor_dtors = @{thms xtor_xtor},
c6f93b8d2d8e moved old 'size' generator together with 'old_datatype'
blanchet
parents: 58353
diff changeset
    42
   ctor_injects = @{thms xtor_inject},
c6f93b8d2d8e moved old 'size' generator together with 'old_datatype'
blanchet
parents: 58353
diff changeset
    43
   dtor_injects = @{thms xtor_inject},
58352
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    44
   xtor_map_thms = [xtor_map],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    45
   xtor_set_thmss = [xtor_sets],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    46
   xtor_rel_thms = [xtor_rel],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    47
   xtor_co_rec_thms = [map_id0_of_bnf fp_bnf RS @{thm ctor_rec}],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    48
   xtor_co_rec_o_map_thms = [ctor_rec_o_map],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    49
   rel_xtor_co_induct_thm = xtor_rel_induct,
58446
e89f57d1e46c generate 'rec_transfer' for datatypes
desharna
parents: 58377
diff changeset
    50
   dtor_set_induct_thms = [],
58448
a1d4e7473c98 generate 'corec_transfer' for codatatypes
desharna
parents: 58446
diff changeset
    51
   xtor_co_rec_transfer_thms = []};
58352
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    52
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    53
fun fp_sugar_of_sum ctxt =
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    54
  let
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    55
    val fpT as Type (fpT_name, As) = @{typ "'a + 'b"};
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    56
    val fpBT = @{typ "'c + 'd"};
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    57
    val C = @{typ 'e};
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    58
    val X = @{typ 'sum};
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    59
    val ctr_Tss = map single As;
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    60
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    61
    val fp_bnf = the (bnf_of ctxt fpT_name);
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    62
    val xtor_map = @{thm xtor_map[of "map_sum f1 f2" for f1 f2]};
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    63
    val xtor_sets = @{thms xtor_set[of setl] xtor_set[of setr]};
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    64
    val xtor_rel = @{thm xtor_rel[of "rel_sum R1 R2" for R1 R2]};
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    65
    val ctor_rec_o_map = @{thm ctor_rec_o_map[of _ "map_sum g1 g2" for g1 g2]};
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    66
    val xtor_rel_induct = @{thm xtor_rel_induct[of "rel_sum R1 R2" for R1 R2]};
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    67
  in
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    68
    {T = fpT,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    69
     BT = fpBT,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    70
     X = X,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    71
     fp = Least_FP,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    72
     fp_res_index = 0,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    73
     fp_res =
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    74
       trivial_fp_result_of fp_bnf fpT C xtor_map xtor_sets xtor_rel ctor_rec_o_map xtor_rel_induct,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    75
     pre_bnf = ID_bnf (*wrong*),
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    76
     absT_info = trivial_absT_info_of fpT,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    77
     fp_nesting_bnfs = [],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    78
     live_nesting_bnfs = [],
58460
a88eb33058f7 refactor fp_sugar move theorems
desharna
parents: 58459
diff changeset
    79
     fp_ctr_sugar =
a88eb33058f7 refactor fp_sugar move theorems
desharna
parents: 58459
diff changeset
    80
       {ctrXs_Tss = ctr_Tss,
a88eb33058f7 refactor fp_sugar move theorems
desharna
parents: 58459
diff changeset
    81
        ctr_defs = @{thms Inl_def_alt Inr_def_alt},
a88eb33058f7 refactor fp_sugar move theorems
desharna
parents: 58459
diff changeset
    82
        ctr_sugar = the_frozen_ctr_sugar_of ctxt fpT_name},
58459
f70bffabd7cf refactor fp_sugar move theorems
desharna
parents: 58458
diff changeset
    83
     fp_bnf_sugar =
58462
b46874f2090f refactor fp_sugar move theorems
desharna
parents: 58461
diff changeset
    84
       {map_thms = @{thms map_sum.simps},
58560
ee502a9b38aa add 'map_disc_iffs' to 'fp_sugar'
desharna
parents: 58462
diff changeset
    85
        map_disc_iffs = [],
58561
7d7473b54fe0 add 'map_sels' to 'fp_sugar'
desharna
parents: 58560
diff changeset
    86
        map_sels = [],
58462
b46874f2090f refactor fp_sugar move theorems
desharna
parents: 58461
diff changeset
    87
        rel_injects = @{thms rel_sum_simps(1,4)},
58562
e94cd4f71d0c add 'rel_sels' to 'fp_sugar'
desharna
parents: 58561
diff changeset
    88
        rel_distincts = @{thms rel_sum_simps(2,3)[THEN eq_False[THEN iffD1]]},
58563
f5019700efa5 add 'rel_intros' to 'fp_sugar'
desharna
parents: 58562
diff changeset
    89
        rel_sels = [],
58564
778a80674112 add 'rel_cases' to 'fp_sugar'
desharna
parents: 58563
diff changeset
    90
        rel_intros = [],
58565
97cefa5ef0be add 'set_thms' to 'fp_sugar'
desharna
parents: 58564
diff changeset
    91
        rel_cases = [],
58566
62aa83edad7e add 'set_sels' to 'fp_sugar'
desharna
parents: 58565
diff changeset
    92
        set_thms = [],
62aa83edad7e add 'set_sels' to 'fp_sugar'
desharna
parents: 58565
diff changeset
    93
        set_sels = []},
58459
f70bffabd7cf refactor fp_sugar move theorems
desharna
parents: 58458
diff changeset
    94
     fp_co_induct_sugar =
58461
75ee8d49c724 refactor fp_sugar move theorems
desharna
parents: 58460
diff changeset
    95
       {co_rec = Const (@{const_name case_sum}, map (fn Ts => (Ts ---> C)) ctr_Tss ---> fpT --> C),
75ee8d49c724 refactor fp_sugar move theorems
desharna
parents: 58460
diff changeset
    96
        common_co_inducts = @{thms sum.induct},
75ee8d49c724 refactor fp_sugar move theorems
desharna
parents: 58460
diff changeset
    97
        co_inducts = @{thms sum.induct},
75ee8d49c724 refactor fp_sugar move theorems
desharna
parents: 58460
diff changeset
    98
        co_rec_def = @{thm ctor_rec_def_alt[of "case_sum f1 f2" for f1 f2]},
58459
f70bffabd7cf refactor fp_sugar move theorems
desharna
parents: 58458
diff changeset
    99
        co_rec_thms = @{thms sum.case},
f70bffabd7cf refactor fp_sugar move theorems
desharna
parents: 58458
diff changeset
   100
        co_rec_discs = [],
f70bffabd7cf refactor fp_sugar move theorems
desharna
parents: 58458
diff changeset
   101
        co_rec_selss = []}}
58352
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   102
  end;
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   103
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   104
fun fp_sugar_of_prod ctxt =
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   105
  let
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   106
    val fpT as Type (fpT_name, As) = @{typ "'a * 'b"};
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   107
    val fpBT = @{typ "'c * 'd"};
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   108
    val C = @{typ 'e};
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   109
    val X = @{typ 'prod};
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   110
    val ctr_Ts = As;
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   111
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   112
    val fp_bnf = the (bnf_of ctxt fpT_name);
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   113
    val xtor_map = @{thm xtor_map[of "map_prod f1 f2" for f1 f2]};
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   114
    val xtor_sets = @{thms xtor_set[of fsts] xtor_set[of snds]};
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   115
    val xtor_rel = @{thm xtor_rel[of "rel_prod R1 R2" for R1 R2]};
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   116
    val ctor_rec_o_map = @{thm ctor_rec_o_map[of _ "map_prod g1 g2" for g1 g2]};
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   117
    val xtor_rel_induct = @{thm xtor_rel_induct[of "rel_prod R1 R2" for R1 R2]};
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   118
  in
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   119
    {T = fpT,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   120
     BT = fpBT,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   121
     X = X,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   122
     fp = Least_FP,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   123
     fp_res_index = 0,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   124
     fp_res =
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   125
       trivial_fp_result_of fp_bnf fpT C xtor_map xtor_sets xtor_rel ctor_rec_o_map xtor_rel_induct,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   126
     pre_bnf = ID_bnf (*wrong*),
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   127
     absT_info = trivial_absT_info_of fpT,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   128
     fp_nesting_bnfs = [],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   129
     live_nesting_bnfs = [],
58460
a88eb33058f7 refactor fp_sugar move theorems
desharna
parents: 58459
diff changeset
   130
     fp_ctr_sugar =
a88eb33058f7 refactor fp_sugar move theorems
desharna
parents: 58459
diff changeset
   131
       {ctrXs_Tss = [ctr_Ts],
a88eb33058f7 refactor fp_sugar move theorems
desharna
parents: 58459
diff changeset
   132
        ctr_defs = @{thms Pair_def_alt},
a88eb33058f7 refactor fp_sugar move theorems
desharna
parents: 58459
diff changeset
   133
        ctr_sugar = the_frozen_ctr_sugar_of ctxt fpT_name},
58459
f70bffabd7cf refactor fp_sugar move theorems
desharna
parents: 58458
diff changeset
   134
     fp_bnf_sugar =
58462
b46874f2090f refactor fp_sugar move theorems
desharna
parents: 58461
diff changeset
   135
       {map_thms = @{thms map_prod_simp},
58560
ee502a9b38aa add 'map_disc_iffs' to 'fp_sugar'
desharna
parents: 58462
diff changeset
   136
        map_disc_iffs = [],
58561
7d7473b54fe0 add 'map_sels' to 'fp_sugar'
desharna
parents: 58560
diff changeset
   137
        map_sels = [],
58462
b46874f2090f refactor fp_sugar move theorems
desharna
parents: 58461
diff changeset
   138
        rel_injects = @{thms rel_prod_apply},
58562
e94cd4f71d0c add 'rel_sels' to 'fp_sugar'
desharna
parents: 58561
diff changeset
   139
        rel_distincts = [],
58563
f5019700efa5 add 'rel_intros' to 'fp_sugar'
desharna
parents: 58562
diff changeset
   140
        rel_sels = [],
58564
778a80674112 add 'rel_cases' to 'fp_sugar'
desharna
parents: 58563
diff changeset
   141
        rel_intros = [],
58565
97cefa5ef0be add 'set_thms' to 'fp_sugar'
desharna
parents: 58564
diff changeset
   142
        rel_cases = [],
58566
62aa83edad7e add 'set_sels' to 'fp_sugar'
desharna
parents: 58565
diff changeset
   143
        set_thms = [],
62aa83edad7e add 'set_sels' to 'fp_sugar'
desharna
parents: 58565
diff changeset
   144
        set_sels = []},
58459
f70bffabd7cf refactor fp_sugar move theorems
desharna
parents: 58458
diff changeset
   145
     fp_co_induct_sugar =
58461
75ee8d49c724 refactor fp_sugar move theorems
desharna
parents: 58460
diff changeset
   146
       {co_rec = Const (@{const_name case_prod}, (ctr_Ts ---> C) --> fpT --> C),
75ee8d49c724 refactor fp_sugar move theorems
desharna
parents: 58460
diff changeset
   147
        common_co_inducts = @{thms prod.induct},
75ee8d49c724 refactor fp_sugar move theorems
desharna
parents: 58460
diff changeset
   148
        co_inducts = @{thms prod.induct},
75ee8d49c724 refactor fp_sugar move theorems
desharna
parents: 58460
diff changeset
   149
        co_rec_def = @{thm ctor_rec_def_alt[of "case_prod f" for f]},
58459
f70bffabd7cf refactor fp_sugar move theorems
desharna
parents: 58458
diff changeset
   150
        co_rec_thms = @{thms prod.case},
f70bffabd7cf refactor fp_sugar move theorems
desharna
parents: 58458
diff changeset
   151
        co_rec_discs = [],
f70bffabd7cf refactor fp_sugar move theorems
desharna
parents: 58458
diff changeset
   152
        co_rec_selss = []}}
58352
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   153
  end;
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   154
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   155
val _ = Theory.setup (map_local_theory (fn lthy =>
58377
c6f93b8d2d8e moved old 'size' generator together with 'old_datatype'
blanchet
parents: 58353
diff changeset
   156
  fold (BNF_FP_Def_Sugar.register_fp_sugars (K true) o single o (fn f => f lthy))
58352
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   157
    [fp_sugar_of_sum, fp_sugar_of_prod] lthy));
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   158
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   159
end;