src/Tools/Code/code_ml.ML
changeset 55681 7714287dc044
parent 55679 59244fc1a7ca
child 55682 def6575032df
equal deleted inserted replaced
55680:bc5edc5dbf18 55681:7714287dc044
   240                   @ map print_classparam_instance inst_params)
   240                   @ map print_classparam_instance inst_params)
   241               :: str ":"
   241               :: str ":"
   242               @@ print_dicttyp (class, tyco `%% map (ITyVar o fst) vs)
   242               @@ print_dicttyp (class, tyco `%% map (ITyVar o fst) vs)
   243             ))
   243             ))
   244           end;
   244           end;
   245     fun print_stmt (ML_Exc (const, (vs_ty, n))) = pair
   245     fun print_stmt _ (ML_Exc (const, (vs_ty, n))) = pair
   246           [print_val_decl print_typscheme (Constant const, vs_ty)]
   246           [print_val_decl print_typscheme (Constant const, vs_ty)]
   247           ((semicolon o map str) (
   247           ((semicolon o map str) (
   248             (if n = 0 then "val" else "fun")
   248             (if n = 0 then "val" else "fun")
   249             :: deresolve_const const
   249             :: deresolve_const const
   250             :: replicate n "_"
   250             :: replicate n "_"
   251             @ "="
   251             @ "="
   252             :: "raise"
   252             :: "raise"
   253             :: "Fail"
   253             :: "Fail"
   254             @@ ML_Syntax.print_string const
   254             @@ ML_Syntax.print_string const
   255           ))
   255           ))
   256       | print_stmt (ML_Val binding) =
   256       | print_stmt _ (ML_Val binding) =
   257           let
   257           let
   258             val (sig_p, p) = print_def (K false) true "val" binding
   258             val (sig_p, p) = print_def (K false) true "val" binding
   259           in pair
   259           in pair
   260             [sig_p]
   260             [sig_p]
   261             (semicolon [p])
   261             (semicolon [p])
   262           end
   262           end
   263       | print_stmt (ML_Funs (binding :: bindings, pseudo_funs)) =
   263       | print_stmt _ (ML_Funs (binding :: bindings, pseudo_funs)) =
   264           let
   264           let
   265             val print_def' = print_def (member (op =) pseudo_funs) false;
   265             val print_def' = print_def (member (op =) pseudo_funs) false;
   266             fun print_pseudo_fun sym = concat [
   266             fun print_pseudo_fun sym = concat [
   267                 str "val",
   267                 str "val",
   268                 (str o deresolve) sym,
   268                 (str o deresolve) sym,
   275             val pseudo_ps = map print_pseudo_fun pseudo_funs;
   275             val pseudo_ps = map print_pseudo_fun pseudo_funs;
   276           in pair
   276           in pair
   277             sig_ps
   277             sig_ps
   278             (Pretty.chunks (ps @ semicolon [p] :: pseudo_ps))
   278             (Pretty.chunks (ps @ semicolon [p] :: pseudo_ps))
   279           end
   279           end
   280      | print_stmt (ML_Datas [(tyco, (vs, []))]) =
   280      | print_stmt _ (ML_Datas [(tyco, (vs, []))]) =
   281           let
   281           let
   282             val ty_p = print_tyco_expr (Type_Constructor tyco, map ITyVar vs);
   282             val ty_p = print_tyco_expr (Type_Constructor tyco, map ITyVar vs);
   283           in
   283           in
   284             pair
   284             pair
   285             [concat [str "type", ty_p]]
   285             [concat [str "type", ty_p]]
   286             (concat [str "datatype", ty_p, str "=", str "EMPTY__"])
   286             (semicolon [str "datatype", ty_p, str "=", str "EMPTY__"])
   287           end
   287           end
   288      | print_stmt (ML_Datas (data :: datas)) = 
   288      | print_stmt export (ML_Datas (data :: datas)) = 
   289           let
   289           let
   290             val sig_ps = print_datatype_decl "datatype" data
   290             val decl_ps = print_datatype_decl "datatype" data
   291               :: map (print_datatype_decl "and") datas;
   291               :: map (print_datatype_decl "and") datas;
   292             val (ps, p) = split_last sig_ps;
   292             val (ps, p) = split_last decl_ps;
   293           in pair
   293           in pair
   294             sig_ps
   294             (if Code_Namespace.is_public export
       
   295               then decl_ps
       
   296               else map (fn (tyco, (vs, _)) =>
       
   297                 concat [str "type", print_tyco_expr (Type_Constructor tyco, map ITyVar vs)])
       
   298                 (data :: datas))
   295             (Pretty.chunks (ps @| semicolon [p]))
   299             (Pretty.chunks (ps @| semicolon [p]))
   296           end
   300           end
   297      | print_stmt (ML_Class (class, (v, (classrels, classparams)))) =
   301      | print_stmt export (ML_Class (class, (v, (classrels, classparams)))) =
   298           let
   302           let
   299             fun print_field s p = concat [str s, str ":", p];
   303             fun print_field s p = concat [str s, str ":", p];
   300             fun print_proj s p = semicolon
   304             fun print_proj s p = semicolon
   301               (map str ["val", s, "=", "#" ^ s, ":"] @| p);
   305               (map str ["val", s, "=", "#" ^ s, ":"] @| p);
   302             fun print_super_class_decl (classrel as (_, super_class)) =
   306             fun print_super_class_decl (classrel as (_, super_class)) =
   315             fun print_classparam_proj (classparam, ty) =
   319             fun print_classparam_proj (classparam, ty) =
   316               print_proj (deresolve_const classparam)
   320               print_proj (deresolve_const classparam)
   317                 (print_typscheme ([(v, [class])], ty));
   321                 (print_typscheme ([(v, [class])], ty));
   318           in pair
   322           in pair
   319             (concat [str "type", print_dicttyp (class, ITyVar v)]
   323             (concat [str "type", print_dicttyp (class, ITyVar v)]
   320               :: map print_super_class_decl classrels
   324               :: (if Code_Namespace.is_public export
   321               @ map print_classparam_decl classparams)
   325                  then map print_super_class_decl classrels
       
   326                    @ map print_classparam_decl classparams
       
   327                  else []))
   322             (Pretty.chunks (
   328             (Pretty.chunks (
   323               concat [
   329               concat [
   324                 str ("type '" ^ v),
   330                 str "type",
   325                 (str o deresolve_class) class,
   331                 print_dicttyp (class, ITyVar v),
   326                 str "=",
   332                 str "=",
   327                 enum "," "{" "};" (
   333                 enum "," "{" "};" (
   328                   map print_super_class_field classrels
   334                   map print_super_class_field classrels
   329                   @ map print_classparam_field classparams
   335                   @ map print_classparam_field classparams
   330                 )
   336                 )
   580                 str ":",
   586                 str ":",
   581                 print_dicttyp (class, tyco `%% map (ITyVar o fst) vs)
   587                 print_dicttyp (class, tyco `%% map (ITyVar o fst) vs)
   582               ]
   588               ]
   583             ))
   589             ))
   584           end;
   590           end;
   585      fun print_stmt (ML_Exc (const, (vs_ty, n))) = pair
   591      fun print_stmt _ (ML_Exc (const, (vs_ty, n))) = pair
   586           [print_val_decl print_typscheme (Constant const, vs_ty)]
   592           [print_val_decl print_typscheme (Constant const, vs_ty)]
   587           ((doublesemicolon o map str) (
   593           ((doublesemicolon o map str) (
   588             "let"
   594             "let"
   589             :: deresolve_const const
   595             :: deresolve_const const
   590             :: replicate n "_"
   596             :: replicate n "_"
   591             @ "="
   597             @ "="
   592             :: "failwith"
   598             :: "failwith"
   593             @@ ML_Syntax.print_string const
   599             @@ ML_Syntax.print_string const
   594           ))
   600           ))
   595       | print_stmt (ML_Val binding) =
   601       | print_stmt _ (ML_Val binding) =
   596           let
   602           let
   597             val (sig_p, p) = print_def (K false) true "let" binding
   603             val (sig_p, p) = print_def (K false) true "let" binding
   598           in pair
   604           in pair
   599             [sig_p]
   605             [sig_p]
   600             (doublesemicolon [p])
   606             (doublesemicolon [p])
   601           end
   607           end
   602       | print_stmt (ML_Funs (binding :: bindings, pseudo_funs)) =
   608       | print_stmt _ (ML_Funs (binding :: bindings, pseudo_funs)) =
   603           let
   609           let
   604             val print_def' = print_def (member (op =) pseudo_funs) false;
   610             val print_def' = print_def (member (op =) pseudo_funs) false;
   605             fun print_pseudo_fun sym = concat [
   611             fun print_pseudo_fun sym = concat [
   606                 str "let",
   612                 str "let",
   607                 (str o deresolve) sym,
   613                 (str o deresolve) sym,
   614             val pseudo_ps = map print_pseudo_fun pseudo_funs;
   620             val pseudo_ps = map print_pseudo_fun pseudo_funs;
   615           in pair
   621           in pair
   616             sig_ps
   622             sig_ps
   617             (Pretty.chunks (ps @ doublesemicolon [p] :: pseudo_ps))
   623             (Pretty.chunks (ps @ doublesemicolon [p] :: pseudo_ps))
   618           end
   624           end
   619      | print_stmt (ML_Datas [(tyco, (vs, []))]) =
   625      | print_stmt _ (ML_Datas [(tyco, (vs, []))]) =
   620           let
   626           let
   621             val ty_p = print_tyco_expr (Type_Constructor tyco, map ITyVar vs);
   627             val ty_p = print_tyco_expr (Type_Constructor tyco, map ITyVar vs);
   622           in
   628           in
   623             pair
   629             pair
   624             [concat [str "type", ty_p]]
   630             [concat [str "type", ty_p]]
   625             (concat [str "type", ty_p, str "=", str "EMPTY__"])
   631             (doublesemicolon [str "type", ty_p, str "=", str "EMPTY__"])
   626           end
   632           end
   627      | print_stmt (ML_Datas (data :: datas)) = 
   633      | print_stmt export (ML_Datas (data :: datas)) = 
   628           let
   634           let
   629             val sig_ps = print_datatype_decl "type" data
   635             val decl_ps = print_datatype_decl "type" data
   630               :: map (print_datatype_decl "and") datas;
   636               :: map (print_datatype_decl "and") datas;
   631             val (ps, p) = split_last sig_ps;
   637             val (ps, p) = split_last decl_ps;
   632           in pair
   638           in pair
   633             sig_ps
   639             (if Code_Namespace.is_public export
       
   640               then decl_ps
       
   641               else map (fn (tyco, (vs, _)) =>
       
   642                 concat [str "type", print_tyco_expr (Type_Constructor tyco, map ITyVar vs)])
       
   643                 (data :: datas))
   634             (Pretty.chunks (ps @| doublesemicolon [p]))
   644             (Pretty.chunks (ps @| doublesemicolon [p]))
   635           end
   645           end
   636      | print_stmt (ML_Class (class, (v, (classrels, classparams)))) =
   646      | print_stmt export (ML_Class (class, (v, (classrels, classparams)))) =
   637           let
   647           let
   638             fun print_field s p = concat [str s, str ":", p];
   648             fun print_field s p = concat [str s, str ":", p];
   639             fun print_super_class_field (classrel as (_, super_class)) =
   649             fun print_super_class_field (classrel as (_, super_class)) =
   640               print_field (deresolve_classrel classrel) (print_dicttyp (super_class, ITyVar v));
   650               print_field (deresolve_classrel classrel) (print_dicttyp (super_class, ITyVar v));
   641             fun print_classparam_decl (classparam, ty) =
   651             fun print_classparam_decl (classparam, ty) =
   703 
   713 
   704 
   714 
   705 
   715 
   706 (** SML/OCaml generic part **)
   716 (** SML/OCaml generic part **)
   707 
   717 
   708 fun ml_program_of_program ctxt module_name reserved identifiers program =
   718 fun ml_program_of_program ctxt module_name reserved identifiers =
   709   let
   719   let
   710     fun namify_const upper base (nsp_const, nsp_type) =
   720     fun namify_const upper base (nsp_const, nsp_type) =
   711       let
   721       let
   712         val (base', nsp_const') =
   722         val (base', nsp_const') =
   713           Name.variant (if upper then first_upper base else base) nsp_const
   723           Name.variant (if upper then first_upper base else base) nsp_const
   780           (Library.commas o map (Code_Symbol.quote ctxt o fst)) stmts);
   790           (Library.commas o map (Code_Symbol.quote ctxt o fst)) stmts);
   781   in
   791   in
   782     Code_Namespace.hierarchical_program ctxt {
   792     Code_Namespace.hierarchical_program ctxt {
   783       module_name = module_name, reserved = reserved, identifiers = identifiers,
   793       module_name = module_name, reserved = reserved, identifiers = identifiers,
   784       empty_nsp = (reserved, reserved), namify_module = pair, namify_stmt = namify_stmt,
   794       empty_nsp = (reserved, reserved), namify_module = pair, namify_stmt = namify_stmt,
   785       cyclic_modules = false, empty_data = (), memorize_data = K I, modify_stmts = modify_stmts } program
   795       cyclic_modules = false, empty_data = (), memorize_data = K I, modify_stmts = modify_stmts }
   786   end;
   796   end;
   787 
   797 
   788 fun serialize_ml print_ml_module print_ml_stmt ctxt
   798 fun serialize_ml print_ml_module print_ml_stmt ctxt
   789     { module_name, reserved_syms, identifiers, includes,
   799     { module_name, reserved_syms, identifiers, includes,
   790       class_syntax, tyco_syntax, const_syntax } program =
   800       class_syntax, tyco_syntax, const_syntax } program =
   791   let
   801   let
   792 
   802 
   793     (* build program *)
   803     (* build program *)
   794     val { deresolver, hierarchical_program = ml_program } =
   804     val { deresolver, hierarchical_program = ml_program } =
   795       ml_program_of_program ctxt module_name (Name.make_context reserved_syms) identifiers program;
   805       ml_program_of_program ctxt module_name (Name.make_context reserved_syms)
       
   806         identifiers program;
   796 
   807 
   797     (* print statements *)
   808     (* print statements *)
   798     fun print_stmt prefix_fragments (_, (export, stmt)) = print_ml_stmt
   809     fun print_stmt prefix_fragments (_, (export, stmt)) = print_ml_stmt
   799       tyco_syntax const_syntax (make_vars reserved_syms)
   810       tyco_syntax const_syntax (make_vars reserved_syms)
   800       (Code_Thingol.is_constr program) (deresolver prefix_fragments) stmt
   811       (Code_Thingol.is_constr program) (deresolver prefix_fragments) export stmt
   801       |> apfst (fn decl => if export then SOME decl else NONE);
   812       |> apfst (fn decl => if Code_Namespace.not_private export then SOME decl else NONE);
   802 
   813 
   803     (* print modules *)
   814     (* print modules *)
   804     fun print_module _ base _ xs =
   815     fun print_module _ base _ xs =
   805       let
   816       let
   806         val (raw_decls, body) = split_list xs;
   817         val (raw_decls, body) = split_list xs;