| author | wenzelm | 
| Fri, 21 Apr 2017 16:48:58 +0200 | |
| changeset 65538 | a39ef48fbee0 | 
| parent 65532 | febfd9f78bd4 | 
| child 66712 | 4c98c929a12a | 
| permissions | -rw-r--r-- | 
| 65477 | 1  | 
/* Title: Pure/ML/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: 
62633 
diff
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: 
62634 
diff
changeset
 | 
18  | 
dirs: List[Path] = Nil,  | 
| 
62556
 
c115e69f457f
more abstract Session.start, without prover command-line;
 
wenzelm 
parents: 
62548 
diff
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: 
62606 
diff
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,  | 
| 
65415
 
8cd54b18b68b
clarified signature: tree structure is not essential;
 
wenzelm 
parents: 
64557 
diff
changeset
 | 
26  | 
sessions: Option[Sessions.T] = None,  | 
| 
65431
 
4a3e6cda3b94
provide session base for "isabelle build" and "isabelle console" ML process;
 
wenzelm 
parents: 
65420 
diff
changeset
 | 
27  | 
session_base: Option[Sessions.Base] = None,  | 
| 62633 | 28  | 
store: Sessions.Store = Sessions.store()): Bash.Process =  | 
| 62544 | 29  | 
  {
 | 
| 
62637
 
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
 
wenzelm 
parents: 
62634 
diff
changeset
 | 
30  | 
val logic_name = Isabelle_System.default_logic(logic)  | 
| 
 
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
 
wenzelm 
parents: 
62634 
diff
changeset
 | 
31  | 
val heaps: List[String] =  | 
| 62643 | 32  | 
if (raw_ml_system) Nil  | 
| 
62637
 
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
 
wenzelm 
parents: 
62634 
diff
changeset
 | 
33  | 
      else {
 | 
| 65419 | 34  | 
val selection = Sessions.Selection(sessions = List(logic_name))  | 
| 
65415
 
8cd54b18b68b
clarified signature: tree structure is not essential;
 
wenzelm 
parents: 
64557 
diff
changeset
 | 
35  | 
val (_, selected_sessions) =  | 
| 65419 | 36  | 
sessions.getOrElse(Sessions.load(options, dirs)).selection(selection)  | 
| 
65420
 
695d4e22345a
support for static session imports, without affect build hierarchy;
 
wenzelm 
parents: 
65419 
diff
changeset
 | 
37  | 
(selected_sessions.build_ancestors(logic_name) ::: List(logic_name)).  | 
| 
62637
 
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
 
wenzelm 
parents: 
62634 
diff
changeset
 | 
38  | 
map(a => File.platform_path(store.heap(a)))  | 
| 62544 | 39  | 
}  | 
40  | 
||
| 
62637
 
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
 
wenzelm 
parents: 
62634 
diff
changeset
 | 
41  | 
val eval_init =  | 
| 
62634
 
aa3b47b32100
less physical "logic" argument, with option -l like "isabelle console" etc.;
 
wenzelm 
parents: 
62633 
diff
changeset
 | 
42  | 
      if (heaps.isEmpty) {
 | 
| 62544 | 43  | 
List(  | 
| 62912 | 44  | 
"""  | 
45  | 
fun chapter (_: string) = ();  | 
|
46  | 
fun section (_: string) = ();  | 
|
47  | 
fun subsection (_: string) = ();  | 
|
48  | 
fun subsubsection (_: string) = ();  | 
|
49  | 
fun paragraph (_: string) = ();  | 
|
50  | 
fun subparagraph (_: string) = ();  | 
|
51  | 
val ML_file = PolyML.use;  | 
|
52  | 
""",  | 
|
| 
64557
 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 
wenzelm 
parents: 
64304 
diff
changeset
 | 
53  | 
          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: 
64304 
diff
changeset
 | 
54  | 
"""  | 
| 
 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 
wenzelm 
parents: 
64304 
diff
changeset
 | 
55  | 
structure FixedInt = IntInf;  | 
| 
 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 
wenzelm 
parents: 
64304 
diff
changeset
 | 
56  | 
structure RunCall =  | 
| 
 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 
wenzelm 
parents: 
64304 
diff
changeset
 | 
57  | 
struct  | 
| 
 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 
wenzelm 
parents: 
64304 
diff
changeset
 | 
58  | 
open RunCall  | 
| 
 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 
wenzelm 
parents: 
64304 
diff
changeset
 | 
59  | 
val loadWord: word * word -> word =  | 
| 
 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 
wenzelm 
parents: 
64304 
diff
changeset
 | 
60  | 
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: 
64304 
diff
changeset
 | 
61  | 
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: 
64304 
diff
changeset
 | 
62  | 
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: 
64304 
diff
changeset
 | 
63  | 
end;  | 
| 
 
37074e22e8be
more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
 
wenzelm 
parents: 
64304 
diff
changeset
 | 
64  | 
"""  | 
| 
62855
 
82859dac5f59
clarified bootstrap -- avoid conditional compilation in ROOT.ML;
 
wenzelm 
parents: 
62835 
diff
changeset
 | 
65  | 
else "",  | 
| 62544 | 66  | 
if (Platform.is_windows)  | 
67  | 
"fun exit 0 = OS.Process.exit OS.Process.success" +  | 
|
68  | 
" | exit 1 = OS.Process.exit OS.Process.failure" +  | 
|
69  | 
" | exit rc = OS.Process.exit (RunCall.unsafeCast (Word8.fromInt rc))"  | 
|
| 62560 | 70  | 
else  | 
71  | 
"fun exit rc = Posix.Process.exit (Word8.fromInt rc)",  | 
|
| 
62629
 
1815513a57f1
clarified prompt: "ML" usually means Isabelle/ML;
 
wenzelm 
parents: 
62614 
diff
changeset
 | 
72  | 
"PolyML.Compiler.prompt1 := \"Poly/ML> \"",  | 
| 
 
1815513a57f1
clarified prompt: "ML" usually means Isabelle/ML;
 
wenzelm 
parents: 
62614 
diff
changeset
 | 
73  | 
"PolyML.Compiler.prompt2 := \"Poly/ML# \"")  | 
| 62544 | 74  | 
}  | 
| 
62637
 
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
 
wenzelm 
parents: 
62634 
diff
changeset
 | 
75  | 
else  | 
| 
 
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
 
wenzelm 
parents: 
62634 
diff
changeset
 | 
76  | 
List(  | 
| 
 
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
 
wenzelm 
parents: 
62634 
diff
changeset
 | 
77  | 
"(PolyML.SaveState.loadHierarchy " +  | 
| 62638 | 78  | 
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: 
62634 
diff
changeset
 | 
79  | 
"; PolyML.print_depth 0) handle exn => (TextIO.output (TextIO.stdErr, General.exnMessage exn ^ " +  | 
| 62638 | 80  | 
          ML_Syntax.print_string0(": " + logic_name + "\n") +
 | 
| 
62637
 
0189fe0f6452
support for Poly/ML heap hierarchy, which saves a lot of disk space;
 
wenzelm 
parents: 
62634 
diff
changeset
 | 
81  | 
"); OS.Process.exit OS.Process.failure)")  | 
| 62544 | 82  | 
|
83  | 
val eval_modes =  | 
|
84  | 
if (modes.isEmpty) Nil  | 
|
| 62638 | 85  | 
      else List("Print_Mode.add_modes " + ML_Syntax.print_list(ML_Syntax.print_string0)(modes))
 | 
| 62544 | 86  | 
|
87  | 
// options  | 
|
88  | 
    val isabelle_process_options = Isabelle_System.tmp_file("options")
 | 
|
89  | 
File.write(isabelle_process_options, YXML.string_of_body(options.encode))  | 
|
| 62573 | 90  | 
    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: 
62633 
diff
changeset
 | 
91  | 
    val eval_options = if (heaps.isEmpty) Nil else List("Options.load_default ()")
 | 
| 62544 | 92  | 
|
| 
65431
 
4a3e6cda3b94
provide session base for "isabelle build" and "isabelle console" ML process;
 
wenzelm 
parents: 
65420 
diff
changeset
 | 
93  | 
// session base  | 
| 
 
4a3e6cda3b94
provide session base for "isabelle build" and "isabelle console" ML process;
 
wenzelm 
parents: 
65420 
diff
changeset
 | 
94  | 
val eval_session_base =  | 
| 
 
4a3e6cda3b94
provide session base for "isabelle build" and "isabelle console" ML process;
 
wenzelm 
parents: 
65420 
diff
changeset
 | 
95  | 
      session_base match {
 | 
| 
 
4a3e6cda3b94
provide session base for "isabelle build" and "isabelle console" ML process;
 
wenzelm 
parents: 
65420 
diff
changeset
 | 
96  | 
case None => Nil  | 
| 
 
4a3e6cda3b94
provide session base for "isabelle build" and "isabelle console" ML process;
 
wenzelm 
parents: 
65420 
diff
changeset
 | 
97  | 
case Some(base) =>  | 
| 65441 | 98  | 
def print_table(table: List[(String, String)]): String =  | 
| 
65431
 
4a3e6cda3b94
provide session base for "isabelle build" and "isabelle console" ML process;
 
wenzelm 
parents: 
65420 
diff
changeset
 | 
99  | 
ML_Syntax.print_list(  | 
| 
 
4a3e6cda3b94
provide session base for "isabelle build" and "isabelle console" ML process;
 
wenzelm 
parents: 
65420 
diff
changeset
 | 
100  | 
ML_Syntax.print_pair(  | 
| 65441 | 101  | 
ML_Syntax.print_string, ML_Syntax.print_string))(table)  | 
| 65532 | 102  | 
          List("Resources.init_session_base" +
 | 
103  | 
            " {global_theories = " + print_table(base.global_theories.toList) +
 | 
|
| 
65471
 
05e5bffcf1d8
clarified loaded_theories: map to qualified theory name;
 
wenzelm 
parents: 
65457 
diff
changeset
 | 
104  | 
", loaded_theories = " + print_table(base.loaded_theories.toList) +  | 
| 65441 | 105  | 
", known_theories = " + print_table(base.dest_known_theories) + "}")  | 
| 
65431
 
4a3e6cda3b94
provide session base for "isabelle build" and "isabelle console" ML process;
 
wenzelm 
parents: 
65420 
diff
changeset
 | 
106  | 
}  | 
| 
 
4a3e6cda3b94
provide session base for "isabelle build" and "isabelle console" ML process;
 
wenzelm 
parents: 
65420 
diff
changeset
 | 
107  | 
|
| 
 
4a3e6cda3b94
provide session base for "isabelle build" and "isabelle console" ML process;
 
wenzelm 
parents: 
65420 
diff
changeset
 | 
108  | 
// process  | 
| 62544 | 109  | 
val eval_process =  | 
| 
62634
 
aa3b47b32100
less physical "logic" argument, with option -l like "isabelle console" etc.;
 
wenzelm 
parents: 
62633 
diff
changeset
 | 
110  | 
if (heaps.isEmpty)  | 
| 62712 | 111  | 
        List("PolyML.print_depth " + ML_Syntax.print_int(options.int("ML_print_depth")))
 | 
| 62565 | 112  | 
else  | 
113  | 
        channel match {
 | 
|
114  | 
case None =>  | 
|
| 62712 | 115  | 
            List("Isabelle_Process.init_options ()")
 | 
| 62565 | 116  | 
case Some(ch) =>  | 
| 62712 | 117  | 
            List("Isabelle_Process.init_protocol " + ML_Syntax.print_string0(ch.server_name))
 | 
| 62565 | 118  | 
}  | 
| 62544 | 119  | 
|
| 
62601
 
a937889f0086
create ISABELLE_TMP in Scala (despite odd/obsolete chmod in d84b4d39bce1);
 
wenzelm 
parents: 
62600 
diff
changeset
 | 
120  | 
// ISABELLE_TMP  | 
| 
 
a937889f0086
create ISABELLE_TMP in Scala (despite odd/obsolete chmod in d84b4d39bce1);
 
wenzelm 
parents: 
62600 
diff
changeset
 | 
121  | 
    val isabelle_tmp = Isabelle_System.tmp_dir("process")
 | 
| 
 
a937889f0086
create ISABELLE_TMP in Scala (despite odd/obsolete chmod in d84b4d39bce1);
 
wenzelm 
parents: 
62600 
diff
changeset
 | 
122  | 
    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: 
62600 
diff
changeset
 | 
123  | 
|
| 64274 | 124  | 
val ml_runtime_options =  | 
125  | 
    {
 | 
|
126  | 
      val ml_options = Word.explode(Isabelle_System.getenv("ML_OPTIONS"))
 | 
|
127  | 
      if (ml_options.exists(_.containsSlice("gcthreads"))) ml_options
 | 
|
128  | 
      else ml_options ::: List("--gcthreads", options.int("threads").toString)
 | 
|
129  | 
}  | 
|
130  | 
||
| 
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: 
62544 
diff
changeset
 | 
131  | 
// 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: 
62544 
diff
changeset
 | 
132  | 
val bash_args =  | 
| 64274 | 133  | 
ml_runtime_options :::  | 
| 
65431
 
4a3e6cda3b94
provide session base for "isabelle build" and "isabelle console" ML process;
 
wenzelm 
parents: 
65420 
diff
changeset
 | 
134  | 
(eval_init ::: eval_modes ::: eval_options ::: eval_session_base ::: 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: 
62544 
diff
changeset
 | 
135  | 
        map(eval => List("--eval", eval)).flatten ::: args
 | 
| 62546 | 136  | 
|
| 
63986
 
c7a4b03727ae
options for process policy, notably for multiprocessor machines;
 
wenzelm 
parents: 
62912 
diff
changeset
 | 
137  | 
Bash.process(  | 
| 
 
c7a4b03727ae
options for process policy, notably for multiprocessor machines;
 
wenzelm 
parents: 
62912 
diff
changeset
 | 
138  | 
      "exec " + options.string("ML_process_policy") + """ "$ML_HOME/poly" -q """ +
 | 
| 64304 | 139  | 
Bash.strings(bash_args),  | 
| 62614 | 140  | 
cwd = cwd,  | 
141  | 
env =  | 
|
142  | 
Isabelle_System.library_path(env ++ env_options ++ env_tmp,  | 
|
143  | 
          Isabelle_System.getenv_strict("ML_HOME")),
 | 
|
144  | 
redirect = redirect,  | 
|
| 62602 | 145  | 
cleanup = () =>  | 
146  | 
        {
 | 
|
147  | 
isabelle_process_options.delete  | 
|
148  | 
Isabelle_System.rm_tree(isabelle_tmp)  | 
|
| 62603 | 149  | 
cleanup()  | 
| 62602 | 150  | 
})  | 
| 62544 | 151  | 
}  | 
| 62586 | 152  | 
|
153  | 
||
| 
62835
 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 
wenzelm 
parents: 
62712 
diff
changeset
 | 
154  | 
/* Isabelle tool wrapper */  | 
| 62586 | 155  | 
|
| 
62835
 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 
wenzelm 
parents: 
62712 
diff
changeset
 | 
156  | 
  val isabelle_tool = Isabelle_Tool("process", "raw ML process (batch mode)", args =>
 | 
| 62586 | 157  | 
  {
 | 
| 
62835
 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 
wenzelm 
parents: 
62712 
diff
changeset
 | 
158  | 
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: 
62712 
diff
changeset
 | 
159  | 
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: 
62712 
diff
changeset
 | 
160  | 
    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: 
62712 
diff
changeset
 | 
161  | 
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: 
62712 
diff
changeset
 | 
162  | 
var options = Options.init()  | 
| 62586 | 163  | 
|
| 
62835
 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 
wenzelm 
parents: 
62712 
diff
changeset
 | 
164  | 
    val getopts = Getopts("""
 | 
| 
62634
 
aa3b47b32100
less physical "logic" argument, with option -l like "isabelle console" etc.;
 
wenzelm 
parents: 
62633 
diff
changeset
 | 
165  | 
Usage: isabelle process [OPTIONS]  | 
| 62586 | 166  | 
|
167  | 
Options are:  | 
|
| 62677 | 168  | 
-T THEORY load theory  | 
| 62639 | 169  | 
-d DIR include session directory  | 
| 62586 | 170  | 
-e ML_EXPR evaluate ML expression on startup  | 
171  | 
-f ML_FILE evaluate ML file on startup  | 
|
| 
62634
 
aa3b47b32100
less physical "logic" argument, with option -l like "isabelle console" etc.;
 
wenzelm 
parents: 
62633 
diff
changeset
 | 
172  | 
-l NAME logic session name (default ISABELLE_LOGIC=""" + quote(logic) + """)  | 
| 62586 | 173  | 
-m MODE add print mode for output  | 
174  | 
-o OPTION override Isabelle system OPTION (via NAME=VAL or NAME)  | 
|
175  | 
||
| 
62634
 
aa3b47b32100
less physical "logic" argument, with option -l like "isabelle console" etc.;
 
wenzelm 
parents: 
62633 
diff
changeset
 | 
176  | 
Run the raw Isabelle ML process in batch mode.  | 
| 62586 | 177  | 
""",  | 
| 
62835
 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 
wenzelm 
parents: 
62712 
diff
changeset
 | 
178  | 
"T:" -> (arg =>  | 
| 
 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 
wenzelm 
parents: 
62712 
diff
changeset
 | 
179  | 
        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: 
62712 
diff
changeset
 | 
180  | 
"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: 
62712 
diff
changeset
 | 
181  | 
      "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: 
62712 
diff
changeset
 | 
182  | 
      "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: 
62712 
diff
changeset
 | 
183  | 
"l:" -> (arg => logic = arg),  | 
| 
 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 
wenzelm 
parents: 
62712 
diff
changeset
 | 
184  | 
"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: 
62712 
diff
changeset
 | 
185  | 
"o:" -> (arg => options = options + arg))  | 
| 62586 | 186  | 
|
| 
62835
 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 
wenzelm 
parents: 
62712 
diff
changeset
 | 
187  | 
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: 
62712 
diff
changeset
 | 
188  | 
if (args.isEmpty || !more_args.isEmpty) getopts.usage()  | 
| 62586 | 189  | 
|
| 62888 | 190  | 
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: 
62712 
diff
changeset
 | 
191  | 
result().print_stdout.rc  | 
| 62888 | 192  | 
sys.exit(rc)  | 
| 
62835
 
1a9ce1b13b20
prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
 
wenzelm 
parents: 
62712 
diff
changeset
 | 
193  | 
})  | 
| 62544 | 194  | 
}  |