src/Pure/System/isabelle_process.scala
changeset 65310 da9f1ef8ef7c
parent 65225 ec9ec04546fc
child 65316 c0fb8405416c
--- a/src/Pure/System/isabelle_process.scala	Sat Mar 18 18:57:14 2017 +0100
+++ b/src/Pure/System/isabelle_process.scala	Sat Mar 18 20:24:12 2017 +0100
@@ -7,6 +7,9 @@
 package isabelle
 
 
+import java.io.{File => JFile}
+
+
 object Isabelle_Process
 {
   def start(session: Session,
@@ -15,6 +18,9 @@
     args: List[String] = Nil,
     dirs: List[Path] = Nil,
     modes: List[String] = Nil,
+    cwd: JFile = null,
+    env: Map[String, String] = Isabelle_System.settings(),
+    tree: Option[Sessions.Tree] = None,
     store: Sessions.Store = Sessions.store(),
     phase_changed: Session.Phase => Unit = null)
   {
@@ -23,7 +29,8 @@
 
     session.start(receiver =>
       Isabelle_Process(options, logic = logic, args = args, dirs = dirs, modes = modes,
-        receiver = receiver, xml_cache = session.xml_cache, store = store))
+        cwd = cwd, env = env, receiver = receiver, xml_cache = session.xml_cache,
+        tree = tree, store = store))
   }
 
   def apply(
@@ -32,15 +39,18 @@
     args: List[String] = Nil,
     dirs: List[Path] = Nil,
     modes: List[String] = Nil,
+    cwd: JFile = null,
+    env: Map[String, String] = Isabelle_System.settings(),
     receiver: Prover.Receiver = Console.println(_),
     xml_cache: XML.Cache = new XML.Cache(),
+    tree: Option[Sessions.Tree] = None,
     store: Sessions.Store = Sessions.store()): Isabelle_Process =
   {
     val channel = System_Channel()
     val process =
       try {
-        ML_Process(options, logic = logic, args = args, dirs = dirs,
-          modes = modes, store = store, channel = Some(channel))
+        ML_Process(options, logic = logic, args = args, dirs = dirs, modes = modes,
+          cwd = cwd, env = env, tree = tree, store = store, channel = Some(channel))
       }
       catch { case exn @ ERROR(_) => channel.accepted(); throw exn }
     process.stdin.close