src/Pure/ML/ml_process.scala
changeset 76655 b3d458a90aeb
parent 76654 a3177042863d
child 76656 a8f452f7c503
equal deleted inserted replaced
76654:a3177042863d 76655:b3d458a90aeb
    11 import java.io.{File => JFile}
    11 import java.io.{File => JFile}
    12 
    12 
    13 
    13 
    14 object ML_Process {
    14 object ML_Process {
    15   def apply(options: Options,
    15   def apply(options: Options,
    16     sessions_structure: Sessions.Structure,
    16     base_info: Sessions.Base_Info,
    17     store: Sessions.Store,
    17     store: Sessions.Store,
    18     logic: String = "",
    18     logic: String = "",
    19     raw_ml_system: Boolean = false,
    19     raw_ml_system: Boolean = false,
    20     use_prelude: List[String] = Nil,
    20     use_prelude: List[String] = Nil,
    21     eval_main: String = "",
    21     eval_main: String = "",
    22     args: List[String] = Nil,
    22     args: List[String] = Nil,
    23     modes: List[String] = Nil,
    23     modes: List[String] = Nil,
    24     cwd: JFile = null,
    24     cwd: JFile = null,
    25     env: JMap[String, String] = Isabelle_System.settings(),
    25     env: JMap[String, String] = Isabelle_System.settings(),
    26     redirect: Boolean = false,
    26     redirect: Boolean = false,
    27     cleanup: () => Unit = () => (),
    27     cleanup: () => Unit = () => ()
    28     session_base: Option[Sessions.Base] = None
       
    29   ): Bash.Process = {
    28   ): Bash.Process = {
    30     val logic_name = Isabelle_System.default_logic(logic)
    29     val logic_name = Isabelle_System.default_logic(logic)
    31 
    30 
    32     val heaps: List[String] =
    31     val heaps: List[String] =
    33       if (raw_ml_system) Nil
    32       if (raw_ml_system) Nil
    34       else {
    33       else {
    35         sessions_structure.selection(logic_name).
    34         base_info.sessions_structure.selection(logic_name).
    36           build_requirements(List(logic_name)).
    35           build_requirements(List(logic_name)).
    37           map(a => File.platform_path(store.the_heap(a)))
    36           map(a => File.platform_path(store.the_heap(a)))
    38       }
    37       }
    39 
    38 
    40     val eval_init =
    39     val eval_init =
    75     val eval_options = if (heaps.isEmpty) Nil else List("Options.load_default ()")
    74     val eval_options = if (heaps.isEmpty) Nil else List("Options.load_default ()")
    76     val isabelle_process_options = Isabelle_System.tmp_file("options")
    75     val isabelle_process_options = Isabelle_System.tmp_file("options")
    77     Isabelle_System.chmod("600", File.path(isabelle_process_options))
    76     Isabelle_System.chmod("600", File.path(isabelle_process_options))
    78     File.write(isabelle_process_options, YXML.string_of_body(options.encode))
    77     File.write(isabelle_process_options, YXML.string_of_body(options.encode))
    79 
    78 
    80     // session base
    79     // session resources
    81     val (init_session_base, eval_init_session) =
    80     val eval_init_session = if (raw_ml_system) Nil else List("Resources.init_session_env ()")
    82       session_base match {
       
    83         case None => (Sessions.Base.bootstrap, Nil)
       
    84         case Some(base) => (base, List("Resources.init_session_env ()"))
       
    85       }
       
    86     val init_session = Isabelle_System.tmp_file("init_session")
    81     val init_session = Isabelle_System.tmp_file("init_session")
    87     Isabelle_System.chmod("600", File.path(init_session))
    82     Isabelle_System.chmod("600", File.path(init_session))
    88     File.write(init_session, new Resources(sessions_structure, init_session_base).init_session_yxml)
    83     val resources = new Resources(base_info.sessions_structure, base_info.base)
       
    84     File.write(init_session, resources.init_session_yxml)
    89 
    85 
    90     // process
    86     // process
    91     val eval_process =
    87     val eval_process =
    92       proper_string(eval_main).getOrElse(
    88       proper_string(eval_main).getOrElse(
    93         if (heaps.isEmpty) {
    89         if (heaps.isEmpty) {
   174       if (args.isEmpty || more_args.nonEmpty) getopts.usage()
   170       if (args.isEmpty || more_args.nonEmpty) getopts.usage()
   175 
   171 
   176       val base_info = Sessions.base_info(options, logic, dirs = dirs).check_errors
   172       val base_info = Sessions.base_info(options, logic, dirs = dirs).check_errors
   177       val store = Sessions.store(options)
   173       val store = Sessions.store(options)
   178       val result =
   174       val result =
   179         ML_Process(options, base_info.sessions_structure, store, logic = logic, args = eval_args,
   175         ML_Process(options, base_info, store, logic = logic, args = eval_args, modes = modes)
   180           modes = modes, session_base = Some(base_info.base)).result(
   176           .result(
   181             progress_stdout = Output.writeln(_, stdout = true),
   177             progress_stdout = Output.writeln(_, stdout = true),
   182             progress_stderr = Output.writeln(_))
   178             progress_stderr = Output.writeln(_))
   183 
   179 
   184       sys.exit(result.rc)
   180       sys.exit(result.rc)
   185     })
   181     })