remove unused selector field from type arg
authorhuffman
Tue, 02 Mar 2010 15:06:02 -0800
changeset 35519 abf45a91d24d
parent 35518 3b20559d809b
child 35520 f433f18d4c41
remove unused selector field from type arg
src/HOLCF/Tools/Domain/domain_extender.ML
src/HOLCF/Tools/Domain/domain_library.ML
--- a/src/HOLCF/Tools/Domain/domain_extender.ML	Tue Mar 02 14:59:24 2010 -0800
+++ b/src/HOLCF/Tools/Domain/domain_extender.ML	Tue Mar 02 15:06:02 2010 -0800
@@ -158,8 +158,7 @@
         (Binding.name_of con,  (* FIXME preverse binding (!?) *)
          mx,
          ListPair.map (fn ((lazy,sel,tp),vn) =>
-           mk_arg ((lazy, Datatype_Aux.dtyp_of_typ new_dts tp),
-                   Option.map Binding.name_of sel,vn))
+           mk_arg ((lazy, Datatype_Aux.dtyp_of_typ new_dts tp), vn))
                       (args, Datatype_Prop.make_tnames (map third args))
         ) : cons;
     val eqs : eq list =
@@ -230,8 +229,7 @@
         (Binding.name_of con,   (* FIXME preverse binding (!?) *)
          mx,
          ListPair.map (fn ((lazy,sel,tp),vn) =>
-           mk_arg ((lazy, Datatype_Aux.dtyp_of_typ new_dts tp),
-                   Option.map Binding.name_of sel,vn))
+           mk_arg ((lazy, Datatype_Aux.dtyp_of_typ new_dts tp), vn))
                       (args, Datatype_Prop.make_tnames (map third args))
         ) : cons;
     val eqs : eq list =
--- a/src/HOLCF/Tools/Domain/domain_library.ML	Tue Mar 02 14:59:24 2010 -0800
+++ b/src/HOLCF/Tools/Domain/domain_library.ML	Tue Mar 02 15:06:02 2010 -0800
@@ -95,11 +95,10 @@
       eqtype arg;
   type cons = string * mixfix * arg list;
   type eq = (string * typ list) * cons list;
-  val mk_arg : (bool * Datatype.dtyp) * string option * string -> arg;
+  val mk_arg : (bool * Datatype.dtyp) * string -> arg;
   val is_lazy : arg -> bool;
   val rec_of : arg -> int;
   val dtyp_of : arg -> Datatype.dtyp;
-  val sel_of : arg -> string option;
   val vname : arg -> string;
   val upd_vname : (string -> string) -> arg -> arg;
   val is_rec : arg -> bool;
@@ -186,7 +185,6 @@
 
 type arg =
      (bool * Datatype.dtyp) *   (*  (lazy, recursive element) *)
-     string option *               (*   selector name    *)
      string;                       (*   argument name    *)
 
 type cons =
@@ -201,15 +199,14 @@
 
 val mk_arg = I;
 
-fun rec_of ((_,dtyp),_,_) =
+fun rec_of ((_,dtyp),_) =
     case dtyp of Datatype_Aux.DtRec i => i | _ => ~1;
 (* FIXME: what about indirect recursion? *)
 
-fun is_lazy arg = fst (first arg);
-fun dtyp_of arg = snd (first arg);
-val sel_of    =       second;
-val     vname =       third;
-val upd_vname =   upd_third;
+fun is_lazy arg = fst (fst arg);
+fun dtyp_of arg = snd (fst arg);
+val     vname =       snd;
+val upd_vname =   apsnd;
 fun is_rec         arg = rec_of arg >=0;
 fun is_nonlazy_rec arg = is_rec arg andalso not (is_lazy arg);
 fun nonlazy     args   = map vname (filter_out is_lazy args);
@@ -229,7 +226,7 @@
 fun big_sprodD ds = case ds of [] => oneD | _ => foldr1 mk_sprodD ds;
 fun big_ssumD ds = case ds of [] => unitD | _ => foldr1 mk_ssumD ds;
 
-fun dtyp_of_arg ((lazy, D), _, _) = if lazy then mk_uD D else D;
+fun dtyp_of_arg ((lazy, D), _) = if lazy then mk_uD D else D;
 fun dtyp_of_cons (_, _, args) = big_sprodD (map dtyp_of_arg args);
 fun dtyp_of_eq (_, cons) = big_ssumD (map dtyp_of_cons cons);