src/Pure/System/isabelle_process.scala
changeset 65218 102b8e092860
parent 65216 060a8a1f2dec
child 65225 ec9ec04546fc
--- a/src/Pure/System/isabelle_process.scala	Mon Mar 13 23:14:44 2017 +0100
+++ b/src/Pure/System/isabelle_process.scala	Mon Mar 13 23:24:20 2017 +0100
@@ -19,7 +19,7 @@
   {
     session.start(receiver =>
       Isabelle_Process(options, logic = logic, args = args, dirs = dirs, modes = modes,
-        receiver = receiver, store = store))
+        receiver = receiver, xml_cache = session.xml_cache, store = store))
   }
 
   def apply(
@@ -29,6 +29,7 @@
     dirs: List[Path] = Nil,
     modes: List[String] = Nil,
     receiver: Prover.Receiver = Console.println(_),
+    xml_cache: XML.Cache = new XML.Cache(),
     store: Sessions.Store = Sessions.store()): Isabelle_Process =
   {
     val channel = System_Channel()
@@ -40,13 +41,16 @@
       catch { case exn @ ERROR(_) => channel.accepted(); throw exn }
     process.stdin.close
 
-    new Isabelle_Process(receiver, channel, process)
+    new Isabelle_Process(receiver, xml_cache, channel, process)
   }
 }
 
 class Isabelle_Process private(
-    receiver: Prover.Receiver, channel: System_Channel, process: Prover.System_Process)
-  extends Prover(receiver, channel, process)
+    receiver: Prover.Receiver,
+    xml_cache: XML.Cache,
+    channel: System_Channel,
+    process: Prover.System_Process)
+  extends Prover(receiver, xml_cache, channel, process)
 {
   def encode(s: String): String = Symbol.encode(s)
   def decode(s: String): String = Symbol.decode(s)