src/Tools/Code/code_ml.ML
changeset 48003 1d11af40b106
parent 47609 b3dab1892cda
child 48072 ace701efe203
equal deleted inserted replaced
48002:6de952f4069f 48003:1d11af40b106
    34 
    34 
    35 datatype ml_stmt =
    35 datatype ml_stmt =
    36     ML_Exc of string * (typscheme * int)
    36     ML_Exc of string * (typscheme * int)
    37   | ML_Val of ml_binding
    37   | ML_Val of ml_binding
    38   | ML_Funs of ml_binding list * string list
    38   | ML_Funs of ml_binding list * string list
    39   | ML_Datas of (string * ((vname * sort) list * ((string * vname list) * itype list) list)) list
    39   | ML_Datas of (string * (vname list * ((string * vname list) * itype list) list)) list
    40   | ML_Class of string * (vname * ((class * string) list * (string * itype) list));
    40   | ML_Class of string * (vname * ((class * string) list * (string * itype) list));
    41 
    41 
    42 fun print_product _ [] = NONE
    42 fun print_product _ [] = NONE
    43   | print_product print [x] = SOME (print x)
    43   | print_product print [x] = SOME (print x)
    44   | print_product print xs = (SOME o enum " *" "" "") (map print xs);
    44   | print_product print xs = (SOME o enum " *" "" "") (map print xs);
   165           | print_co ((co, _), tys) = concat [str (deresolve co), str "of",
   165           | print_co ((co, _), tys) = concat [str (deresolve co), str "of",
   166               enum " *" "" "" (map (print_typ (INFX (2, X))) tys)];
   166               enum " *" "" "" (map (print_typ (INFX (2, X))) tys)];
   167       in
   167       in
   168         concat (
   168         concat (
   169           str definer
   169           str definer
   170           :: print_tyco_expr (tyco, map (ITyVar o fst) vs)
   170           :: print_tyco_expr (tyco, map ITyVar vs)
   171           :: str "="
   171           :: str "="
   172           :: separate (str "|") (map print_co cos)
   172           :: separate (str "|") (map print_co cos)
   173         )
   173         )
   174       end;
   174       end;
   175     fun print_def is_pseudo_fun needs_typ definer
   175     fun print_def is_pseudo_fun needs_typ definer
   271             sig_ps
   271             sig_ps
   272             (Pretty.chunks (ps @ semicolon [p] :: pseudo_ps))
   272             (Pretty.chunks (ps @ semicolon [p] :: pseudo_ps))
   273           end
   273           end
   274      | print_stmt (ML_Datas [(tyco, (vs, []))]) =
   274      | print_stmt (ML_Datas [(tyco, (vs, []))]) =
   275           let
   275           let
   276             val ty_p = print_tyco_expr (tyco, map (ITyVar o fst) vs);
   276             val ty_p = print_tyco_expr (tyco, map ITyVar vs);
   277           in
   277           in
   278             pair
   278             pair
   279             [concat [str "type", ty_p]]
   279             [concat [str "type", ty_p]]
   280             (concat [str "datatype", ty_p, str "=", str "EMPTY__"])
   280             (concat [str "datatype", ty_p, str "=", str "EMPTY__"])
   281           end
   281           end
   460           | print_co ((co, _), tys) = concat [str (deresolve co), str "of",
   460           | print_co ((co, _), tys) = concat [str (deresolve co), str "of",
   461               enum " *" "" "" (map (print_typ (INFX (2, X))) tys)];
   461               enum " *" "" "" (map (print_typ (INFX (2, X))) tys)];
   462       in
   462       in
   463         concat (
   463         concat (
   464           str definer
   464           str definer
   465           :: print_tyco_expr (tyco, map (ITyVar o fst) vs)
   465           :: print_tyco_expr (tyco, map ITyVar vs)
   466           :: str "="
   466           :: str "="
   467           :: separate (str "|") (map print_co cos)
   467           :: separate (str "|") (map print_co cos)
   468         )
   468         )
   469       end;
   469       end;
   470     fun print_def is_pseudo_fun needs_typ definer
   470     fun print_def is_pseudo_fun needs_typ definer
   611             sig_ps
   611             sig_ps
   612             (Pretty.chunks (ps @ doublesemicolon [p] :: pseudo_ps))
   612             (Pretty.chunks (ps @ doublesemicolon [p] :: pseudo_ps))
   613           end
   613           end
   614      | print_stmt (ML_Datas [(tyco, (vs, []))]) =
   614      | print_stmt (ML_Datas [(tyco, (vs, []))]) =
   615           let
   615           let
   616             val ty_p = print_tyco_expr (tyco, map (ITyVar o fst) vs);
   616             val ty_p = print_tyco_expr (tyco, map ITyVar vs);
   617           in
   617           in
   618             pair
   618             pair
   619             [concat [str "type", ty_p]]
   619             [concat [str "type", ty_p]]
   620             (concat [str "type", ty_p, str "=", str "EMPTY__"])
   620             (concat [str "type", ty_p, str "=", str "EMPTY__"])
   621           end
   621           end