| author | eberlm | 
| Tue, 05 Jan 2016 17:54:21 +0100 | |
| changeset 62066 | 4db2d39aa76c | 
| parent 61424 | c3658c18b7bc | 
| child 62321 | 1abe81758619 | 
| permissions | -rw-r--r-- | 
| 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 | 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 = | 
| 59819 | 20 |   {absT = fpT, repT = fpT, abs = Const (@{const_name id_bnf}, fpT --> fpT),
 | 
| 58353 | 21 |    rep = Const (@{const_name id_bnf}, fpT --> fpT),
 | 
| 22 |    abs_inject = @{thm type_definition.Abs_inject[OF type_definition_id_bnf_UNIV UNIV_I UNIV_I]},
 | |
| 23 |    abs_inverse = @{thm type_definition.Abs_inverse[OF type_definition_id_bnf_UNIV UNIV_I]},
 | |
| 24 |    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 | 25 | |
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 26 | 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 | 27 | the (ctr_sugar_of ctxt fpT_name) | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 28 | |> 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 | 29 | $> 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 | 30 | |
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 31 | fun trivial_fp_result_of fp_bnf fpT C xtor_map xtor_sets xtor_rel ctor_rec_o_map xtor_rel_induct = | 
| 59819 | 32 | let | 
| 33 |     val xtors = [Const (@{const_name xtor}, fpT --> fpT)];
 | |
| 34 |     val co_recs = [Const (@{const_name ctor_rec}, (fpT --> C) --> (fpT --> C))];
 | |
| 35 |     val co_rec_thms = [map_id0_of_bnf fp_bnf RS @{thm ctor_rec}];
 | |
| 36 | in | |
| 37 |     {Ts = [fpT], bnfs = [fp_bnf], ctors = xtors, dtors = xtors, xtor_un_folds = co_recs,
 | |
| 38 |      xtor_co_recs = co_recs, xtor_co_induct = @{thm xtor_induct}, dtor_ctors = @{thms xtor_xtor},
 | |
| 39 |      ctor_dtors = @{thms xtor_xtor}, ctor_injects = @{thms xtor_inject},
 | |
| 59856 
ed0ca9029021
export more low-level theorems in data structure (partly for 'corec')
 blanchet parents: 
59819diff
changeset | 40 |      dtor_injects = @{thms xtor_inject}, xtor_maps = [xtor_map], xtor_map_uniques = [],
 | 
| 
ed0ca9029021
export more low-level theorems in data structure (partly for 'corec')
 blanchet parents: 
59819diff
changeset | 41 | xtor_setss = [xtor_sets], xtor_rels = [xtor_rel], xtor_un_fold_thms = co_rec_thms, | 
| 
ed0ca9029021
export more low-level theorems in data structure (partly for 'corec')
 blanchet parents: 
59819diff
changeset | 42 | xtor_co_rec_thms = co_rec_thms, xtor_un_fold_uniques = [], xtor_co_rec_uniques = [], | 
| 
ed0ca9029021
export more low-level theorems in data structure (partly for 'corec')
 blanchet parents: 
59819diff
changeset | 43 | xtor_un_fold_o_maps = [ctor_rec_o_map], xtor_co_rec_o_maps = [ctor_rec_o_map], | 
| 
ed0ca9029021
export more low-level theorems in data structure (partly for 'corec')
 blanchet parents: 
59819diff
changeset | 44 | xtor_un_fold_transfers = [], xtor_co_rec_transfers = [], xtor_rel_co_induct = xtor_rel_induct, | 
| 
ed0ca9029021
export more low-level theorems in data structure (partly for 'corec')
 blanchet parents: 
59819diff
changeset | 45 | dtor_set_inducts = []} | 
| 59819 | 46 | end; | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 47 | |
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 48 | fun fp_sugar_of_sum ctxt = | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 49 | let | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 50 |     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 | 51 |     val fpBT = @{typ "'c + 'd"};
 | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 52 |     val C = @{typ 'e};
 | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 53 |     val X = @{typ 'sum};
 | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 54 | val ctr_Tss = map single As; | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 55 | |
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 56 | 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 | 57 |     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 | 58 |     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 | 59 |     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 | 60 |     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 | 61 |     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 | 62 | in | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 63 |     {T = fpT,
 | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 64 | BT = fpBT, | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 65 | X = X, | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 66 | fp = Least_FP, | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 67 | fp_res_index = 0, | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 68 | fp_res = | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 69 | 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 | 70 | pre_bnf = ID_bnf (*wrong*), | 
| 58674 | 71 | fp_bnf = fp_bnf, | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 72 | absT_info = trivial_absT_info_of fpT, | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 73 | fp_nesting_bnfs = [], | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 74 | live_nesting_bnfs = [], | 
| 58460 | 75 | fp_ctr_sugar = | 
| 76 |        {ctrXs_Tss = ctr_Tss,
 | |
| 77 |         ctr_defs = @{thms Inl_def_alt Inr_def_alt},
 | |
| 58569 | 78 | ctr_sugar = the_frozen_ctr_sugar_of ctxt fpT_name, | 
| 58916 | 79 |         ctr_transfers = @{thms Inl_transfer Inr_transfer},
 | 
| 80 |         case_transfers = @{thms case_sum_transfer},
 | |
| 81 |         disc_transfers = @{thms isl_transfer},
 | |
| 58676 | 82 | sel_transfers = []}, | 
| 58459 | 83 | fp_bnf_sugar = | 
| 58462 | 84 |        {map_thms = @{thms map_sum.simps},
 | 
| 58916 | 85 |         map_disc_iffs = @{thms isl_map_sum},
 | 
| 86 |         map_selss = map single @{thms map_sum_sel},
 | |
| 58462 | 87 |         rel_injects = @{thms rel_sum_simps(1,4)},
 | 
| 58562 | 88 |         rel_distincts = @{thms rel_sum_simps(2,3)[THEN eq_False[THEN iffD1]]},
 | 
| 58916 | 89 |         rel_sels = @{thms rel_sum_sel},
 | 
| 90 |         rel_intros = @{thms rel_sum.intros},
 | |
| 91 |         rel_cases = @{thms rel_sum.cases},
 | |
| 92 |         set_thms = @{thms sum_set_simps},
 | |
| 93 |         set_selssss = [[[@{thms set_sum_sel(1)}], [[]]], [[[]], [@{thms set_sum_sel(2)}]]],
 | |
| 94 |         set_introssss = [[[@{thms setl.intros[OF refl]}], [[]]],
 | |
| 95 |           [[[]], [@{thms setr.intros[OF refl]}]]],
 | |
| 96 |         set_cases = @{thms setl.cases[unfolded hypsubst_in_prems]
 | |
| 97 | setr.cases[unfolded hypsubst_in_prems]}}, | |
| 58459 | 98 | fp_co_induct_sugar = | 
| 58461 | 99 |        {co_rec = Const (@{const_name case_sum}, map (fn Ts => (Ts ---> C)) ctr_Tss ---> fpT --> C),
 | 
| 100 |         common_co_inducts = @{thms sum.induct},
 | |
| 101 |         co_inducts = @{thms sum.induct},
 | |
| 102 |         co_rec_def = @{thm ctor_rec_def_alt[of "case_sum f1 f2" for f1 f2]},
 | |
| 58459 | 103 |         co_rec_thms = @{thms sum.case},
 | 
| 104 | co_rec_discs = [], | |
| 58572 | 105 | co_rec_disc_iffs = [], | 
| 58573 | 106 | co_rec_selss = [], | 
| 107 | co_rec_codes = [], | |
| 58916 | 108 |         co_rec_transfers = @{thms case_sum_transfer},
 | 
| 58734 | 109 |         co_rec_o_maps = @{thms case_sum_o_map_sum},
 | 
| 58916 | 110 |         common_rel_co_inducts = @{thms rel_sum.inducts},
 | 
| 111 |         rel_co_inducts = @{thms rel_sum.induct},
 | |
| 58577 | 112 | common_set_inducts = [], | 
| 113 | set_inducts = []}} | |
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 114 | end; | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 115 | |
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 116 | fun fp_sugar_of_prod ctxt = | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 117 | let | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 118 |     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 | 119 |     val fpBT = @{typ "'c * 'd"};
 | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 120 |     val C = @{typ 'e};
 | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 121 |     val X = @{typ 'prod};
 | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 122 | val ctr_Ts = As; | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 123 | |
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 124 | 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 | 125 |     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 | 126 |     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 | 127 |     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 | 128 |     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 | 129 |     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 | 130 | in | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 131 |     {T = fpT,
 | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 132 | BT = fpBT, | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 133 | X = X, | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 134 | fp = Least_FP, | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 135 | fp_res_index = 0, | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 136 | fp_res = | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 137 | 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 | 138 | pre_bnf = ID_bnf (*wrong*), | 
| 58674 | 139 | fp_bnf = fp_bnf, | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 140 | absT_info = trivial_absT_info_of fpT, | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 141 | fp_nesting_bnfs = [], | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 142 | live_nesting_bnfs = [], | 
| 58460 | 143 | fp_ctr_sugar = | 
| 144 |        {ctrXs_Tss = [ctr_Ts],
 | |
| 145 |         ctr_defs = @{thms Pair_def_alt},
 | |
| 58569 | 146 | ctr_sugar = the_frozen_ctr_sugar_of ctxt fpT_name, | 
| 58916 | 147 |         ctr_transfers = @{thms Pair_transfer},
 | 
| 148 |         case_transfers = @{thms case_prod_transfer},
 | |
| 58676 | 149 | disc_transfers = [], | 
| 58916 | 150 |         sel_transfers = @{thms fst_transfer snd_transfer}},
 | 
| 58459 | 151 | fp_bnf_sugar = | 
| 58462 | 152 |        {map_thms = @{thms map_prod_simp},
 | 
| 58560 | 153 | map_disc_iffs = [], | 
| 58916 | 154 |         map_selss = [@{thms fst_map_prod snd_map_prod}],
 | 
| 58462 | 155 |         rel_injects = @{thms rel_prod_apply},
 | 
| 58562 | 156 | rel_distincts = [], | 
| 58916 | 157 |         rel_sels = @{thms rel_prod_sel},
 | 
| 158 |         rel_intros = @{thms rel_prod.intros},
 | |
| 159 |         rel_cases = @{thms rel_prod.cases},
 | |
| 160 |         set_thms = @{thms prod_set_simps},
 | |
| 161 |         set_selssss = [[[@{thms fsts.intros}, []]], [[[], @{thms snds.intros}]]],
 | |
| 162 |         set_introssss = [[[@{thms fsts.intros[of "(a, b)" for a b, unfolded fst_conv]}, []]],
 | |
| 163 |           [[[], @{thms snds.intros[of "(a, b)" for a b, unfolded snd_conv]}]]],
 | |
| 164 |         set_cases = @{thms fsts.cases[unfolded eq_fst_iff ex_neg_all_pos]
 | |
| 165 | snds.cases[unfolded eq_snd_iff ex_neg_all_pos]}}, | |
| 58459 | 166 | fp_co_induct_sugar = | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61125diff
changeset | 167 |        {co_rec = Const (@{const_name case_prod}, (ctr_Ts ---> C) --> fpT --> C),
 | 
| 58461 | 168 |         common_co_inducts = @{thms prod.induct},
 | 
| 169 |         co_inducts = @{thms prod.induct},
 | |
| 170 |         co_rec_def = @{thm ctor_rec_def_alt[of "case_prod f" for f]},
 | |
| 58459 | 171 |         co_rec_thms = @{thms prod.case},
 | 
| 172 | co_rec_discs = [], | |
| 58572 | 173 | co_rec_disc_iffs = [], | 
| 58573 | 174 | co_rec_selss = [], | 
| 175 | co_rec_codes = [], | |
| 58916 | 176 |         co_rec_transfers = @{thms case_prod_transfer},
 | 
| 58734 | 177 |         co_rec_o_maps = @{thms case_prod_o_map_prod},
 | 
| 58916 | 178 |         common_rel_co_inducts = @{thms rel_prod.inducts},
 | 
| 179 |         rel_co_inducts = @{thms rel_prod.induct},
 | |
| 58577 | 180 | common_set_inducts = [], | 
| 181 | set_inducts = []}} | |
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 182 | end; | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 183 | |
| 58665 | 184 | val _ = Theory.setup (Named_Target.theory_map (fn lthy => | 
| 58377 
c6f93b8d2d8e
moved old 'size' generator together with 'old_datatype'
 blanchet parents: 
58353diff
changeset | 185 | 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 | 186 | [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 | 187 | |
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 188 | end; |