src/Pure/ML/ml_process.scala
author wenzelm
Mon, 21 Jan 2019 16:50:48 +0100
changeset 69702 1adc89c4a795
parent 69573 c7a69b6cd405
child 69831 b35c3839d5d1
permissions -rw-r--r--
clarified ML_OPTIONS on Windows;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
65477
64e61b0f6972 clarified directories;
wenzelm
parents: 65471
diff changeset
     1
/*  Title:      Pure/ML/ml_process.scala
62544
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
     3
62586
a522a5692832 clarified modules;
wenzelm
parents: 62577
diff changeset
     4
The raw ML process.
62544
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
     5
*/
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
     6
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
     7
package isabelle
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
     8
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
     9
62573
27f90319a499 isabelle.Build uses ML_Process directly;
wenzelm
parents: 62565
diff changeset
    10
import java.io.{File => JFile}
27f90319a499 isabelle.Build uses ML_Process directly;
wenzelm
parents: 62565
diff changeset
    11
27f90319a499 isabelle.Build uses ML_Process directly;
wenzelm
parents: 62565
diff changeset
    12
62544
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
    13
object ML_Process
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
    14
{
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
    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
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
    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
6f7ac44365d7 eliminated without magic name;
wenzelm
parents: 62639
diff changeset
    20
    raw_ml_system: Boolean = false,
62573
27f90319a499 isabelle.Build uses ML_Process directly;
wenzelm
parents: 62565
diff changeset
    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
a4ea3a222e0e tuned signature;
wenzelm
parents: 62556
diff changeset
    23
    redirect: Boolean = false,
62603
c077eb5e0b56 clarified cleanup;
wenzelm
parents: 62602
diff changeset
    24
    cleanup: () => Unit = () => (),
68204
a554da2811f2 clarified signature;
wenzelm
parents: 67586
diff changeset
    25
    sessions_structure: Option[Sessions.Structure] = None,
65431
4a3e6cda3b94 provide session base for "isabelle build" and "isabelle console" ML process;
wenzelm
parents: 65420
diff changeset
    26
    session_base: Option[Sessions.Base] = None,
68209
aeffd8f1f079 support Store with options;
wenzelm
parents: 68204
diff changeset
    27
    store: Option[Sessions.Store] = None): Bash.Process =
62544
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
    28
  {
62637
0189fe0f6452 support for Poly/ML heap hierarchy, which saves a lot of disk space;
wenzelm
parents: 62634
diff changeset
    29
    val logic_name = Isabelle_System.default_logic(logic)
68211
wenzelm
parents: 68209
diff changeset
    30
    val _store = store.getOrElse(Sessions.store(options))
68209
aeffd8f1f079 support Store with options;
wenzelm
parents: 68204
diff changeset
    31
62637
0189fe0f6452 support for Poly/ML heap hierarchy, which saves a lot of disk space;
wenzelm
parents: 62634
diff changeset
    32
    val heaps: List[String] =
62643
6f7ac44365d7 eliminated without magic name;
wenzelm
parents: 62639
diff changeset
    33
      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
    34
      else {
65419
457e4fbed731 explicit Sessions.Selection;
wenzelm
parents: 65415
diff changeset
    35
        val selection = Sessions.Selection(sessions = List(logic_name))
67025
961285f581e6 clarifified selection: always wrt. build_graph structure;
wenzelm
parents: 66961
diff changeset
    36
        val selected_sessions =
68204
a554da2811f2 clarified signature;
wenzelm
parents: 67586
diff changeset
    37
          sessions_structure.getOrElse(Sessions.load_structure(options, dirs = dirs))
a554da2811f2 clarified signature;
wenzelm
parents: 67586
diff changeset
    38
            .selection(selection)
66848
982baed14542 clarified signature;
wenzelm
parents: 66782
diff changeset
    39
        selected_sessions.build_requirements(List(logic_name)).
68221
dbef88c2b6c5 support for build_database_server (PostgreSQL);
wenzelm
parents: 68211
diff changeset
    40
          map(a => File.platform_path(_store.the_heap(a)))
62544
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
    41
      }
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
    42
62637
0189fe0f6452 support for Poly/ML heap hierarchy, which saves a lot of disk space;
wenzelm
parents: 62634
diff changeset
    43
    val eval_init =
62634
aa3b47b32100 less physical "logic" argument, with option -l like "isabelle console" etc.;
wenzelm
parents: 62633
diff changeset
    44
      if (heaps.isEmpty) {
62544
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
    45
        List(
62912
745d31e63c21 section headings for ROOT.ML;
wenzelm
parents: 62902
diff changeset
    46
          """
745d31e63c21 section headings for ROOT.ML;
wenzelm
parents: 62902
diff changeset
    47
          fun chapter (_: string) = ();
745d31e63c21 section headings for ROOT.ML;
wenzelm
parents: 62902
diff changeset
    48
          fun section (_: string) = ();
745d31e63c21 section headings for ROOT.ML;
wenzelm
parents: 62902
diff changeset
    49
          fun subsection (_: string) = ();
745d31e63c21 section headings for ROOT.ML;
wenzelm
parents: 62902
diff changeset
    50
          fun subsubsection (_: string) = ();
745d31e63c21 section headings for ROOT.ML;
wenzelm
parents: 62902
diff changeset
    51
          fun paragraph (_: string) = ();
745d31e63c21 section headings for ROOT.ML;
wenzelm
parents: 62902
diff changeset
    52
          fun subparagraph (_: string) = ();
745d31e63c21 section headings for ROOT.ML;
wenzelm
parents: 62902
diff changeset
    53
          val ML_file = PolyML.use;
745d31e63c21 section headings for ROOT.ML;
wenzelm
parents: 62902
diff changeset
    54
          """,
64557
37074e22e8be more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
wenzelm
parents: 64304
diff changeset
    55
          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
    56
            """
37074e22e8be more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
wenzelm
parents: 64304
diff changeset
    57
              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
    58
              structure RunCall =
37074e22e8be more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
wenzelm
parents: 64304
diff changeset
    59
              struct
37074e22e8be more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
wenzelm
parents: 64304
diff changeset
    60
                open RunCall
37074e22e8be more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
wenzelm
parents: 64304
diff changeset
    61
                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
    62
                  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
    63
                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
    64
                  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
    65
              end;
37074e22e8be more tight thread attributes, based in internal word arithmetic instead of symbolic datatypes: measurable performance improvement;
wenzelm
parents: 64304
diff changeset
    66
            """
62855
82859dac5f59 clarified bootstrap -- avoid conditional compilation in ROOT.ML;
wenzelm
parents: 62835
diff changeset
    67
          else "",
62544
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
    68
          if (Platform.is_windows)
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
    69
            "fun exit 0 = OS.Process.exit OS.Process.success" +
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
    70
            " | exit 1 = OS.Process.exit OS.Process.failure" +
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
    71
            " | exit rc = OS.Process.exit (RunCall.unsafeCast (Word8.fromInt rc))"
62560
498f6ff16804 clarified initial ML;
wenzelm
parents: 62557
diff changeset
    72
          else
498f6ff16804 clarified initial ML;
wenzelm
parents: 62557
diff changeset
    73
            "fun exit rc = Posix.Process.exit (Word8.fromInt rc)",
62629
1815513a57f1 clarified prompt: "ML" usually means Isabelle/ML;
wenzelm
parents: 62614
diff changeset
    74
          "PolyML.Compiler.prompt1 := \"Poly/ML> \"",
1815513a57f1 clarified prompt: "ML" usually means Isabelle/ML;
wenzelm
parents: 62614
diff changeset
    75
          "PolyML.Compiler.prompt2 := \"Poly/ML# \"")
62544
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
    76
      }
62637
0189fe0f6452 support for Poly/ML heap hierarchy, which saves a lot of disk space;
wenzelm
parents: 62634
diff changeset
    77
      else
0189fe0f6452 support for Poly/ML heap hierarchy, which saves a lot of disk space;
wenzelm
parents: 62634
diff changeset
    78
        List(
0189fe0f6452 support for Poly/ML heap hierarchy, which saves a lot of disk space;
wenzelm
parents: 62634
diff changeset
    79
          "(PolyML.SaveState.loadHierarchy " +
66782
193c31b79a33 clarified signature;
wenzelm
parents: 66781
diff changeset
    80
            ML_Syntax.print_list(ML_Syntax.print_string_bytes)(heaps) +
62637
0189fe0f6452 support for Poly/ML heap hierarchy, which saves a lot of disk space;
wenzelm
parents: 62634
diff changeset
    81
          "; PolyML.print_depth 0) handle exn => (TextIO.output (TextIO.stdErr, General.exnMessage exn ^ " +
66782
193c31b79a33 clarified signature;
wenzelm
parents: 66781
diff changeset
    82
          ML_Syntax.print_string_bytes(": " + logic_name + "\n") +
62637
0189fe0f6452 support for Poly/ML heap hierarchy, which saves a lot of disk space;
wenzelm
parents: 62634
diff changeset
    83
          "); OS.Process.exit OS.Process.failure)")
62544
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
    84
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
    85
    val eval_modes =
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
    86
      if (modes.isEmpty) Nil
66782
193c31b79a33 clarified signature;
wenzelm
parents: 66781
diff changeset
    87
      else List("Print_Mode.add_modes " + ML_Syntax.print_list(ML_Syntax.print_string_bytes)(modes))
62544
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
    88
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
    89
    // options
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
    90
    val isabelle_process_options = Isabelle_System.tmp_file("options")
69573
c7a69b6cd405 strict bash invocation: proper error checking;
wenzelm
parents: 69572
diff changeset
    91
    Isabelle_System.bash("chmod 600 " + File.bash_path(File.path(isabelle_process_options))).check
62544
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
    92
    File.write(isabelle_process_options, YXML.string_of_body(options.encode))
62573
27f90319a499 isabelle.Build uses ML_Process directly;
wenzelm
parents: 62565
diff changeset
    93
    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
    94
    val eval_options = if (heaps.isEmpty) Nil else List("Options.load_default ()")
62544
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
    95
65431
4a3e6cda3b94 provide session base for "isabelle build" and "isabelle console" ML process;
wenzelm
parents: 65420
diff changeset
    96
    // session base
4a3e6cda3b94 provide session base for "isabelle build" and "isabelle console" ML process;
wenzelm
parents: 65420
diff changeset
    97
    val eval_session_base =
4a3e6cda3b94 provide session base for "isabelle build" and "isabelle console" ML process;
wenzelm
parents: 65420
diff changeset
    98
      session_base match {
4a3e6cda3b94 provide session base for "isabelle build" and "isabelle console" ML process;
wenzelm
parents: 65420
diff changeset
    99
        case None => Nil
4a3e6cda3b94 provide session base for "isabelle build" and "isabelle console" ML process;
wenzelm
parents: 65420
diff changeset
   100
        case Some(base) =>
65441
9425e4d8bdb6 more session_base information in ML;
wenzelm
parents: 65432
diff changeset
   101
          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
   102
            ML_Syntax.print_list(
4a3e6cda3b94 provide session base for "isabelle build" and "isabelle console" ML process;
wenzelm
parents: 65420
diff changeset
   103
              ML_Syntax.print_pair(
66782
193c31b79a33 clarified signature;
wenzelm
parents: 66781
diff changeset
   104
                ML_Syntax.print_string_bytes, ML_Syntax.print_string_bytes))(table)
66712
4c98c929a12a session-qualified theory names are mandatory;
wenzelm
parents: 65532
diff changeset
   105
          def print_list(list: List[String]): String =
66782
193c31b79a33 clarified signature;
wenzelm
parents: 66781
diff changeset
   106
            ML_Syntax.print_list(ML_Syntax.print_string_bytes)(list)
67493
c4e9e0c50487 treat sessions as entities with defining position;
wenzelm
parents: 67471
diff changeset
   107
          def print_sessions(list: List[(String, Position.T)]): String =
c4e9e0c50487 treat sessions as entities with defining position;
wenzelm
parents: 67471
diff changeset
   108
            ML_Syntax.print_list(
c4e9e0c50487 treat sessions as entities with defining position;
wenzelm
parents: 67471
diff changeset
   109
              ML_Syntax.print_pair(ML_Syntax.print_string_bytes, ML_Syntax.print_properties))(list)
c4e9e0c50487 treat sessions as entities with defining position;
wenzelm
parents: 67471
diff changeset
   110
65532
febfd9f78bd4 eliminated default_qualifier: just a constant;
wenzelm
parents: 65478
diff changeset
   111
          List("Resources.init_session_base" +
67493
c4e9e0c50487 treat sessions as entities with defining position;
wenzelm
parents: 67471
diff changeset
   112
            " {sessions = " + print_sessions(base.known.sessions.toList) +
67520
6ff47e27c32d proper signature (amending c4e9e0c50487);
wenzelm
parents: 67493
diff changeset
   113
            ", docs = " + print_list(base.doc_names) +
67219
81e9804b2014 added document antiquotation @{session name};
wenzelm
parents: 67052
diff changeset
   114
            ", global_theories = " + print_table(base.global_theories.toList) +
66717
67dbf5cdc056 more informative loaded_theories: dependencies and syntax;
wenzelm
parents: 66712
diff changeset
   115
            ", loaded_theories = " + print_list(base.loaded_theories.keys) +
65441
9425e4d8bdb6 more session_base information in ML;
wenzelm
parents: 65432
diff changeset
   116
            ", 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
   117
      }
4a3e6cda3b94 provide session base for "isabelle build" and "isabelle console" ML process;
wenzelm
parents: 65420
diff changeset
   118
4a3e6cda3b94 provide session base for "isabelle build" and "isabelle console" ML process;
wenzelm
parents: 65420
diff changeset
   119
    // process
62544
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
   120
    val eval_process =
62634
aa3b47b32100 less physical "logic" argument, with option -l like "isabelle console" etc.;
wenzelm
parents: 62633
diff changeset
   121
      if (heaps.isEmpty)
62712
22a17cec2efe clarified use of options;
wenzelm
parents: 62711
diff changeset
   122
        List("PolyML.print_depth " + ML_Syntax.print_int(options.int("ML_print_depth")))
69572
09a6a7c04b45 more robust system channel via options that are private to the user;
wenzelm
parents: 68221
diff changeset
   123
      else List("Isabelle_Process.init ()")
62544
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
   124
62601
a937889f0086 create ISABELLE_TMP in Scala (despite odd/obsolete chmod in d84b4d39bce1);
wenzelm
parents: 62600
diff changeset
   125
    // ISABELLE_TMP
a937889f0086 create ISABELLE_TMP in Scala (despite odd/obsolete chmod in d84b4d39bce1);
wenzelm
parents: 62600
diff changeset
   126
    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
   127
    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
   128
64274
c8990e5feac9 re-use "threads" for --gcthreads;
wenzelm
parents: 63986
diff changeset
   129
    val ml_runtime_options =
c8990e5feac9 re-use "threads" for --gcthreads;
wenzelm
parents: 63986
diff changeset
   130
    {
c8990e5feac9 re-use "threads" for --gcthreads;
wenzelm
parents: 63986
diff changeset
   131
      val ml_options = Word.explode(Isabelle_System.getenv("ML_OPTIONS"))
69702
1adc89c4a795 clarified ML_OPTIONS on Windows;
wenzelm
parents: 69573
diff changeset
   132
      val ml_options1 =
1adc89c4a795 clarified ML_OPTIONS on Windows;
wenzelm
parents: 69573
diff changeset
   133
        if (ml_options.exists(_.containsSlice("gcthreads"))) ml_options
1adc89c4a795 clarified ML_OPTIONS on Windows;
wenzelm
parents: 69573
diff changeset
   134
        else ml_options ::: List("--gcthreads", options.int("threads").toString)
1adc89c4a795 clarified ML_OPTIONS on Windows;
wenzelm
parents: 69573
diff changeset
   135
      val ml_options2 =
1adc89c4a795 clarified ML_OPTIONS on Windows;
wenzelm
parents: 69573
diff changeset
   136
        if (!Platform.is_windows || ml_options.exists(_.containsSlice("codepage"))) ml_options1
1adc89c4a795 clarified ML_OPTIONS on Windows;
wenzelm
parents: 69573
diff changeset
   137
        else ml_options1 ::: List("--codepage", "utf8")
1adc89c4a795 clarified ML_OPTIONS on Windows;
wenzelm
parents: 69573
diff changeset
   138
      ml_options2
64274
c8990e5feac9 re-use "threads" for --gcthreads;
wenzelm
parents: 63986
diff changeset
   139
    }
c8990e5feac9 re-use "threads" for --gcthreads;
wenzelm
parents: 63986
diff changeset
   140
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
   141
    // 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
   142
    val bash_args =
64274
c8990e5feac9 re-use "threads" for --gcthreads;
wenzelm
parents: 63986
diff changeset
   143
      ml_runtime_options :::
65431
4a3e6cda3b94 provide session base for "isabelle build" and "isabelle console" ML process;
wenzelm
parents: 65420
diff changeset
   144
      (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
   145
        map(eval => List("--eval", eval)).flatten ::: args
62546
wenzelm
parents: 62545
diff changeset
   146
63986
c7a4b03727ae options for process policy, notably for multiprocessor machines;
wenzelm
parents: 62912
diff changeset
   147
    Bash.process(
c7a4b03727ae options for process policy, notably for multiprocessor machines;
wenzelm
parents: 62912
diff changeset
   148
      "exec " + options.string("ML_process_policy") + """ "$ML_HOME/poly" -q """ +
64304
96bc94c87a81 clarified modules;
wenzelm
parents: 64274
diff changeset
   149
        Bash.strings(bash_args),
62614
0a01bc7f0946 prefer Scala over bash function;
wenzelm
parents: 62610
diff changeset
   150
      cwd = cwd,
67586
8b19a8a7f029 more robust: avoid global change of LD_LIBRARY_PATH (e.g. relevant for subprocesses);
wenzelm
parents: 67520
diff changeset
   151
      env = env ++ env_options ++ env_tmp,
62614
0a01bc7f0946 prefer Scala over bash function;
wenzelm
parents: 62610
diff changeset
   152
      redirect = redirect,
62602
96e679f042ec more thorough cleanup -- in Scala;
wenzelm
parents: 62601
diff changeset
   153
      cleanup = () =>
96e679f042ec more thorough cleanup -- in Scala;
wenzelm
parents: 62601
diff changeset
   154
        {
96e679f042ec more thorough cleanup -- in Scala;
wenzelm
parents: 62601
diff changeset
   155
          isabelle_process_options.delete
96e679f042ec more thorough cleanup -- in Scala;
wenzelm
parents: 62601
diff changeset
   156
          Isabelle_System.rm_tree(isabelle_tmp)
62603
c077eb5e0b56 clarified cleanup;
wenzelm
parents: 62602
diff changeset
   157
          cleanup()
62602
96e679f042ec more thorough cleanup -- in Scala;
wenzelm
parents: 62601
diff changeset
   158
        })
62544
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
   159
  }
62586
a522a5692832 clarified modules;
wenzelm
parents: 62577
diff changeset
   160
a522a5692832 clarified modules;
wenzelm
parents: 62577
diff changeset
   161
62835
1a9ce1b13b20 prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
wenzelm
parents: 62712
diff changeset
   162
  /* Isabelle tool wrapper */
62586
a522a5692832 clarified modules;
wenzelm
parents: 62577
diff changeset
   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 isabelle_tool = Isabelle_Tool("process", "raw ML process (batch mode)", args =>
62586
a522a5692832 clarified modules;
wenzelm
parents: 62577
diff changeset
   165
  {
62835
1a9ce1b13b20 prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
wenzelm
parents: 62712
diff changeset
   166
    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
   167
    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
   168
    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
   169
    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
   170
    var options = Options.init()
62586
a522a5692832 clarified modules;
wenzelm
parents: 62577
diff changeset
   171
62835
1a9ce1b13b20 prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
wenzelm
parents: 62712
diff changeset
   172
    val getopts = Getopts("""
62634
aa3b47b32100 less physical "logic" argument, with option -l like "isabelle console" etc.;
wenzelm
parents: 62633
diff changeset
   173
Usage: isabelle process [OPTIONS]
62586
a522a5692832 clarified modules;
wenzelm
parents: 62577
diff changeset
   174
a522a5692832 clarified modules;
wenzelm
parents: 62577
diff changeset
   175
  Options are:
62677
0df43889f496 isabelle process -T THEORY;
wenzelm
parents: 62672
diff changeset
   176
    -T THEORY    load theory
62639
699e86051e35 isabelle process -d;
wenzelm
parents: 62638
diff changeset
   177
    -d DIR       include session directory
62586
a522a5692832 clarified modules;
wenzelm
parents: 62577
diff changeset
   178
    -e ML_EXPR   evaluate ML expression on startup
a522a5692832 clarified modules;
wenzelm
parents: 62577
diff changeset
   179
    -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
   180
    -l NAME      logic session name (default ISABELLE_LOGIC=""" + quote(logic) + """)
62586
a522a5692832 clarified modules;
wenzelm
parents: 62577
diff changeset
   181
    -m MODE      add print mode for output
a522a5692832 clarified modules;
wenzelm
parents: 62577
diff changeset
   182
    -o OPTION    override Isabelle system OPTION (via NAME=VAL or NAME)
a522a5692832 clarified modules;
wenzelm
parents: 62577
diff changeset
   183
62634
aa3b47b32100 less physical "logic" argument, with option -l like "isabelle console" etc.;
wenzelm
parents: 62633
diff changeset
   184
  Run the raw Isabelle ML process in batch mode.
62586
a522a5692832 clarified modules;
wenzelm
parents: 62577
diff changeset
   185
""",
62835
1a9ce1b13b20 prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
wenzelm
parents: 62712
diff changeset
   186
      "T:" -> (arg =>
66782
193c31b79a33 clarified signature;
wenzelm
parents: 66781
diff changeset
   187
        eval_args = eval_args ::: List("--eval", "use_thy " + ML_Syntax.print_string_bytes(arg))),
62835
1a9ce1b13b20 prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
wenzelm
parents: 62712
diff changeset
   188
      "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
   189
      "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
   190
      "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
   191
      "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
   192
      "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
   193
      "o:" -> (arg => options = options + arg))
62586
a522a5692832 clarified modules;
wenzelm
parents: 62577
diff changeset
   194
62835
1a9ce1b13b20 prefer internal tool -- assuming that ISABELLE_TMP_PREFIX is created properly by Isabelle_System.isabelle_tmp_prefix;
wenzelm
parents: 62712
diff changeset
   195
    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
   196
    if (args.isEmpty || !more_args.isEmpty) getopts.usage()
62586
a522a5692832 clarified modules;
wenzelm
parents: 62577
diff changeset
   197
62888
64f44d7279e5 proper return code;
wenzelm
parents: 62855
diff changeset
   198
    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
   199
      result().print_stdout.rc
62888
64f44d7279e5 proper return code;
wenzelm
parents: 62855
diff changeset
   200
    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
   201
  })
62544
efa178abe023 manage the underlying ML process in Scala;
wenzelm
parents:
diff changeset
   202
}