src/HOL/Tools/BNF/bnf_lfp_basic_sugar.ML
author blanchet
Tue, 16 Sep 2014 19:23:37 +0200
changeset 58353 c9f374b64d99
parent 58352 37745650a3f4
child 58377 c6f93b8d2d8e
permissions -rw-r--r--
tuned fact visibility
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
open BNF_LFP_Size
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    19
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    20
fun trivial_absT_info_of fpT =
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    21
  {absT = fpT,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    22
   repT = fpT,
58353
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
    23
   abs = Const (@{const_name id_bnf}, fpT --> fpT),
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
    24
   rep = Const (@{const_name id_bnf}, fpT --> fpT),
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
    25
   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
    26
   abs_inverse = @{thm type_definition.Abs_inverse[OF type_definition_id_bnf_UNIV UNIV_I]},
c9f374b64d99 tuned fact visibility
blanchet
parents: 58352
diff changeset
    27
   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
    28
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    29
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
    30
  the (ctr_sugar_of ctxt fpT_name)
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    31
  |> 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
    32
    $> 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
    33
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    34
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
    35
  {Ts = [fpT],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    36
   bnfs = [fp_bnf],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    37
   ctors = [Const (@{const_name xtor}, fpT --> fpT)],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    38
   dtors = [Const (@{const_name xtor}, fpT --> fpT)],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    39
   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
    40
   xtor_co_induct = @{thm xtor_induct},
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    41
   dtor_ctors = [@{thm xtor_xtor}],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    42
   ctor_dtors = [@{thm xtor_xtor}],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    43
   ctor_injects = [@{thm xtor_inject}],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    44
   dtor_injects = [@{thm xtor_inject}],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    45
   xtor_map_thms = [xtor_map],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    46
   xtor_set_thmss = [xtor_sets],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    47
   xtor_rel_thms = [xtor_rel],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    48
   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
    49
   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
    50
   rel_xtor_co_induct_thm = xtor_rel_induct,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    51
   dtor_set_induct_thms = []};
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 = [],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    79
     ctrXs_Tss = ctr_Tss,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    80
     ctr_defs = @{thms Inl_def_alt Inr_def_alt},
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    81
     ctr_sugar = the_frozen_ctr_sugar_of ctxt fpT_name,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    82
     co_rec = Const (@{const_name case_sum}, map (fn Ts => (Ts ---> C)) ctr_Tss ---> fpT --> C),
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    83
     co_rec_def = @{thm case_sum_def},
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    84
     maps = @{thms map_sum.simps},
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    85
     common_co_inducts = [@{thm sum.induct}],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    86
     co_inducts = [@{thm sum.induct}],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    87
     co_rec_thms = @{thms sum.case},
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    88
     co_rec_discs = [],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    89
     co_rec_selss = [],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    90
     rel_injects = @{thms rel_sum_simps(1,4)},
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    91
     rel_distincts = @{thms rel_sum_simps(2,3)[THEN eq_False[THEN iffD1]]}}
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    92
  end;
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    93
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    94
fun fp_sugar_of_prod ctxt =
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    95
  let
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    96
    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
    97
    val fpBT = @{typ "'c * 'd"};
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    98
    val C = @{typ 'e};
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
    99
    val X = @{typ 'prod};
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   100
    val ctr_Ts = As;
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   101
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   102
    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
   103
    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
   104
    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
   105
    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
   106
    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
   107
    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
   108
  in
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   109
    {T = fpT,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   110
     BT = fpBT,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   111
     X = X,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   112
     fp = Least_FP,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   113
     fp_res_index = 0,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   114
     fp_res =
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   115
       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
   116
     pre_bnf = ID_bnf (*wrong*),
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   117
     absT_info = trivial_absT_info_of fpT,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   118
     fp_nesting_bnfs = [],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   119
     live_nesting_bnfs = [],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   120
     ctrXs_Tss = [ctr_Ts],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   121
     ctr_defs = [@{thm Pair_def_alt}],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   122
     ctr_sugar = the_frozen_ctr_sugar_of ctxt fpT_name,
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   123
     co_rec = Const (@{const_name case_prod}, (ctr_Ts ---> C) --> fpT --> C),
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   124
     co_rec_def = @{thm case_prod_def},
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   125
     maps = [@{thm map_prod_simp}],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   126
     common_co_inducts = [@{thm prod.induct}],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   127
     co_inducts = [@{thm prod.induct}],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   128
     co_rec_thms = [@{thm prod.case}],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   129
     co_rec_discs = [],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   130
     co_rec_selss = [],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   131
     rel_injects = [@{thm rel_prod_apply}],
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   132
     rel_distincts = []}
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   133
  end;
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   134
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   135
val _ = Theory.setup (map_local_theory (fn lthy =>
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   136
  fold (BNF_FP_Def_Sugar.register_fp_sugars (fn s => s <> size_plugin) o single o (fn f => f lthy))
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   137
    [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
   138
37745650a3f4 register 'prod' and 'sum' as datatypes, to allow N2M through them
blanchet
parents:
diff changeset
   139
end;