| author | wenzelm | 
| Mon, 13 Mar 2017 20:33:42 +0100 | |
| changeset 65213 | 51c0f094dc02 | 
| parent 64557 | 37074e22e8be | 
| child 65415 | 8cd54b18b68b | 
| permissions | -rw-r--r-- | 
| 62586 | 1 | /* Title: Pure/Tools/ml_process.scala | 
| 62544 | 2 | Author: Makarius | 
| 3 | ||
| 62586 | 4 | The raw ML process. | 
| 62544 | 5 | */ | 
| 6 | ||
| 7 | package isabelle | |
| 8 | ||
| 9 | ||
| 62573 | 10 | import java.io.{File => JFile}
 | 
| 11 | ||
| 12 | ||
| 62544 | 13 | object ML_Process | 
| 14 | {
 | |
| 15 | def apply(options: Options, | |
| 62634 
aa3b47b32100
less physical "logic" argument, with option -l like "isabelle console" etc.;
 wenzelm parents: 
62633diff
changeset | 16 | logic: String = "", | 
| 62544 | 17 | args: List[String] = Nil, | 
| 62637 
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
 wenzelm parents: 
62634diff
changeset | 18 | dirs: List[Path] = Nil, | 
| 62556 
c115e69f457f
more abstract Session.start, without prover command-line;
 wenzelm parents: 
62548diff
changeset | 19 | modes: List[String] = Nil, | 
| 62643 | 20 | raw_ml_system: Boolean = false, | 
| 62573 | 21 | cwd: JFile = null, | 
| 62610 
4c89504c76fb
more uniform signature for various process invocations;
 wenzelm parents: 
62606diff
changeset | 22 | env: Map[String, String] = Isabelle_System.settings(), | 
| 62557 | 23 | redirect: Boolean = false, | 
| 62603 | 24 | cleanup: () => Unit = () => (), | 
| 62633 | 25 | channel: Option[System_Channel] = None, | 
| 62637 
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
 wenzelm parents: 
62634diff
changeset | 26 | tree: Option[Sessions.Tree] = None, | 
| 62633 | 27 | store: Sessions.Store = Sessions.store()): Bash.Process = | 
| 62544 | 28 |   {
 | 
| 62637 
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
 wenzelm parents: 
62634diff
changeset | 29 | val logic_name = Isabelle_System.default_logic(logic) | 
| 
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
 wenzelm parents: 
62634diff
changeset | 30 | val heaps: List[String] = | 
| 62643 | 31 | if (raw_ml_system) Nil | 
| 62637 
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
 wenzelm parents: 
62634diff
changeset | 32 |       else {
 | 
| 62644 | 33 | val (_, session_tree) = | 
| 34 | tree.getOrElse(Sessions.load(options, dirs)).selection(sessions = List(logic_name)) | |
| 62637 
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
 wenzelm parents: 
62634diff
changeset | 35 | (session_tree.ancestors(logic_name) ::: List(logic_name)). | 
| 
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
 wenzelm parents: 
62634diff
changeset | 36 | map(a => File.platform_path(store.heap(a))) | 
| 62544 | 37 | } | 
| 38 | ||
| 62637 
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
 wenzelm parents: 
62634diff
changeset | 39 | val eval_init = | 
| 62634 
aa3b47b32100
less physical "logic" argument, with option -l like "isabelle console" etc.;
 wenzelm parents: 
62633diff
changeset | 40 |       if (heaps.isEmpty) {
 | 
| 62544 | 41 | List( | 
| 62912 | 42 | """ | 
| 43 | fun chapter (_: string) = (); | |
| 44 | fun section (_: string) = (); | |
| 45 | fun subsection (_: string) = (); | |
| 46 | fun subsubsection (_: string) = (); | |
| 47 | fun paragraph (_: string) = (); | |
| 48 | fun subparagraph (_: string) = (); | |
| 49 | val ML_file = PolyML.use; | |
| 50 | """, | |
| 64557 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 wenzelm parents: 
64304diff
changeset | 51 |           if (Isabelle_System.getenv("ML_SYSTEM") == "polyml-5.6")
 | 
| 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 wenzelm parents: 
64304diff
changeset | 52 | """ | 
| 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 wenzelm parents: 
64304diff
changeset | 53 | structure FixedInt = IntInf; | 
| 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 wenzelm parents: 
64304diff
changeset | 54 | structure RunCall = | 
| 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 wenzelm parents: 
64304diff
changeset | 55 | struct | 
| 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 wenzelm parents: 
64304diff
changeset | 56 | open RunCall | 
| 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 wenzelm parents: 
64304diff
changeset | 57 | val loadWord: word * word -> word = | 
| 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 wenzelm parents: 
64304diff
changeset | 58 | run_call2 RuntimeCalls.POLY_SYS_load_word; | 
| 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 wenzelm parents: 
64304diff
changeset | 59 | val storeWord: word * word * word -> unit = | 
| 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 wenzelm parents: 
64304diff
changeset | 60 | run_call3 RuntimeCalls.POLY_SYS_assign_word; | 
| 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 wenzelm parents: 
64304diff
changeset | 61 | end; | 
| 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 wenzelm parents: 
64304diff
changeset | 62 | """ | 
| 62855 
82859dac5f59
clarified bootstrap -- avoid conditional compilation in ROOT.ML;
 wenzelm parents: 
62835diff
changeset | 63 | else "", | 
| 62544 | 64 | if (Platform.is_windows) | 
| 65 | "fun exit 0 = OS.Process.exit OS.Process.success" + | |
| 66 | " | exit 1 = OS.Process.exit OS.Process.failure" + | |
| 67 | " | exit rc = OS.Process.exit (RunCall.unsafeCast (Word8.fromInt rc))" | |
| 62560 | 68 | else | 
| 69 | "fun exit rc = Posix.Process.exit (Word8.fromInt rc)", | |
| 62629 
1815513a57f1
clarified prompt: "ML" usually means Isabelle/ML;
 wenzelm parents: 
62614diff
changeset | 70 | "PolyML.Compiler.prompt1 := \"Poly/ML> \"", | 
| 
1815513a57f1
clarified prompt: "ML" usually means Isabelle/ML;
 wenzelm parents: 
62614diff
changeset | 71 | "PolyML.Compiler.prompt2 := \"Poly/ML# \"") | 
| 62544 | 72 | } | 
| 62637 
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
 wenzelm parents: 
62634diff
changeset | 73 | else | 
| 
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
 wenzelm parents: 
62634diff
changeset | 74 | List( | 
| 
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
 wenzelm parents: 
62634diff
changeset | 75 | "(PolyML.SaveState.loadHierarchy " + | 
| 62638 | 76 | ML_Syntax.print_list(ML_Syntax.print_string0)(heaps) + | 
| 62637 
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
 wenzelm parents: 
62634diff
changeset | 77 | "; PolyML.print_depth 0) handle exn => (TextIO.output (TextIO.stdErr, General.exnMessage exn ^ " + | 
| 62638 | 78 |           ML_Syntax.print_string0(": " + logic_name + "\n") +
 | 
| 62637 
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
 wenzelm parents: 
62634diff
changeset | 79 | "); OS.Process.exit OS.Process.failure)") | 
| 62544 | 80 | |
| 81 | val eval_modes = | |
| 82 | if (modes.isEmpty) Nil | |
| 62638 | 83 |       else List("Print_Mode.add_modes " + ML_Syntax.print_list(ML_Syntax.print_string0)(modes))
 | 
| 62544 | 84 | |
| 85 | // options | |
| 86 |     val isabelle_process_options = Isabelle_System.tmp_file("options")
 | |
| 87 | File.write(isabelle_process_options, YXML.string_of_body(options.encode)) | |
| 62573 | 88 |     val env_options = Map("ISABELLE_PROCESS_OPTIONS" -> File.standard_path(isabelle_process_options))
 | 
| 62634 
aa3b47b32100
less physical "logic" argument, with option -l like "isabelle console" etc.;
 wenzelm parents: 
62633diff
changeset | 89 |     val eval_options = if (heaps.isEmpty) Nil else List("Options.load_default ()")
 | 
| 62544 | 90 | |
| 91 | val eval_process = | |
| 62634 
aa3b47b32100
less physical "logic" argument, with option -l like "isabelle console" etc.;
 wenzelm parents: 
62633diff
changeset | 92 | if (heaps.isEmpty) | 
| 62712 | 93 |         List("PolyML.print_depth " + ML_Syntax.print_int(options.int("ML_print_depth")))
 | 
| 62565 | 94 | else | 
| 95 |         channel match {
 | |
| 96 | case None => | |
| 62712 | 97 |             List("Isabelle_Process.init_options ()")
 | 
| 62565 | 98 | case Some(ch) => | 
| 62712 | 99 |             List("Isabelle_Process.init_protocol " + ML_Syntax.print_string0(ch.server_name))
 | 
| 62565 | 100 | } | 
| 62544 | 101 | |
| 62601 
a937889f0086
create ISABELLE_TMP in Scala (despite odd/obsolete chmod in d84b4d39bce1);
 wenzelm parents: 
62600diff
changeset | 102 | // ISABELLE_TMP | 
| 
a937889f0086
create ISABELLE_TMP in Scala (despite odd/obsolete chmod in d84b4d39bce1);
 wenzelm parents: 
62600diff
changeset | 103 |     val isabelle_tmp = Isabelle_System.tmp_dir("process")
 | 
| 
a937889f0086
create ISABELLE_TMP in Scala (despite odd/obsolete chmod in d84b4d39bce1);
 wenzelm parents: 
62600diff
changeset | 104 |     val env_tmp = Map("ISABELLE_TMP" -> File.standard_path(isabelle_tmp))
 | 
| 
a937889f0086
create ISABELLE_TMP in Scala (despite odd/obsolete chmod in d84b4d39bce1);
 wenzelm parents: 
62600diff
changeset | 105 | |
| 64274 | 106 | val ml_runtime_options = | 
| 107 |     {
 | |
| 108 |       val ml_options = Word.explode(Isabelle_System.getenv("ML_OPTIONS"))
 | |
| 109 |       if (ml_options.exists(_.containsSlice("gcthreads"))) ml_options
 | |
| 110 |       else ml_options ::: List("--gcthreads", options.int("threads").toString)
 | |
| 111 | } | |
| 112 | ||
| 62545 
8ebffdaf2ce2
Bash.process always uses a closed script instead of an open argument list, for extra robustness on Windows, where quoting is not well-defined;
 wenzelm parents: 
62544diff
changeset | 113 | // bash | 
| 
8ebffdaf2ce2
Bash.process always uses a closed script instead of an open argument list, for extra robustness on Windows, where quoting is not well-defined;
 wenzelm parents: 
62544diff
changeset | 114 | val bash_args = | 
| 64274 | 115 | ml_runtime_options ::: | 
| 62668 | 116 | (eval_init ::: eval_modes ::: eval_options ::: eval_process). | 
| 62545 
8ebffdaf2ce2
Bash.process always uses a closed script instead of an open argument list, for extra robustness on Windows, where quoting is not well-defined;
 wenzelm parents: 
62544diff
changeset | 117 |         map(eval => List("--eval", eval)).flatten ::: args
 | 
| 62546 | 118 | |
| 63986 
c7a4b03727ae
options for process policy, notably for multiprocessor machines;
 wenzelm parents: 
62912diff
changeset | 119 | Bash.process( | 
| 
c7a4b03727ae
options for process policy, notably for multiprocessor machines;
 wenzelm parents: 
62912diff
changeset | 120 |       "exec " + options.string("ML_process_policy") + """ "$ML_HOME/poly" -q """ +
 | 
| 64304 | 121 | Bash.strings(bash_args), | 
| 62614 | 122 | cwd = cwd, | 
| 123 | env = | |
| 124 | Isabelle_System.library_path(env ++ env_options ++ env_tmp, | |
| 125 |           Isabelle_System.getenv_strict("ML_HOME")),
 | |
| 126 | redirect = redirect, | |
| 62602 | 127 | cleanup = () => | 
| 128 |         {
 | |
| 129 | isabelle_process_options.delete | |
| 130 | Isabelle_System.rm_tree(isabelle_tmp) | |
| 62603 | 131 | cleanup() | 
| 62602 | 132 | }) | 
| 62544 | 133 | } | 
| 62586 | 134 | |
| 135 | ||
| 62835 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 wenzelm parents: 
62712diff
changeset | 136 | /* Isabelle tool wrapper */ | 
| 62586 | 137 | |
| 62835 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 wenzelm parents: 
62712diff
changeset | 138 |   val isabelle_tool = Isabelle_Tool("process", "raw ML process (batch mode)", args =>
 | 
| 62586 | 139 |   {
 | 
| 62835 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 wenzelm parents: 
62712diff
changeset | 140 | var dirs: List[Path] = Nil | 
| 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 wenzelm parents: 
62712diff
changeset | 141 | var eval_args: List[String] = Nil | 
| 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 wenzelm parents: 
62712diff
changeset | 142 |     var logic = Isabelle_System.getenv("ISABELLE_LOGIC")
 | 
| 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 wenzelm parents: 
62712diff
changeset | 143 | var modes: List[String] = Nil | 
| 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 wenzelm parents: 
62712diff
changeset | 144 | var options = Options.init() | 
| 62586 | 145 | |
| 62835 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 wenzelm parents: 
62712diff
changeset | 146 |     val getopts = Getopts("""
 | 
| 62634 
aa3b47b32100
less physical "logic" argument, with option -l like "isabelle console" etc.;
 wenzelm parents: 
62633diff
changeset | 147 | Usage: isabelle process [OPTIONS] | 
| 62586 | 148 | |
| 149 | Options are: | |
| 62677 | 150 | -T THEORY load theory | 
| 62639 | 151 | -d DIR include session directory | 
| 62586 | 152 | -e ML_EXPR evaluate ML expression on startup | 
| 153 | -f ML_FILE evaluate ML file on startup | |
| 62634 
aa3b47b32100
less physical "logic" argument, with option -l like "isabelle console" etc.;
 wenzelm parents: 
62633diff
changeset | 154 | -l NAME logic session name (default ISABELLE_LOGIC=""" + quote(logic) + """) | 
| 62586 | 155 | -m MODE add print mode for output | 
| 156 | -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) | |
| 157 | ||
| 62634 
aa3b47b32100
less physical "logic" argument, with option -l like "isabelle console" etc.;
 wenzelm parents: 
62633diff
changeset | 158 | Run the raw Isabelle ML process in batch mode. | 
| 62586 | 159 | """, | 
| 62835 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 wenzelm parents: 
62712diff
changeset | 160 | "T:" -> (arg => | 
| 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 wenzelm parents: 
62712diff
changeset | 161 |         eval_args = eval_args ::: List("--eval", "use_thy " + ML_Syntax.print_string0(arg))),
 | 
| 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 wenzelm parents: 
62712diff
changeset | 162 | "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))), | 
| 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 wenzelm parents: 
62712diff
changeset | 163 |       "e:" -> (arg => eval_args = eval_args ::: List("--eval", arg)),
 | 
| 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 wenzelm parents: 
62712diff
changeset | 164 |       "f:" -> (arg => eval_args = eval_args ::: List("--use", arg)),
 | 
| 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 wenzelm parents: 
62712diff
changeset | 165 | "l:" -> (arg => logic = arg), | 
| 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 wenzelm parents: 
62712diff
changeset | 166 | "m:" -> (arg => modes = arg :: modes), | 
| 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 wenzelm parents: 
62712diff
changeset | 167 | "o:" -> (arg => options = options + arg)) | 
| 62586 | 168 | |
| 62835 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 wenzelm parents: 
62712diff
changeset | 169 | val more_args = getopts(args) | 
| 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 wenzelm parents: 
62712diff
changeset | 170 | if (args.isEmpty || !more_args.isEmpty) getopts.usage() | 
| 62586 | 171 | |
| 62888 | 172 | val rc = ML_Process(options, logic = logic, args = eval_args, dirs = dirs, modes = modes). | 
| 62835 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 wenzelm parents: 
62712diff
changeset | 173 | result().print_stdout.rc | 
| 62888 | 174 | sys.exit(rc) | 
| 62835 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 wenzelm parents: 
62712diff
changeset | 175 | }) | 
| 62544 | 176 | } |