| author | wenzelm | 
| Thu, 18 May 2017 14:14:20 +0200 | |
| changeset 65865 | 177b90f33f40 | 
| parent 63859 | dca6fabd8060 | 
| child 69593 | 3dda49e08b9d | 
| 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 | |
| 63045 | 31 | fun trivial_fp_result_of fp_bnf fpT C xtor_map xtor_sets xtor_rel ctor_rec_o_map | 
| 32 | xtor_rel_induct ctor_rec_transfer = | |
| 59819 | 33 | let | 
| 34 |     val xtors = [Const (@{const_name xtor}, fpT --> fpT)];
 | |
| 35 |     val co_recs = [Const (@{const_name ctor_rec}, (fpT --> C) --> (fpT --> C))];
 | |
| 36 |     val co_rec_thms = [map_id0_of_bnf fp_bnf RS @{thm ctor_rec}];
 | |
| 62863 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 traytel parents: 
62721diff
changeset | 37 |     val co_rec_unique_thm = map_id0_of_bnf fp_bnf RS @{thm ctor_rec_unique};
 | 
| 59819 | 38 | in | 
| 62684 
cb20e8828196
document that n2m does not depend on most things in fp_sugar in its type
 traytel parents: 
62335diff
changeset | 39 |     {Ts = [fpT], bnfs = [fp_bnf], pre_bnfs = [ID_bnf], absT_infos = [trivial_absT_info_of fpT],
 | 
| 62907 | 40 | ctors = xtors, dtors = xtors, xtor_un_folds = co_recs, | 
| 59819 | 41 |      xtor_co_recs = co_recs, xtor_co_induct = @{thm xtor_induct}, dtor_ctors = @{thms xtor_xtor},
 | 
| 42 |      ctor_dtors = @{thms xtor_xtor}, ctor_injects = @{thms xtor_inject},
 | |
| 62863 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 traytel parents: 
62721diff
changeset | 43 |      dtor_injects = @{thms xtor_inject}, xtor_maps = [xtor_map],
 | 
| 
e0b894bba6ff
single uniqueness theorems for map, (un)fold, (co)rec for mutual (co)datatypes
 traytel parents: 
62721diff
changeset | 44 |      xtor_map_unique = @{thm xtor_map_unique}, xtor_setss = [xtor_sets], xtor_rels = [xtor_rel],
 | 
| 62907 | 45 | xtor_un_fold_thms = co_rec_thms, xtor_co_rec_thms = co_rec_thms, | 
| 46 | xtor_un_fold_unique = co_rec_unique_thm, xtor_co_rec_unique = co_rec_unique_thm, | |
| 47 | xtor_un_fold_o_maps = [ctor_rec_o_map], xtor_co_rec_o_maps = [ctor_rec_o_map], | |
| 63045 | 48 | xtor_un_fold_transfers = [ctor_rec_transfer], | 
| 49 | xtor_co_rec_transfers = [ctor_rec_transfer], | |
| 62721 | 50 | xtor_rel_co_induct = xtor_rel_induct, dtor_set_inducts = []} | 
| 59819 | 51 | end; | 
| 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]};
 | 
| 63045 | 67 |     val ctor_rec_transfer = @{thm ctor_rec_transfer[of "rel_sum R1 R2" for R1 R2]};
 | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 68 | in | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 69 |     {T = fpT,
 | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 70 | BT = fpBT, | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 71 | X = X, | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 72 | fp = Least_FP, | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 73 | fp_res_index = 0, | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 74 | fp_res = | 
| 63045 | 75 | trivial_fp_result_of fp_bnf fpT C xtor_map xtor_sets xtor_rel ctor_rec_o_map xtor_rel_induct | 
| 76 | ctor_rec_transfer, | |
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 77 | pre_bnf = ID_bnf (*wrong*), | 
| 58674 | 78 | fp_bnf = fp_bnf, | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 79 | absT_info = trivial_absT_info_of fpT, | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 80 | fp_nesting_bnfs = [], | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 81 | live_nesting_bnfs = [], | 
| 58460 | 82 | fp_ctr_sugar = | 
| 83 |        {ctrXs_Tss = ctr_Tss,
 | |
| 62321 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 blanchet parents: 
61424diff
changeset | 84 |         ctor_iff_dtor = @{thm xtor_iff_xtor},
 | 
| 58460 | 85 |         ctr_defs = @{thms Inl_def_alt Inr_def_alt},
 | 
| 58569 | 86 | ctr_sugar = the_frozen_ctr_sugar_of ctxt fpT_name, | 
| 58916 | 87 |         ctr_transfers = @{thms Inl_transfer Inr_transfer},
 | 
| 88 |         case_transfers = @{thms case_sum_transfer},
 | |
| 89 |         disc_transfers = @{thms isl_transfer},
 | |
| 58676 | 90 | sel_transfers = []}, | 
| 58459 | 91 | fp_bnf_sugar = | 
| 58462 | 92 |        {map_thms = @{thms map_sum.simps},
 | 
| 58916 | 93 |         map_disc_iffs = @{thms isl_map_sum},
 | 
| 94 |         map_selss = map single @{thms map_sum_sel},
 | |
| 58462 | 95 |         rel_injects = @{thms rel_sum_simps(1,4)},
 | 
| 58562 | 96 |         rel_distincts = @{thms rel_sum_simps(2,3)[THEN eq_False[THEN iffD1]]},
 | 
| 58916 | 97 |         rel_sels = @{thms rel_sum_sel},
 | 
| 98 |         rel_intros = @{thms rel_sum.intros},
 | |
| 99 |         rel_cases = @{thms rel_sum.cases},
 | |
| 62335 | 100 |         pred_injects = @{thms pred_sum_inject},
 | 
| 58916 | 101 |         set_thms = @{thms sum_set_simps},
 | 
| 102 |         set_selssss = [[[@{thms set_sum_sel(1)}], [[]]], [[[]], [@{thms set_sum_sel(2)}]]],
 | |
| 103 |         set_introssss = [[[@{thms setl.intros[OF refl]}], [[]]],
 | |
| 104 |           [[[]], [@{thms setr.intros[OF refl]}]]],
 | |
| 63170 | 105 |         set_cases = @{thms setl.cases[simplified hypsubst_in_prems]
 | 
| 106 | setr.cases[simplified hypsubst_in_prems]}}, | |
| 63859 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 blanchet parents: 
63170diff
changeset | 107 | fp_co_induct_sugar = SOME | 
| 58461 | 108 |        {co_rec = Const (@{const_name case_sum}, map (fn Ts => (Ts ---> C)) ctr_Tss ---> fpT --> C),
 | 
| 109 |         common_co_inducts = @{thms sum.induct},
 | |
| 110 |         co_inducts = @{thms sum.induct},
 | |
| 111 |         co_rec_def = @{thm ctor_rec_def_alt[of "case_sum f1 f2" for f1 f2]},
 | |
| 58459 | 112 |         co_rec_thms = @{thms sum.case},
 | 
| 113 | co_rec_discs = [], | |
| 58572 | 114 | co_rec_disc_iffs = [], | 
| 58573 | 115 | co_rec_selss = [], | 
| 116 | co_rec_codes = [], | |
| 58916 | 117 |         co_rec_transfers = @{thms case_sum_transfer},
 | 
| 58734 | 118 |         co_rec_o_maps = @{thms case_sum_o_map_sum},
 | 
| 58916 | 119 |         common_rel_co_inducts = @{thms rel_sum.inducts},
 | 
| 120 |         rel_co_inducts = @{thms rel_sum.induct},
 | |
| 58577 | 121 | common_set_inducts = [], | 
| 122 | set_inducts = []}} | |
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 123 | end; | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 124 | |
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 125 | fun fp_sugar_of_prod ctxt = | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 126 | let | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 127 |     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 | 128 |     val fpBT = @{typ "'c * 'd"};
 | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 129 |     val C = @{typ 'e};
 | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 130 |     val X = @{typ 'prod};
 | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 131 | val ctr_Ts = As; | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 132 | |
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 133 | 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 | 134 |     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 | 135 |     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 | 136 |     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 | 137 |     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 | 138 |     val xtor_rel_induct = @{thm xtor_rel_induct[of "rel_prod R1 R2" for R1 R2]};
 | 
| 63045 | 139 |     val ctor_rec_transfer = @{thm ctor_rec_transfer[of "rel_prod R1 R2" for R1 R2]};
 | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 140 | in | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 141 |     {T = fpT,
 | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 142 | BT = fpBT, | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 143 | X = X, | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 144 | fp = Least_FP, | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 145 | fp_res_index = 0, | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 146 | fp_res = | 
| 63045 | 147 | trivial_fp_result_of fp_bnf fpT C xtor_map xtor_sets xtor_rel ctor_rec_o_map xtor_rel_induct | 
| 148 | ctor_rec_transfer, | |
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 149 | pre_bnf = ID_bnf (*wrong*), | 
| 58674 | 150 | fp_bnf = fp_bnf, | 
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 151 | absT_info = trivial_absT_info_of fpT, | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 152 | fp_nesting_bnfs = [], | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 153 | live_nesting_bnfs = [], | 
| 58460 | 154 | fp_ctr_sugar = | 
| 155 |        {ctrXs_Tss = [ctr_Ts],
 | |
| 62321 
1abe81758619
keep 'ctor_iff_dtor' theorem around in BNF FP database
 blanchet parents: 
61424diff
changeset | 156 |         ctor_iff_dtor = @{thm xtor_iff_xtor},
 | 
| 58460 | 157 |         ctr_defs = @{thms Pair_def_alt},
 | 
| 58569 | 158 | ctr_sugar = the_frozen_ctr_sugar_of ctxt fpT_name, | 
| 58916 | 159 |         ctr_transfers = @{thms Pair_transfer},
 | 
| 160 |         case_transfers = @{thms case_prod_transfer},
 | |
| 58676 | 161 | disc_transfers = [], | 
| 58916 | 162 |         sel_transfers = @{thms fst_transfer snd_transfer}},
 | 
| 58459 | 163 | fp_bnf_sugar = | 
| 58462 | 164 |        {map_thms = @{thms map_prod_simp},
 | 
| 58560 | 165 | map_disc_iffs = [], | 
| 58916 | 166 |         map_selss = [@{thms fst_map_prod snd_map_prod}],
 | 
| 62335 | 167 |         rel_injects = @{thms rel_prod_inject},
 | 
| 58562 | 168 | rel_distincts = [], | 
| 58916 | 169 |         rel_sels = @{thms rel_prod_sel},
 | 
| 170 |         rel_intros = @{thms rel_prod.intros},
 | |
| 171 |         rel_cases = @{thms rel_prod.cases},
 | |
| 62335 | 172 |         pred_injects = @{thms pred_prod_inject},
 | 
| 58916 | 173 |         set_thms = @{thms prod_set_simps},
 | 
| 174 |         set_selssss = [[[@{thms fsts.intros}, []]], [[[], @{thms snds.intros}]]],
 | |
| 63170 | 175 |         set_introssss = [[[@{thms fsts.intros[of "(a, b)" for a b, simplified fst_conv]}, []]],
 | 
| 176 |           [[[], @{thms snds.intros[of "(a, b)" for a b, simplified snd_conv]}]]],
 | |
| 177 |         set_cases = @{thms fsts.cases[simplified eq_fst_iff ex_neg_all_pos]
 | |
| 178 | snds.cases[simplified eq_snd_iff ex_neg_all_pos]}}, | |
| 63859 
dca6fabd8060
make (co)induct component of 'fp_sugar' optional, for the benefit of nonuniform (co)datatypes and other similar extensions
 blanchet parents: 
63170diff
changeset | 179 | fp_co_induct_sugar = SOME | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61125diff
changeset | 180 |        {co_rec = Const (@{const_name case_prod}, (ctr_Ts ---> C) --> fpT --> C),
 | 
| 58461 | 181 |         common_co_inducts = @{thms prod.induct},
 | 
| 182 |         co_inducts = @{thms prod.induct},
 | |
| 183 |         co_rec_def = @{thm ctor_rec_def_alt[of "case_prod f" for f]},
 | |
| 58459 | 184 |         co_rec_thms = @{thms prod.case},
 | 
| 185 | co_rec_discs = [], | |
| 58572 | 186 | co_rec_disc_iffs = [], | 
| 58573 | 187 | co_rec_selss = [], | 
| 188 | co_rec_codes = [], | |
| 58916 | 189 |         co_rec_transfers = @{thms case_prod_transfer},
 | 
| 58734 | 190 |         co_rec_o_maps = @{thms case_prod_o_map_prod},
 | 
| 58916 | 191 |         common_rel_co_inducts = @{thms rel_prod.inducts},
 | 
| 192 |         rel_co_inducts = @{thms rel_prod.induct},
 | |
| 58577 | 193 | common_set_inducts = [], | 
| 194 | set_inducts = []}} | |
| 58352 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 195 | end; | 
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 196 | |
| 58665 | 197 | val _ = Theory.setup (Named_Target.theory_map (fn lthy => | 
| 58377 
c6f93b8d2d8e
moved old 'size' generator together with 'old_datatype'
 blanchet parents: 
58353diff
changeset | 198 | 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 | 199 | [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 | 200 | |
| 
37745650a3f4
register 'prod' and 'sum' as datatypes, to allow N2M through them
 blanchet parents: diff
changeset | 201 | end; |