diff --git a/src/Pure/PIDE/resources.scala b/src/Pure/PIDE/resources.scala --- a/src/Pure/PIDE/resources.scala +++ b/src/Pure/PIDE/resources.scala @@ -1,479 +1,479 @@ /* Title: Pure/PIDE/resources.scala Author: Makarius Resources for theories and auxiliary files. */ package isabelle import scala.util.parsing.input.Reader import java.io.{File => JFile} object Resources { def empty: Resources = new Resources(Sessions.Structure.empty, Sessions.Structure.empty.bootstrap) def hidden_node(name: Document.Node.Name): Boolean = !name.is_theory || name.theory == Sessions.root_name || File_Format.registry.is_theory(name) def html_document(snapshot: Document.Snapshot): Option[Browser_Info.HTML_Document] = File_Format.registry.get(snapshot.node_name).flatMap(_.html_document(snapshot)) } class Resources( val sessions_structure: Sessions.Structure, val session_base: Sessions.Base, val log: Logger = No_Logger, command_timings: List[Properties.T] = Nil ) { resources => override def toString: String = "Resources(" + session_base.toString + ")" /* init session */ def init_session_yxml: String = { import XML.Encode._ YXML.string_of_body( pair(list(pair(string, properties)), pair(list(pair(string, string)), pair(list(pair(string, list(string))), pair(list(properties), pair(list(pair(string, properties)), pair(list(Scala.encode_fun), pair(list(pair(string, string)), list(string))))))))( (sessions_structure.session_positions, (sessions_structure.dest_session_directories, (sessions_structure.bibtex_entries, (command_timings, (Command_Span.load_commands.map(cmd => (cmd.name, cmd.position)), (Scala.functions, (sessions_structure.global_theories.toList, session_base.loaded_theories.keys))))))))) } /* file formats */ def make_theory_name(name: Document.Node.Name): Option[Document.Node.Name] = File_Format.registry.get(name).flatMap(_.make_theory_name(resources, name)) def make_theory_content(thy_name: Document.Node.Name): Option[String] = File_Format.registry.get_theory(thy_name).flatMap(_.make_theory_content(resources, thy_name)) /* file-system operations */ def append(dir: String, source_path: Path): String = (Path.explode(dir) + source_path).expand.implode def append(node_name: Document.Node.Name, source_path: Path): String = append(node_name.master_dir, source_path) def file_node(file: Path, dir: String = "", theory: String = ""): Document.Node.Name = { val node = append(dir, file) val master_dir = append(dir, file.dir) Document.Node.Name(node, master_dir, theory) } def loaded_theory_node(theory: String): Document.Node.Name = Document.Node.Name(theory, "", theory) /* source files of Isabelle/ML bootstrap */ def source_file(raw_name: String): Option[String] = { if (Path.is_wellformed(raw_name)) { if (Path.is_valid(raw_name)) { def check(p: Path): Option[Path] = if (p.is_file) Some(p) else None val path = Path.explode(raw_name) val path1 = if (path.is_absolute || path.is_current) check(path) else { check(Path.explode("~~/src/Pure") + path) orElse (if (Isabelle_System.getenv("ML_SOURCES") == "") None else check(Path.explode("$ML_SOURCES") + path)) } Some(File.platform_path(path1 getOrElse path)) } else None } else Some(raw_name) } /* theory files */ def load_commands( syntax: Outer_Syntax, name: Document.Node.Name ) : () => List[Command_Span.Span] = { val (is_utf8, raw_text) = with_thy_reader(name, reader => (Scan.reader_is_utf8(reader), reader.source.toString)) () => { if (syntax.has_load_commands(raw_text)) { val text = Symbol.decode(Scan.reader_decode_utf8(is_utf8, raw_text)) syntax.parse_spans(text).filter(_.is_load_command(syntax)) } else Nil } } def loaded_files( syntax: Outer_Syntax, name: Document.Node.Name, spans: List[Command_Span.Span] ) : List[Path] = { val dir = name.master_dir_path for { span <- spans; file <- span.loaded_files(syntax).files } yield (dir + Path.explode(file)).expand } def pure_files(syntax: Outer_Syntax): List[Path] = { val pure_dir = Path.explode("~~/src/Pure") for { (name, theory) <- Thy_Header.ml_roots path = (pure_dir + Path.explode(name)).expand node_name = Document.Node.Name(path.implode, path.dir.implode, theory) file <- loaded_files(syntax, node_name, load_commands(syntax, node_name)()) } yield file } def global_theory(theory: String): Boolean = sessions_structure.global_theories.isDefinedAt(theory) def literal_theory(theory: String): Boolean = Long_Name.is_qualified(theory) || global_theory(theory) def theory_name(qualifier: String, theory: String): String = if (literal_theory(theory)) theory else Long_Name.qualify(qualifier, theory) def find_theory_node(theory: String): Option[Document.Node.Name] = { val thy_file = Path.basic(Long_Name.base_name(theory)).thy val session = sessions_structure.theory_qualifier(theory) val dirs = sessions_structure.get(session) match { case Some(info) => info.dirs case None => Nil } dirs.collectFirst({ case dir if (dir + thy_file).is_file => file_node(dir + thy_file, theory = theory) }) } def import_name(qualifier: String, dir: String, s: String): Document.Node.Name = { val theory = theory_name(qualifier, Thy_Header.import_name(s)) val literal_import = literal_theory(theory) && qualifier != sessions_structure.theory_qualifier(theory) if (literal_import && !Thy_Header.is_base_name(s)) { error("Bad import of theory from other session via file-path: " + quote(s)) } if (session_base.loaded_theory(theory)) loaded_theory_node(theory) else { find_theory_node(theory) match { case Some(node_name) => node_name case None => if (Thy_Header.is_base_name(s) && literal_theory(s)) loaded_theory_node(theory) else file_node(Path.explode(s).thy, dir = dir, theory = theory) } } } def import_name(name: Document.Node.Name, s: String): Document.Node.Name = import_name(sessions_structure.theory_qualifier(name), name.master_dir, s) def import_name(info: Sessions.Info, s: String): Document.Node.Name = import_name(info.name, info.dir.implode, s) def find_theory(file: JFile): Option[Document.Node.Name] = { for { qualifier <- sessions_structure.session_directories.get(File.canonical(file).getParentFile) theory_base <- proper_string(Thy_Header.theory_name(file.getName)) theory = theory_name(qualifier, theory_base) theory_node <- find_theory_node(theory) if File.eq(theory_node.path.file, file) } yield theory_node } def complete_import_name(context_name: Document.Node.Name, s: String): List[String] = { val context_session = sessions_structure.theory_qualifier(context_name) val context_dir = try { Some(context_name.master_dir_path) } catch { case ERROR(_) => None } (for { (session, (info, _)) <- sessions_structure.imports_graph.iterator dir <- (if (session == context_session) context_dir.toList else info.dirs).iterator theory <- Thy_Header.try_read_dir(dir).iterator if Completion.completed(s)(theory) } yield { if (session == context_session || global_theory(theory)) theory else Long_Name.qualify(session, theory) }).toList.sorted } def with_thy_reader[A](name: Document.Node.Name, f: Reader[Char] => A): A = { val path = name.path if (path.is_file) using(Scan.byte_reader(path.file))(f) else if (name.node == name.theory) error("Cannot load theory " + quote(name.theory)) else error ("Cannot load theory file " + path) } def check_thy( node_name: Document.Node.Name, reader: Reader[Char], command: Boolean = true, strict: Boolean = true ): Document.Node.Header = { if (node_name.is_theory && reader.source.length > 0) { try { val header = Thy_Header.read(node_name, reader, command = command, strict = strict) val imports = header.imports.map({ case (s, pos) => val name = import_name(node_name, s) if (Sessions.illegal_theory(name.theory_base_name)) { error("Illegal theory name " + quote(name.theory_base_name) + Position.here(pos)) } else (name, pos) }) Document.Node.Header(imports, header.keywords, header.abbrevs) } catch { case exn: Throwable => Document.Node.bad_header(Exn.message(exn)) } } else Document.Node.no_header } /* special header */ def special_header(name: Document.Node.Name): Option[Document.Node.Header] = { val imports = if (name.theory == Sessions.root_name) List(import_name(name, Sessions.theory_import)) else if (Thy_Header.is_ml_root(name.theory)) List(import_name(name, Thy_Header.ML_BOOTSTRAP)) else if (Thy_Header.is_bootstrap(name.theory)) List(import_name(name, Thy_Header.PURE)) else Nil if (imports.isEmpty) None else Some(Document.Node.Header(imports.map((_, Position.none)))) } /* blobs */ - def undefined_blobs(nodes: Document.Nodes): List[Document.Node.Name] = + def undefined_blobs(version: Document.Version): List[Document.Node.Name] = (for { - (node_name, node) <- nodes.iterator + (node_name, node) <- version.nodes.iterator if !session_base.loaded_theory(node_name) cmd <- node.load_commands.iterator name <- cmd.blobs_undefined.iterator } yield name).toList /* document changes */ def parse_change( reparse_limit: Int, previous: Document.Version, doc_blobs: Document.Blobs, edits: List[Document.Edit_Text], consolidate: List[Document.Node.Name]): Session.Change = Thy_Syntax.parse_change(resources, reparse_limit, previous, doc_blobs, edits, consolidate) def commit(change: Session.Change): Unit = {} /* theory and file dependencies */ def dependencies( thys: List[(Document.Node.Name, Position.T)], progress: Progress = new Progress): Dependencies[Unit] = Dependencies.empty[Unit].require_thys((), thys, progress = progress) def session_dependencies( info: Sessions.Info, progress: Progress = new Progress ) : Dependencies[Options] = { info.theories.foldLeft(Dependencies.empty[Options]) { case (dependencies, (options, thys)) => dependencies.require_thys(options, for { (thy, pos) <- thys } yield (import_name(info, thy), pos), progress = progress) } } object Dependencies { def empty[A]: Dependencies[A] = new Dependencies[A](Nil, Map.empty) private def show_path(names: List[Document.Node.Name]): String = names.map(name => quote(name.theory)).mkString(" via ") private def cycle_msg(names: List[Document.Node.Name]): String = "Cyclic dependency of " + show_path(names) private def required_by(initiators: List[Document.Node.Name]): String = if (initiators.isEmpty) "" else "\n(required by " + show_path(initiators.reverse) + ")" } final class Dependencies[A] private( rev_entries: List[Document.Node.Entry], seen: Map[Document.Node.Name, A] ) { private def cons(entry: Document.Node.Entry): Dependencies[A] = new Dependencies[A](entry :: rev_entries, seen) def require_thy(adjunct: A, thy: (Document.Node.Name, Position.T), initiators: List[Document.Node.Name] = Nil, progress: Progress = new Progress): Dependencies[A] = { val (name, pos) = thy def message: String = "The error(s) above occurred for theory " + quote(name.theory) + Dependencies.required_by(initiators) + Position.here(pos) if (seen.isDefinedAt(name)) this else { val dependencies1 = new Dependencies[A](rev_entries, seen + (name -> adjunct)) if (session_base.loaded_theory(name)) dependencies1 else { try { if (initiators.contains(name)) error(Dependencies.cycle_msg(initiators)) progress.expose_interrupt() val header = try { with_thy_reader(name, check_thy(name, _, command = false)).cat_errors(message) } catch { case ERROR(msg) => cat_error(msg, message) } val entry = Document.Node.Entry(name, header) dependencies1.require_thys(adjunct, header.imports_pos, initiators = name :: initiators, progress = progress).cons(entry) } catch { case e: Throwable => dependencies1.cons(Document.Node.Entry(name, Document.Node.bad_header(Exn.message(e)))) } } } } def require_thys(adjunct: A, thys: List[(Document.Node.Name, Position.T)], progress: Progress = new Progress, initiators: List[Document.Node.Name] = Nil): Dependencies[A] = thys.foldLeft(this)(_.require_thy(adjunct, _, progress = progress, initiators = initiators)) def entries: List[Document.Node.Entry] = rev_entries.reverse def theories: List[Document.Node.Name] = entries.map(_.name) def theories_adjunct: List[(Document.Node.Name, A)] = theories.map(name => (name, seen(name))) def errors: List[String] = entries.flatMap(_.header.errors) def check_errors: Dependencies[A] = errors match { case Nil => this case errs => error(cat_lines(errs)) } lazy val theory_graph: Document.Node.Name.Graph[Unit] = { val regular = theories.toSet val irregular = (for { entry <- entries.iterator imp <- entry.header.imports if !regular(imp) } yield imp).toSet Document.Node.Name.make_graph( irregular.toList.map(name => ((name, ()), Nil)) ::: entries.map(entry => ((entry.name, ()), entry.header.imports))) } lazy val loaded_theories: Graph[String, Outer_Syntax] = entries.foldLeft(session_base.loaded_theories) { case (graph, entry) => val name = entry.name.theory val imports = entry.header.imports.map(_.theory) val graph1 = (name :: imports).foldLeft(graph)(_.default_node(_, Outer_Syntax.empty)) val graph2 = imports.foldLeft(graph1)(_.add_edge(_, name)) val syntax0 = if (name == Thy_Header.PURE) List(Thy_Header.bootstrap_syntax) else Nil val syntax1 = (name :: graph2.imm_preds(name).toList).map(graph2.get_node) val syntax = Outer_Syntax.merge(syntax0 ::: syntax1) + entry.header graph2.map_node(name, _ => syntax) } def get_syntax(name: Document.Node.Name): Outer_Syntax = loaded_theories.get_node(name.theory) lazy val load_commands: List[(Document.Node.Name, List[Command_Span.Span])] = theories.zip( Par_List.map((e: () => List[Command_Span.Span]) => e(), theories.map(name => resources.load_commands(get_syntax(name), name)))) .filter(p => p._2.nonEmpty) def loaded_files( name: Document.Node.Name, spans: List[Command_Span.Span] ) : (String, List[Path]) = { val theory = name.theory val syntax = get_syntax(name) val files1 = resources.loaded_files(syntax, name, spans) val files2 = if (theory == Thy_Header.PURE) pure_files(syntax) else Nil (theory, files1 ::: files2) } def loaded_files: List[Path] = for { (name, spans) <- load_commands file <- loaded_files(name, spans)._2 } yield file def imported_files: List[Path] = { val base_theories = loaded_theories.all_preds(theories.map(_.theory)). filter(session_base.loaded_theories.defined) base_theories.map(theory => session_base.known_theories(theory).name.path) ::: base_theories.flatMap(session_base.known_loaded_files.withDefaultValue(Nil)) } lazy val overall_syntax: Outer_Syntax = Outer_Syntax.merge(loaded_theories.maximals.map(loaded_theories.get_node)) override def toString: String = entries.toString } /* resolve implicit theory dependencies */ def resolve_dependencies[A]( models: Map[A, Document.Model], theories: List[(Document.Node.Name, Position.T)] ): List[Document.Node.Name] = { val model_theories = (for (model <- models.valuesIterator if model.is_theory) yield (model.node_name, Position.none)).toList val thy_files1 = dependencies(model_theories ::: theories).theories val thy_files2 = (for { model <- models.valuesIterator if !model.is_theory thy_name <- make_theory_name(model.node_name) } yield thy_name).toList thy_files1 ::: thy_files2 } } diff --git a/src/Pure/PIDE/session.scala b/src/Pure/PIDE/session.scala --- a/src/Pure/PIDE/session.scala +++ b/src/Pure/PIDE/session.scala @@ -1,756 +1,760 @@ /* Title: Pure/PIDE/session.scala Author: Makarius Options: :folding=explicit: PIDE editor session, potentially with running prover process. */ package isabelle import scala.collection.immutable.Queue import scala.collection.mutable import scala.annotation.tailrec object Session { /* outlets */ object Consumer { def apply[A](name: String)(consume: A => Unit): Consumer[A] = new Consumer[A](name, consume) } final class Consumer[-A] private(val name: String, val consume: A => Unit) class Outlet[A](dispatcher: Consumer_Thread[() => Unit]) { private val consumers = Synchronized[List[Consumer[A]]](Nil) def += (c: Consumer[A]): Unit = consumers.change(Library.update(c)) def -= (c: Consumer[A]): Unit = consumers.change(Library.remove(c)) def post(a: A): Unit = { for (c <- consumers.value.iterator) { dispatcher.send(() => try { c.consume(a) } catch { case exn: Throwable => Output.error_message("Consumer failed: " + quote(c.name) + "\n" + Exn.message(exn)) }) } } } /* change */ sealed case class Change( previous: Document.Version, syntax_changed: List[Document.Node.Name], deps_changed: Boolean, doc_edits: List[Document.Edit_Command], consolidate: List[Document.Node.Name], version: Document.Version) case object Change_Flush /* events */ //{{{ case class Command_Timing(props: Properties.T) case class Theory_Timing(props: Properties.T) case class Runtime_Statistics(props: Properties.T) case class Task_Statistics(props: Properties.T) case class Global_Options(options: Options) case object Caret_Focus case class Raw_Edits(doc_blobs: Document.Blobs, edits: List[Document.Edit_Text]) case class Dialog_Result(id: Document_ID.Generic, serial: Long, result: String) case class Commands_Changed( assignment: Boolean, nodes: Set[Document.Node.Name], commands: Set[Command]) sealed abstract class Phase { def print: String = this match { case Terminated(result) => if (result.ok) "finished" else "failed" case _ => Word.lowercase(this.toString) } } case object Inactive extends Phase // stable case object Startup extends Phase // transient case object Ready extends Phase // metastable case object Shutdown extends Phase // transient case class Terminated(result: Process_Result) extends Phase // stable //}}} /* syslog */ class Syslog(limit: Int) { private var queue = Queue.empty[String] private var length = 0 def += (msg: String): Unit = synchronized { queue = queue.enqueue(msg) length += 1 if (length > limit) queue = queue.dequeue._2 } def content(): String = synchronized { cat_lines(queue.iterator) + (if (length > limit) "\n(A total of " + length + " messages...)" else "") } override def toString: String = "Syslog(" + length + ")" } /* protocol handlers */ type Protocol_Function = Prover.Protocol_Output => Boolean type Protocol_Functions = List[(String, Protocol_Function)] abstract class Protocol_Handler extends Isabelle_System.Service { def init(session: Session): Unit = {} def exit(): Unit = {} def functions: Protocol_Functions = Nil def prover_options(options: Options): Options = options } } class Session(_session_options: => Options, val resources: Resources) extends Document.Session { session => val init_time: Time = Time.now() val cache: Term.Cache = Term.Cache.make() def build_blobs_info(name: Document.Node.Name): Command.Blobs_Info = Command.Blobs_Info.none /* global flags */ @volatile var timing: Boolean = false @volatile var verbose: Boolean = false /* dynamic session options */ def session_options: Options = _session_options def output_delay: Time = session_options.seconds("editor_output_delay") def consolidate_delay: Time = session_options.seconds("editor_consolidate_delay") def prune_delay: Time = session_options.seconds("editor_prune_delay") def prune_size: Int = session_options.int("editor_prune_size") def syslog_limit: Int = session_options.int("editor_syslog_limit") def reparse_limit: Int = session_options.int("editor_reparse_limit") /* dispatcher */ private val dispatcher = Consumer_Thread.fork[() => Unit]("Session.dispatcher", daemon = true) { case e => e(); true } def assert_dispatcher[A](body: => A): A = { assert(dispatcher.check_thread()) body } def require_dispatcher[A](body: => A): A = { require(dispatcher.check_thread(), "not on dispatcher thread") body } def send_dispatcher(body: => Unit): Unit = { if (dispatcher.check_thread()) body else dispatcher.send(() => body) } def send_wait_dispatcher(body: => Unit): Unit = { if (dispatcher.check_thread()) body else dispatcher.send_wait(() => body) } /* outlets */ val finished_theories = new Session.Outlet[Document.Snapshot](dispatcher) val command_timings = new Session.Outlet[Session.Command_Timing](dispatcher) val theory_timings = new Session.Outlet[Session.Theory_Timing](dispatcher) val runtime_statistics = new Session.Outlet[Session.Runtime_Statistics](dispatcher) val task_statistics = new Session.Outlet[Session.Task_Statistics](dispatcher) val global_options = new Session.Outlet[Session.Global_Options](dispatcher) val caret_focus = new Session.Outlet[Session.Caret_Focus.type](dispatcher) val raw_edits = new Session.Outlet[Session.Raw_Edits](dispatcher) val commands_changed = new Session.Outlet[Session.Commands_Changed](dispatcher) val phase_changed = new Session.Outlet[Session.Phase](dispatcher) val syslog_messages = new Session.Outlet[Prover.Output](dispatcher) val raw_output_messages = new Session.Outlet[Prover.Output](dispatcher) val trace_events = new Session.Outlet[Simplifier_Trace.Event.type](dispatcher) val debugger_updates = new Session.Outlet[Debugger.Update.type](dispatcher) val all_messages = new Session.Outlet[Prover.Message](dispatcher) // potential bottle-neck! /** main protocol manager **/ /* internal messages */ private case class Start(start_prover: Prover.Receiver => Prover) private case object Stop private case class Get_State(promise: Promise[Document.State]) private case class Cancel_Exec(exec_id: Document_ID.Exec) private case class Protocol_Command_Raw(name: String, args: List[Bytes]) private case class Protocol_Command_Args(name: String, args: List[String]) private case class Update_Options(options: Options) private case object Consolidate_Execution private case object Prune_History /* phase */ private def post_phase(new_phase: Session.Phase): Session.Phase = { phase_changed.post(new_phase) new_phase } private val _phase = Synchronized[Session.Phase](Session.Inactive) private def phase_=(new_phase: Session.Phase): Unit = _phase.change(_ => post_phase(new_phase)) def phase: Session.Phase = _phase.value def is_ready: Boolean = phase == Session.Ready /* syslog */ def make_syslog(): Session.Syslog = new Session.Syslog(syslog_limit) val syslog: Session.Syslog = make_syslog() /* pipelined change parsing */ private case class Text_Edits( previous: Future[Document.Version], doc_blobs: Document.Blobs, text_edits: List[Document.Edit_Text], consolidate: List[Document.Node.Name], version_result: Promise[Document.Version]) private val change_parser = Consumer_Thread.fork[Text_Edits]("change_parser", daemon = true) { case Text_Edits(previous, doc_blobs, text_edits, consolidate, version_result) => val prev = previous.get_finished val change = Timing.timeit("parse_change", timing) { resources.parse_change(reparse_limit, prev, doc_blobs, text_edits, consolidate) } version_result.fulfill(change.version) manager.send(change) true } /* buffered changes */ private object change_buffer { private var assignment: Boolean = false private var nodes: Set[Document.Node.Name] = Set.empty private var commands: Set[Command] = Set.empty def flush(): Unit = synchronized { if (assignment || nodes.nonEmpty || commands.nonEmpty) commands_changed.post(Session.Commands_Changed(assignment, nodes, commands)) if (nodes.nonEmpty) consolidation.update(more_nodes = nodes) assignment = false nodes = Set.empty commands = Set.empty } private val delay_flush = Delay.first(output_delay) { flush() } def invoke(assign: Boolean, edited_nodes: List[Document.Node.Name], cmds: List[Command]): Unit = synchronized { assignment |= assign for (node <- edited_nodes) { nodes += node } for (command <- cmds) { nodes += command.node_name command.blobs_names.foreach(nodes += _) commands += command } delay_flush.invoke() } def shutdown(): Unit = { delay_flush.revoke() flush() } } /* postponed changes */ private object postponed_changes { private var postponed: List[Session.Change] = Nil def store(change: Session.Change): Unit = synchronized { postponed ::= change } def flush(state: Document.State): List[Session.Change] = synchronized { val (assigned, unassigned) = postponed.partition(change => state.is_assigned(change.previous)) postponed = unassigned assigned.reverse } } /* node consolidation */ private object consolidation { private val delay = Delay.first(consolidate_delay) { manager.send(Consolidate_Execution) } private val init_state: Option[Set[Document.Node.Name]] = Some(Set.empty) private val state = Synchronized(init_state) def exit(): Unit = { delay.revoke() state.change(_ => None) } def update(more_nodes: Set[Document.Node.Name] = Set.empty): Unit = { val active = state.change_result(st => (st.isDefined, st.map(nodes => if (nodes.isEmpty) more_nodes else nodes ++ more_nodes))) if (active) delay.invoke() } def flush(): Set[Document.Node.Name] = state.change_result(st => if (st.isDefined) (st.get, init_state) else (Set.empty, None)) } /* prover process */ private object prover { private val variable = Synchronized[Option[Prover]](None) def defined: Boolean = variable.value.isDefined def get: Prover = variable.value.get def set(p: Prover): Unit = variable.change(_ => Some(p)) def reset(): Unit = variable.change(_ => None) def await_reset(): Unit = variable.guarded_access({ case None => Some((), None) case _ => None }) } /* file formats and protocol handlers */ private lazy val file_formats: File_Format.Session = File_Format.registry.start_session(session) private val protocol_handlers = Protocol_Handlers.init(session) def get_protocol_handler[C <: Session.Protocol_Handler](c: Class[C]): Option[C] = protocol_handlers.get(c.getName) match { case Some(h) if Library.is_subclass(h.getClass, c) => Some(h.asInstanceOf[C]) case _ => None } def init_protocol_handler(handler: Session.Protocol_Handler): Unit = protocol_handlers.init(handler) def prover_options(options: Options): Options = protocol_handlers.prover_options(file_formats.prover_options(options)) /* debugger */ private val debugger_handler = new Debugger.Handler(this) init_protocol_handler(debugger_handler) def debugger: Debugger = debugger_handler.debugger /* manager thread */ private lazy val delay_prune = Delay.first(prune_delay) { manager.send(Prune_History) } private val manager: Consumer_Thread[Any] = { /* global state */ val global_state = Synchronized(Document.State.init) /* raw edits */ def handle_raw_edits( doc_blobs: Document.Blobs = Document.Blobs.empty, edits: List[Document.Edit_Text] = Nil, consolidate: List[Document.Node.Name] = Nil): Unit = { //{{{ require(prover.defined, "prover process not defined (handle_raw_edits)") if (edits.nonEmpty) prover.get.discontinue_execution() val previous = global_state.value.history.tip.version val version = Future.promise[Document.Version] global_state.change(_.continue_history(previous, edits, version)) raw_edits.post(Session.Raw_Edits(doc_blobs, edits)) change_parser.send(Text_Edits(previous, doc_blobs, edits, consolidate, version)) //}}} } /* resulting changes */ def handle_change(change: Session.Change): Unit = { //{{{ require(prover.defined, "prover process not defined (handle_change)") // define commands { val id_commands = new mutable.ListBuffer[Command] def id_command(command: Command): Unit = { for { (name, digest) <- command.blobs_defined if !global_state.value.defined_blob(digest) } { change.version.nodes(name).get_blob match { case Some(blob) => global_state.change(_.define_blob(digest)) prover.get.define_blob(digest, blob.bytes) case None => Output.error_message("Missing blob " + quote(name.toString)) } } if (!global_state.value.defined_command(command.id)) { global_state.change(_.define_command(command)) id_commands += command } } for { (_, edit) <- change.doc_edits } { edit.foreach({ case (c1, c2) => c1.foreach(id_command); c2.foreach(id_command) }) } if (id_commands.nonEmpty) { prover.get.define_commands_bulk(resources, id_commands.toList) } } val assignment = global_state.value.the_assignment(change.previous).check_finished global_state.change(_.define_version(change.version, assignment)) prover.get.update(change.previous.id, change.version.id, change.doc_edits, change.consolidate) resources.commit(change) //}}} } /* prover output */ def handle_output(output: Prover.Output): Unit = { //{{{ def bad_output(): Unit = { if (verbose) Output.warning("Ignoring bad prover output: " + output.message.toString) } def change_command(f: Document.State => (Command.State, Document.State)): Unit = { try { val st = global_state.change_result(f) if (!st.command.span.is_theory) { change_buffer.invoke(false, Nil, List(st.command)) } } catch { case _: Document.State.Fail => bad_output() } } output match { case msg: Prover.Protocol_Output => val handled = protocol_handlers.invoke(msg) if (!handled) { msg.properties match { case Protocol.Command_Timing(props, state_id, timing) if prover.defined => command_timings.post(Session.Command_Timing(props)) val message = XML.elem(Markup.STATUS, List(XML.Elem(Markup.Timing(timing), Nil))) change_command(_.accumulate(state_id, cache.elem(message), cache)) case Markup.Theory_Timing(props) => theory_timings.post(Session.Theory_Timing(props)) case Markup.Task_Statistics(props) => task_statistics.post(Session.Task_Statistics(props)) case Protocol.Export(args) if args.id.isDefined && Value.Long.unapply(args.id.get).isDefined => val id = Value.Long.unapply(args.id.get).get val entry = Export.make_entry(Sessions.DRAFT, args, msg.chunk, cache) change_command(_.add_export(id, (args.serial, entry))) case Protocol.Loading_Theory(node_name, id) => val blobs_info = build_blobs_info(node_name) try { global_state.change(_.begin_theory(node_name, id, msg.text, blobs_info)) } catch { case _: Document.State.Fail => bad_output() } case List(Markup.Commands_Accepted.THIS) => msg.text match { case Protocol.Commands_Accepted(ids) => ids.foreach(id => change_command(_.accumulate(id, Protocol.Commands_Accepted.message, cache))) case _ => bad_output() } case List(Markup.Assign_Update.THIS) => msg.text match { case Protocol.Assign_Update(id, edited, update) => try { val (edited_nodes, cmds) = global_state.change_result(_.assign(id, edited, update)) change_buffer.invoke(true, edited_nodes, cmds) manager.send(Session.Change_Flush) } catch { case _: Document.State.Fail => bad_output() } case _ => bad_output() } delay_prune.invoke() case List(Markup.Removed_Versions.THIS) => msg.text match { case Protocol.Removed(removed) => try { global_state.change(_.removed_versions(removed)) manager.send(Session.Change_Flush) } catch { case _: Document.State.Fail => bad_output() } case _ => bad_output() } case _ => bad_output() } } case _ => output.properties match { case Position.Id(state_id) => change_command(_.accumulate(state_id, output.message, cache)) case _ if output.is_init => val init_ok = try { Isabelle_System.make_services(classOf[Session.Protocol_Handler]) .foreach(init_protocol_handler) true } catch { case exn: Throwable => prover.get.protocol_command("Prover.stop", "1", Exn.message(exn)) false } if (init_ok) { prover.get.options(prover_options(session_options)) prover.get.init_session(resources) phase = Session.Ready debugger.ready() } case Markup.Process_Result(result) if output.is_exit => if (prover.defined) protocol_handlers.exit() for (id <- global_state.value.theories.keys) { val snapshot = global_state.change_result(_.end_theory(id)) finished_theories.post(snapshot) } file_formats.stop_session phase = Session.Terminated(result) prover.reset() case _ => raw_output_messages.post(output) } } //}}} } /* main thread */ Consumer_Thread.fork[Any]("Session.manager", daemon = true) { case arg: Any => //{{{ arg match { case output: Prover.Output => if (output.is_syslog) { syslog += XML.content(output.message) syslog_messages.post(output) } if (output.is_stdout || output.is_stderr) raw_output_messages.post(output) else handle_output(output) all_messages.post(output) case input: Prover.Input => all_messages.post(input) case Start(start_prover) if !prover.defined => prover.set(start_prover(manager.send(_))) case Stop => consolidation.exit() delay_prune.revoke() if (prover.defined) { global_state.change(_ => Document.State.init) prover.get.terminate() } case Get_State(promise) => promise.fulfill(global_state.value) case Consolidate_Execution => if (prover.defined) { val state = global_state.value state.stable_tip_version match { case None => consolidation.update() case Some(version) => val consolidate = consolidation.flush().iterator.filter(name => !resources.session_base.loaded_theory(name) && !state.node_consolidated(version, name) && state.node_maybe_consolidated(version, name)).toList if (consolidate.nonEmpty) handle_raw_edits(consolidate = consolidate) } } case Prune_History => if (prover.defined) { val old_versions = global_state.change_result(_.remove_versions(prune_size)) if (old_versions.nonEmpty) prover.get.remove_versions(old_versions) } case Update_Options(options) => if (prover.defined && is_ready) { prover.get.options(prover_options(options)) handle_raw_edits() } global_options.post(Session.Global_Options(options)) case Cancel_Exec(exec_id) if prover.defined => prover.get.cancel_exec(exec_id) case Session.Raw_Edits(doc_blobs, edits) if prover.defined => handle_raw_edits(doc_blobs = doc_blobs, edits = edits) case Session.Dialog_Result(id, serial, result) if prover.defined => prover.get.dialog_result(serial, result) handle_output(new Prover.Output(Protocol.Dialog_Result(id, serial, result))) case Protocol_Command_Raw(name, args) if prover.defined => prover.get.protocol_command_raw(name, args) case Protocol_Command_Args(name, args) if prover.defined => prover.get.protocol_command_args(name, args) case change: Session.Change if prover.defined => val state = global_state.value if (!state.removing_versions && state.is_assigned(change.previous)) handle_change(change) else postponed_changes.store(change) case Session.Change_Flush if prover.defined => val state = global_state.value if (!state.removing_versions) postponed_changes.flush(state).foreach(handle_change) case bad => if (verbose) Output.warning("Ignoring bad message: " + bad.toString) } true //}}} } } /* main operations */ def get_state(): Document.State = { if (manager.is_active()) { val promise = Future.promise[Document.State] manager.send_wait(Get_State(promise)) promise.join } else Document.State.init } def snapshot(name: Document.Node.Name = Document.Node.Name.empty, pending_edits: List[Text.Edit] = Nil): Document.Snapshot = get_state().snapshot(name, pending_edits) def recent_syntax(name: Document.Node.Name): Outer_Syntax = get_state().recent_finished.version.get_finished.nodes(name).syntax getOrElse resources.session_base.overall_syntax + def stable_tip_version[A](models: Map[A, Document.Model]): Option[Document.Version] = + if (models.valuesIterator.forall(_.is_stable)) get_state().stable_tip_version + else None + @tailrec final def await_stable_snapshot(): Document.Snapshot = { val snapshot = this.snapshot() if (snapshot.is_outdated) { output_delay.sleep() await_stable_snapshot() } else snapshot } def start(start_prover: Prover.Receiver => Prover): Unit = { file_formats _phase.change( { case Session.Inactive => manager.send(Start(start_prover)) post_phase(Session.Startup) case phase => error("Cannot start prover in phase " + quote(phase.print)) }) } def stop(): Process_Result = { val was_ready = _phase.guarded_access( { case Session.Startup | Session.Shutdown => None case Session.Terminated(_) => Some((false, phase)) case Session.Inactive => Some((false, post_phase(Session.Terminated(Process_Result(0))))) case Session.Ready => Some((true, post_phase(Session.Shutdown))) }) if (was_ready) manager.send(Stop) prover.await_reset() change_parser.shutdown() change_buffer.shutdown() manager.shutdown() dispatcher.shutdown() phase match { case Session.Terminated(result) => result case phase => error("Bad session phase after shutdown: " + quote(phase.print)) } } def protocol_command_raw(name: String, args: List[Bytes]): Unit = manager.send(Protocol_Command_Raw(name, args)) def protocol_command_args(name: String, args: List[String]): Unit = manager.send(Protocol_Command_Args(name, args)) def protocol_command(name: String, args: String*): Unit = protocol_command_args(name, args.toList) def cancel_exec(exec_id: Document_ID.Exec): Unit = manager.send(Cancel_Exec(exec_id)) def update(doc_blobs: Document.Blobs, edits: List[Document.Edit_Text]): Unit = if (edits.nonEmpty) manager.send_wait(Session.Raw_Edits(doc_blobs, edits)) def update_options(options: Options): Unit = manager.send_wait(Update_Options(options)) def dialog_result(id: Document_ID.Generic, serial: Long, result: String): Unit = manager.send(Session.Dialog_Result(id, serial, result)) } diff --git a/src/Tools/VSCode/src/language_server.scala b/src/Tools/VSCode/src/language_server.scala --- a/src/Tools/VSCode/src/language_server.scala +++ b/src/Tools/VSCode/src/language_server.scala @@ -1,545 +1,546 @@ /* Title: Tools/VSCode/src/language_server.scala Author: Makarius Server for VS Code Language Server Protocol 2.0/3.0, see also https://github.com/Microsoft/language-server-protocol https://github.com/Microsoft/language-server-protocol/blob/master/protocol.md PIDE protocol extensions depend on system option "vscode_pide_extensions". */ package isabelle.vscode import isabelle._ import java.io.{PrintStream, OutputStream, File => JFile} import scala.annotation.tailrec import scala.collection.mutable object Language_Server { type Editor = isabelle.Editor[Unit] /* Isabelle tool wrapper */ private lazy val default_logic = Isabelle_System.getenv("ISABELLE_LOGIC") val isabelle_tool = Isabelle_Tool("vscode_server", "VSCode Language Server for PIDE", Scala_Project.here, { args => try { var logic_ancestor: Option[String] = None var log_file: Option[Path] = None var logic_requirements = false var dirs: List[Path] = Nil var include_sessions: List[String] = Nil var logic = default_logic var modes: List[String] = Nil var no_build = false var options = Options.init() var verbose = false val getopts = Getopts(""" Usage: isabelle vscode_server [OPTIONS] Options are: -A NAME ancestor session for option -R (default: parent) -L FILE logging on FILE -R NAME build image with requirements from other sessions -d DIR include session directory -i NAME include session in name-space of theories -l NAME logic session name (default ISABELLE_LOGIC=""" + quote(default_logic) + """) -m MODE add print mode for output -n no build of session image on startup -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -v verbose logging Run the VSCode Language Server protocol (JSON RPC) over stdin/stdout. """, "A:" -> (arg => logic_ancestor = Some(arg)), "L:" -> (arg => log_file = Some(Path.explode(File.standard_path(arg)))), "R:" -> (arg => { logic = arg; logic_requirements = true }), "d:" -> (arg => dirs = dirs ::: List(Path.explode(File.standard_path(arg)))), "i:" -> (arg => include_sessions = include_sessions ::: List(arg)), "l:" -> (arg => logic = arg), "m:" -> (arg => modes = arg :: modes), "n" -> (_ => no_build = true), "o:" -> (arg => options = options + arg), "v" -> (_ => verbose = true)) val more_args = getopts(args) if (more_args.nonEmpty) getopts.usage() val log = Logger.make(log_file) val channel = new Channel(System.in, System.out, log, verbose) val server = new Language_Server(channel, options, session_name = logic, session_dirs = dirs, include_sessions = include_sessions, session_ancestor = logic_ancestor, session_requirements = logic_requirements, session_no_build = no_build, modes = modes, log = log) // prevent spurious garbage on the main protocol channel val orig_out = System.out try { System.setOut(new PrintStream(OutputStream.nullOutputStream())) server.start() } finally { System.setOut(orig_out) } } catch { case exn: Throwable => val channel = new Channel(System.in, System.out, No_Logger) channel.error_message(Exn.message(exn)) throw(exn) } }) } class Language_Server( val channel: Channel, options: Options, session_name: String = Language_Server.default_logic, include_sessions: List[String] = Nil, session_dirs: List[Path] = Nil, session_ancestor: Option[String] = None, session_requirements: Boolean = false, session_no_build: Boolean = false, modes: List[String] = Nil, log: Logger = No_Logger ) { server => /* prover session */ private val session_ = Synchronized(None: Option[Session]) def session: Session = session_.value getOrElse error("Server inactive") def resources: VSCode_Resources = session.resources.asInstanceOf[VSCode_Resources] def rendering_offset(node_pos: Line.Node_Position): Option[(VSCode_Rendering, Text.Offset)] = for { model <- resources.get_model(new JFile(node_pos.name)) rendering = model.rendering() offset <- model.content.doc.offset(node_pos.pos) } yield (rendering, offset) private val dynamic_output = Dynamic_Output(server) /* input from client or file-system */ private val file_watcher: File_Watcher = File_Watcher(sync_documents, options.seconds("vscode_load_delay")) private val delay_input: Delay = Delay.last(options.seconds("vscode_input_delay"), channel.Error_Logger) { resources.flush_input(session, channel) } private val delay_load: Delay = Delay.last(options.seconds("vscode_load_delay"), channel.Error_Logger) { val (invoke_input, invoke_load) = resources.resolve_dependencies(session, editor, file_watcher) if (invoke_input) delay_input.invoke() if (invoke_load) delay_load.invoke() } private def close_document(file: JFile): Unit = { if (resources.close_model(file)) { file_watcher.register_parent(file) sync_documents(Set(file)) delay_input.invoke() delay_output.invoke() } } private def sync_documents(changed: Set[JFile]): Unit = { resources.sync_models(changed) delay_input.invoke() delay_output.invoke() } private def change_document( file: JFile, version: Long, changes: List[LSP.TextDocumentChange] ): Unit = { val norm_changes = new mutable.ListBuffer[LSP.TextDocumentChange] @tailrec def norm(chs: List[LSP.TextDocumentChange]): Unit = { if (chs.nonEmpty) { val (full_texts, rest1) = chs.span(_.range.isEmpty) val (edits, rest2) = rest1.span(_.range.nonEmpty) norm_changes ++= full_texts norm_changes ++= edits.sortBy(_.range.get.start)(Line.Position.Ordering).reverse norm(rest2) } } norm(changes) norm_changes.foreach(change => resources.change_model(session, editor, file, version, change.text, change.range)) delay_input.invoke() delay_output.invoke() } /* caret handling */ private val delay_caret_update: Delay = Delay.last(options.seconds("vscode_input_delay"), channel.Error_Logger) { session.caret_focus.post(Session.Caret_Focus) } private def update_caret(caret: Option[(JFile, Line.Position)]): Unit = { resources.update_caret(caret) delay_caret_update.invoke() delay_input.invoke() } /* preview */ private lazy val preview_panel = new Preview_Panel(resources) private lazy val delay_preview: Delay = Delay.last(options.seconds("vscode_output_delay"), channel.Error_Logger) { if (preview_panel.flush(channel)) delay_preview.invoke() } private def request_preview(file: JFile, column: Int): Unit = { preview_panel.request(file, column) delay_preview.invoke() } /* output to client */ private val delay_output: Delay = Delay.last(options.seconds("vscode_output_delay"), channel.Error_Logger) { if (resources.flush_output(channel)) delay_output.invoke() } def update_output(changed_nodes: Iterable[JFile]): Unit = { resources.update_output(changed_nodes) delay_output.invoke() } def update_output_visible(): Unit = { resources.update_output_visible() delay_output.invoke() } private val prover_output = Session.Consumer[Session.Commands_Changed](getClass.getName) { case changed => update_output(changed.nodes.toList.map(resources.node_file(_))) } private val syslog_messages = Session.Consumer[Prover.Output](getClass.getName) { case output => channel.log_writeln(resources.output_xml(output.message)) } /* init and exit */ def init(id: LSP.Id): Unit = { def reply_ok(msg: String): Unit = { channel.write(LSP.Initialize.reply(id, "")) channel.writeln(msg) } def reply_error(msg: String): Unit = { channel.write(LSP.Initialize.reply(id, msg)) channel.error_message(msg) } val try_session = try { val base_info = Sessions.base_info( options, session_name, dirs = session_dirs, include_sessions = include_sessions, session_ancestor = session_ancestor, session_requirements = session_requirements).check_errors def build(no_build: Boolean = false): Build.Results = Build.build(options, selection = Sessions.Selection.session(base_info.session_name), build_heap = true, no_build = no_build, dirs = session_dirs, infos = base_info.infos) if (!session_no_build && !build(no_build = true).ok) { val start_msg = "Build started for Isabelle/" + base_info.session_name + " ..." val fail_msg = "Session build failed -- prover process remains inactive!" val progress = channel.progress(verbose = true) progress.echo(start_msg); channel.writeln(start_msg) if (!build().ok) { progress.echo(fail_msg); error(fail_msg) } } val resources = new VSCode_Resources(options, base_info, log) { override def commit(change: Session.Change): Unit = - if (change.deps_changed || undefined_blobs(change.version.nodes).nonEmpty) + if (change.deps_changed || undefined_blobs(change.version).nonEmpty) { delay_load.invoke() + } } val session_options = options.bool("editor_output_state") = true val session = new Session(session_options, resources) Some((base_info, session)) } catch { case ERROR(msg) => reply_error(msg); None } for ((base_info, session) <- try_session) { session_.change(_ => Some(session)) session.commands_changed += prover_output session.syslog_messages += syslog_messages dynamic_output.init() try { Isabelle_Process.start(session, options, base_info.sessions_structure, Sessions.store(options), modes = modes, logic = base_info.session_name).await_startup() reply_ok("Welcome to Isabelle/" + base_info.session_name + Isabelle_System.isabelle_heading()) } catch { case ERROR(msg) => reply_error(msg) } } } def shutdown(id: LSP.Id): Unit = { def reply(err: String): Unit = channel.write(LSP.Shutdown.reply(id, err)) session_.change({ case Some(session) => session.commands_changed -= prover_output session.syslog_messages -= syslog_messages dynamic_output.exit() delay_load.revoke() file_watcher.shutdown() delay_input.revoke() delay_output.revoke() delay_caret_update.revoke() delay_preview.revoke() val result = session.stop() if (result.ok) reply("") else reply("Prover shutdown failed: " + result.rc) None case None => reply("Prover inactive") None }) } def exit(): Unit = { log("\n") sys.exit(if (session_.value.isEmpty) Process_Result.RC.ok else Process_Result.RC.failure) } /* completion */ def completion(id: LSP.Id, node_pos: Line.Node_Position): Unit = { val result = (for ((rendering, offset) <- rendering_offset(node_pos)) yield rendering.completion(node_pos, offset)) getOrElse Nil channel.write(LSP.Completion.reply(id, result)) } /* spell-checker dictionary */ def update_dictionary(include: Boolean, permanent: Boolean): Unit = { for { spell_checker <- resources.spell_checker.get caret <- resources.get_caret() rendering = caret.model.rendering() range = rendering.before_caret_range(caret.offset) Text.Info(_, word) <- Spell_Checker.current_word(rendering, range) } { spell_checker.update(word, include, permanent) update_output_visible() } } def reset_dictionary(): Unit = { for (spell_checker <- resources.spell_checker.get) { spell_checker.reset() update_output_visible() } } /* hover */ def hover(id: LSP.Id, node_pos: Line.Node_Position): Unit = { val result = for { (rendering, offset) <- rendering_offset(node_pos) info <- rendering.tooltips(VSCode_Rendering.tooltip_elements, Text.Range(offset, offset + 1)) } yield { val range = rendering.model.content.doc.range(info.range) val contents = info.info.map(t => LSP.MarkedString(resources.output_pretty_tooltip(List(t)))) (range, contents) } channel.write(LSP.Hover.reply(id, result)) } /* goto definition */ def goto_definition(id: LSP.Id, node_pos: Line.Node_Position): Unit = { val result = (for ((rendering, offset) <- rendering_offset(node_pos)) yield rendering.hyperlinks(Text.Range(offset, offset + 1))) getOrElse Nil channel.write(LSP.GotoDefinition.reply(id, result)) } /* document highlights */ def document_highlights(id: LSP.Id, node_pos: Line.Node_Position): Unit = { val result = (for ((rendering, offset) <- rendering_offset(node_pos)) yield { val model = rendering.model rendering.caret_focus_ranges(Text.Range(offset, offset + 1), model.content.text_range) .map(r => LSP.DocumentHighlight.text(model.content.doc.range(r))) }) getOrElse Nil channel.write(LSP.DocumentHighlights.reply(id, result)) } /* main loop */ def start(): Unit = { log("Server started " + Date.now()) def handle(json: JSON.T): Unit = { try { json match { case LSP.Initialize(id) => init(id) case LSP.Initialized(()) => case LSP.Shutdown(id) => shutdown(id) case LSP.Exit(()) => exit() case LSP.DidOpenTextDocument(file, _, version, text) => change_document(file, version, List(LSP.TextDocumentChange(None, text))) delay_load.invoke() case LSP.DidChangeTextDocument(file, version, changes) => change_document(file, version, changes) case LSP.DidCloseTextDocument(file) => close_document(file) case LSP.Completion(id, node_pos) => completion(id, node_pos) case LSP.Include_Word(()) => update_dictionary(true, false) case LSP.Include_Word_Permanently(()) => update_dictionary(true, true) case LSP.Exclude_Word(()) => update_dictionary(false, false) case LSP.Exclude_Word_Permanently(()) => update_dictionary(false, true) case LSP.Reset_Words(()) => reset_dictionary() case LSP.Hover(id, node_pos) => hover(id, node_pos) case LSP.GotoDefinition(id, node_pos) => goto_definition(id, node_pos) case LSP.DocumentHighlights(id, node_pos) => document_highlights(id, node_pos) case LSP.Caret_Update(caret) => update_caret(caret) case LSP.State_Init(()) => State_Panel.init(server) case LSP.State_Exit(id) => State_Panel.exit(id) case LSP.State_Locate(id) => State_Panel.locate(id) case LSP.State_Update(id) => State_Panel.update(id) case LSP.State_Auto_Update(id, enabled) => State_Panel.auto_update(id, enabled) case LSP.Preview_Request(file, column) => request_preview(file, column) case _ => if (!LSP.ResponseMessage.is_empty(json)) log("### IGNORED") } } catch { case exn: Throwable => channel.log_error_message(Exn.message(exn)) } } @tailrec def loop(): Unit = { channel.read() match { case Some(json) => json match { case bulk: List[_] => bulk.foreach(handle) case _ => handle(json) } loop() case None => log("### TERMINATE") } } loop() } /* abstract editor operations */ object editor extends Language_Server.Editor { /* session */ override def session: Session = server.session override def flush(): Unit = resources.flush_input(session, channel) override def invoke(): Unit = delay_input.invoke() /* current situation */ override def current_node(context: Unit): Option[Document.Node.Name] = resources.get_caret().map(_.model.node_name) override def current_node_snapshot(context: Unit): Option[Document.Snapshot] = resources.get_caret().map(_.model.snapshot()) override def node_snapshot(name: Document.Node.Name): Document.Snapshot = { resources.get_model(name) match { case Some(model) => model.snapshot() case None => session.snapshot(name) } } def current_command(snapshot: Document.Snapshot): Option[Command] = { resources.get_caret() match { case Some(caret) => snapshot.current_command(caret.node_name, caret.offset) case None => None } } override def current_command(context: Unit, snapshot: Document.Snapshot): Option[Command] = current_command(snapshot) /* overlays */ override def node_overlays(name: Document.Node.Name): Document.Node.Overlays = resources.node_overlays(name) override def insert_overlay(command: Command, fn: String, args: List[String]): Unit = resources.insert_overlay(command, fn, args) override def remove_overlay(command: Command, fn: String, args: List[String]): Unit = resources.remove_overlay(command, fn, args) /* hyperlinks */ override def hyperlink_command( focus: Boolean, snapshot: Document.Snapshot, id: Document_ID.Generic, offset: Symbol.Offset = 0 ): Option[Hyperlink] = { if (snapshot.is_outdated) None else snapshot.find_command_position(id, offset).map(node_pos => new Hyperlink { def follow(unit: Unit): Unit = channel.write(LSP.Caret_Update(node_pos, focus)) }) } /* dispatcher thread */ override def assert_dispatcher[A](body: => A): A = session.assert_dispatcher(body) override def require_dispatcher[A](body: => A): A = session.require_dispatcher(body) override def send_dispatcher(body: => Unit): Unit = session.send_dispatcher(body) override def send_wait_dispatcher(body: => Unit): Unit = session.send_wait_dispatcher(body) } } diff --git a/src/Tools/VSCode/src/vscode_resources.scala b/src/Tools/VSCode/src/vscode_resources.scala --- a/src/Tools/VSCode/src/vscode_resources.scala +++ b/src/Tools/VSCode/src/vscode_resources.scala @@ -1,345 +1,336 @@ /* Title: Tools/VSCode/src/vscode_resources.scala Author: Makarius Resources for VSCode Language Server: file-system access and global state. */ package isabelle.vscode import isabelle._ import java.io.{File => JFile} import scala.util.parsing.input.Reader object VSCode_Resources { /* internal state */ sealed case class State( models: Map[JFile, VSCode_Model] = Map.empty, caret: Option[(JFile, Line.Position)] = None, overlays: Document.Overlays = Document.Overlays.empty, pending_input: Set[JFile] = Set.empty, pending_output: Set[JFile] = Set.empty ) { def update_models(changed: Iterable[(JFile, VSCode_Model)]): State = copy( models = models ++ changed, pending_input = changed.foldLeft(pending_input) { case (set, (file, _)) => set + file }, pending_output = changed.foldLeft(pending_output) { case (set, (file, _)) => set + file }) def update_caret(new_caret: Option[(JFile, Line.Position)]): State = if (caret == new_caret) this else copy( caret = new_caret, pending_input = pending_input ++ caret.map(_._1) ++ new_caret.map(_._1)) def get_caret(file: JFile): Option[Line.Position] = caret match { case Some((caret_file, caret_pos)) if caret_file == file => Some(caret_pos) case _ => None } lazy val document_blobs: Document.Blobs = Document.Blobs( (for { (_, model) <- models.iterator blob <- model.get_blob } yield (model.node_name -> blob)).toMap) def change_overlay(insert: Boolean, file: JFile, command: Command, fn: String, args: List[String]): State = copy( overlays = if (insert) overlays.insert(command, fn, args) else overlays.remove(command, fn, args), pending_input = pending_input + file) } /* caret */ sealed case class Caret(file: JFile, model: VSCode_Model, offset: Text.Offset) { def node_name: Document.Node.Name = model.node_name } } class VSCode_Resources( val options: Options, session_base_info: Sessions.Base_Info, log: Logger = No_Logger) extends Resources( session_base_info.sessions_structure, session_base_info.check_errors.base, log = log ) { resources => private val state = Synchronized(VSCode_Resources.State()) /* options */ def pide_extensions: Boolean = options.bool("vscode_pide_extensions") def unicode_symbols: Boolean = options.bool("vscode_unicode_symbols") def tooltip_margin: Int = options.int("vscode_tooltip_margin") def message_margin: Int = options.int("vscode_message_margin") /* document node name */ def node_file(name: Document.Node.Name): JFile = new JFile(name.node) def node_name(file: JFile): Document.Node.Name = find_theory(file) getOrElse { val node = file.getPath val theory = theory_name(Sessions.DRAFT, Thy_Header.theory_name(node)) if (session_base.loaded_theory(theory)) loaded_theory_node(theory) else { val master_dir = file.getParent Document.Node.Name(node, master_dir, theory) } } override def append(dir: String, source_path: Path): String = { val path = source_path.expand if (dir == "" || path.is_absolute) File.platform_path(path) else if (path.is_current) dir else if (path.is_basic && !dir.endsWith("/") && !dir.endsWith(JFile.separator)) dir + JFile.separator + File.platform_path(path) else if (path.is_basic) dir + File.platform_path(path) else File.absolute_name(new JFile(dir + JFile.separator + File.platform_path(path))) } def get_models: Map[JFile, VSCode_Model] = state.value.models def get_model(file: JFile): Option[VSCode_Model] = get_models.get(file) def get_model(name: Document.Node.Name): Option[VSCode_Model] = get_model(node_file(name)) /* file content */ def read_file_content(name: Document.Node.Name): Option[String] = { make_theory_content(name) orElse { try { Some(Line.normalize(File.read(node_file(name)))) } catch { case ERROR(_) => None } } } def get_file_content(name: Document.Node.Name): Option[String] = get_model(name) match { case Some(model) => Some(model.content.text) case None => read_file_content(name) } override def with_thy_reader[A](name: Document.Node.Name, f: Reader[Char] => A): A = { val file = node_file(name) get_model(file) match { case Some(model) => f(Scan.char_reader(model.content.text)) case None if file.isFile => using(Scan.byte_reader(file))(f) case None => error("No such file: " + quote(file.toString)) } } /* document models */ def visible_node(name: Document.Node.Name): Boolean = get_model(name) match { case Some(model) => model.node_visible case None => false } def change_model( session: Session, editor: Language_Server.Editor, file: JFile, version: Long, text: String, range: Option[Line.Range] = None ): Unit = { state.change { st => val model = st.models.getOrElse(file, VSCode_Model.init(session, editor, node_name(file))) val model1 = (model.change_text(text, range) getOrElse model).set_version(version).external(false) st.update_models(Some(file -> model1)) } } def close_model(file: JFile): Boolean = state.change_result(st => st.models.get(file) match { case None => (false, st) case Some(model) => (true, st.update_models(Some(file -> model.external(true)))) }) def sync_models(changed_files: Set[JFile]): Unit = state.change { st => val changed_models = (for { (file, model) <- st.models.iterator if changed_files(file) && model.external_file text <- read_file_content(model.node_name) model1 <- model.change_text(text) } yield (file, model1)).toList st.update_models(changed_models) } /* overlays */ def node_overlays(name: Document.Node.Name): Document.Node.Overlays = state.value.overlays(name) def insert_overlay(command: Command, fn: String, args: List[String]): Unit = state.change(_.change_overlay(true, node_file(command.node_name), command, fn, args)) def remove_overlay(command: Command, fn: String, args: List[String]): Unit = state.change(_.change_overlay(false, node_file(command.node_name), command, fn, args)) /* resolve dependencies */ def resolve_dependencies( session: Session, editor: Language_Server.Editor, file_watcher: File_Watcher ): (Boolean, Boolean) = { state.change_result { st => val thy_files = resources.resolve_dependencies(st.models, Nil) - val stable_tip_version = - if (st.models.valuesIterator.forall(_.is_stable)) { - session.get_state().stable_tip_version - } - else None - - val aux_files = - stable_tip_version match { - case Some(version) => undefined_blobs(version.nodes) - case None => Nil - } + val stable_tip_version = session.stable_tip_version(st.models) + val aux_files = stable_tip_version.toList.flatMap(undefined_blobs) val loaded_models = (for { node_name <- thy_files.iterator ++ aux_files.iterator file = node_file(node_name) if !st.models.isDefinedAt(file) text <- { file_watcher.register_parent(file); read_file_content(node_name) } } yield { val model = VSCode_Model.init(session, editor, node_name) val model1 = (model.change_text(text) getOrElse model).external(true) (file, model1) }).toList val invoke_input = loaded_models.nonEmpty val invoke_load = stable_tip_version.isEmpty ((invoke_input, invoke_load), st.update_models(loaded_models)) } } /* pending input */ def flush_input(session: Session, channel: Channel): Unit = { state.change { st => val changed_models = (for { file <- st.pending_input.iterator model <- st.models.get(file) (edits, model1) <- model.flush_edits(false, st.document_blobs, file, st.get_caret(file)) } yield (edits, (file, model1))).toList for { ((workspace_edits, _), _) <- changed_models if workspace_edits.nonEmpty } channel.write(LSP.WorkspaceEdit(workspace_edits)) session.update(st.document_blobs, changed_models.flatMap(res => res._1._2)) st.copy( models = st.models ++ changed_models.iterator.map(_._2), pending_input = Set.empty) } } /* pending output */ def update_output(changed_nodes: Iterable[JFile]): Unit = state.change(st => st.copy(pending_output = st.pending_output ++ changed_nodes)) def update_output_visible(): Unit = state.change(st => st.copy(pending_output = st.pending_output ++ (for ((file, model) <- st.models.iterator if model.node_visible) yield file))) def flush_output(channel: Channel): Boolean = { state.change_result { st => val (postponed, flushed) = (for { file <- st.pending_output.iterator model <- st.models.get(file) } yield (file, model, model.rendering())).toList.partition(_._3.snapshot.is_outdated) val changed_iterator = for { (file, model, rendering) <- flushed.iterator (changed_diags, changed_decos, model1) = model.publish(rendering) if changed_diags.isDefined || changed_decos.isDefined } yield { for (diags <- changed_diags) channel.write(LSP.PublishDiagnostics(file, rendering.diagnostics_output(diags))) if (pide_extensions) { for (decos <- changed_decos) channel.write(rendering.decoration_output(decos).json(file)) } (file, model1) } (postponed.nonEmpty, st.copy( models = st.models ++ changed_iterator, pending_output = postponed.map(_._1).toSet)) } } /* output text */ def output_text(text: String): String = Symbol.output(unicode_symbols, text) def output_xml(xml: XML.Tree): String = output_text(XML.content(xml)) def output_pretty(body: XML.Body, margin: Int): String = output_text(Pretty.string_of(body, margin = margin, metric = Symbol.Metric)) def output_pretty_tooltip(body: XML.Body): String = output_pretty(body, tooltip_margin) def output_pretty_message(body: XML.Body): String = output_pretty(body, message_margin) /* caret handling */ def update_caret(caret: Option[(JFile, Line.Position)]): Unit = state.change(_.update_caret(caret)) def get_caret(): Option[VSCode_Resources.Caret] = { val st = state.value for { (file, pos) <- st.caret model <- st.models.get(file) offset <- model.content.doc.offset(pos) } yield VSCode_Resources.Caret(file, model, offset) } /* spell checker */ val spell_checker = new Spell_Checker_Variable spell_checker.update(options) } diff --git a/src/Tools/jEdit/src/jedit_resources.scala b/src/Tools/jEdit/src/jedit_resources.scala --- a/src/Tools/jEdit/src/jedit_resources.scala +++ b/src/Tools/jEdit/src/jedit_resources.scala @@ -1,139 +1,139 @@ /* Title: Tools/jEdit/src/jedit_resources.scala Author: Makarius Resources for theories and auxiliary files, based on jEdit buffer content and virtual file-systems. */ package isabelle.jedit import isabelle._ import java.io.{File => JFile, ByteArrayOutputStream} import javax.swing.text.Segment import scala.util.parsing.input.Reader import org.gjt.sp.jedit.io.{VFS, FileVFS, VFSManager} import org.gjt.sp.jedit.MiscUtilities import org.gjt.sp.jedit.{jEdit, View, Buffer} import org.gjt.sp.jedit.bufferio.BufferIORequest object JEdit_Resources { def apply(options: Options): JEdit_Resources = new JEdit_Resources(JEdit_Sessions.session_base_info(options)) } class JEdit_Resources private(val session_base_info: Sessions.Base_Info) extends Resources(session_base_info.sessions_structure, session_base_info.base) { def session_errors: List[String] = session_base_info.errors /* document node name */ def node_name(path: String): Document.Node.Name = JEdit_Lib.check_file(path).flatMap(find_theory) getOrElse { val vfs = VFSManager.getVFSForPath(path) val node = if (vfs.isInstanceOf[FileVFS]) MiscUtilities.resolveSymlinks(path) else path val theory = theory_name(Sessions.DRAFT, Thy_Header.theory_name(node)) if (session_base.loaded_theory(theory)) loaded_theory_node(theory) else { val master_dir = vfs.getParentOfPath(path) Document.Node.Name(node, master_dir, theory) } } def node_name(buffer: Buffer): Document.Node.Name = node_name(JEdit_Lib.buffer_name(buffer)) def theory_node_name(buffer: Buffer): Option[Document.Node.Name] = { val name = node_name(buffer) if (name.is_theory) Some(name) else None } override def append(dir: String, source_path: Path): String = { val path = source_path.expand if (dir == "" || path.is_absolute) MiscUtilities.resolveSymlinks(File.platform_path(path)) else if (path.is_current) dir else { val vfs = VFSManager.getVFSForPath(dir) if (vfs.isInstanceOf[FileVFS]) MiscUtilities.resolveSymlinks( vfs.constructPath(dir, File.platform_path(path))) else vfs.constructPath(dir, File.standard_path(path)) } } /* file content */ def read_file_content(node_name: Document.Node.Name): Option[String] = { make_theory_content(node_name) orElse { val name = node_name.node try { val text = if (Url.is_wellformed(name)) Url.read(Url(name)) else File.read(new JFile(name)) Some(Symbol.decode(Line.normalize(text))) } catch { case ERROR(_) => None } } } def get_file_content(node_name: Document.Node.Name): Option[String] = Document_Model.get(node_name) match { case Some(model: Buffer_Model) => Some(JEdit_Lib.buffer_text(model.buffer)) case Some(model: File_Model) => Some(model.content.text) case None => read_file_content(node_name) } override def with_thy_reader[A](name: Document.Node.Name, f: Reader[Char] => A): A = { GUI_Thread.now { Document_Model.get(name) match { case Some(model: Buffer_Model) => JEdit_Lib.buffer_lock(model.buffer) { Some(f(JEdit_Lib.buffer_reader(model.buffer))) } case Some(model: File_Model) => Some(f(Scan.char_reader(model.content.text))) case None => None } } getOrElse { if (Url.is_wellformed(name.node)) f(Scan.byte_reader(Url(name.node))) else { val file = new JFile(name.node) if (file.isFile) f(Scan.byte_reader(file)) else error("No such file: " + quote(file.toString)) } } } private class File_Content_Output(buffer: Buffer) extends ByteArrayOutputStream(buffer.getLength + 1) { def content(): Bytes = Bytes(this.buf, 0, this.count) } private class File_Content(buffer: Buffer) extends BufferIORequest(null, buffer, null, VFSManager.getVFSForPath(buffer.getPath), buffer.getPath) { def _run(): Unit = {} def content(): Bytes = { val out = new File_Content_Output(buffer) write(buffer, out) out.content() } } def make_file_content(buffer: Buffer): Bytes = (new File_Content(buffer)).content() /* theory text edits */ override def commit(change: Session.Change): Unit = { if (change.syntax_changed.nonEmpty) GUI_Thread.later { Document_Model.syntax_changed(change.syntax_changed) } if (change.deps_changed || - PIDE.options.bool("jedit_auto_resolve") && undefined_blobs(change.version.nodes).nonEmpty) + PIDE.options.bool("jedit_auto_resolve") && undefined_blobs(change.version).nonEmpty) PIDE.plugin.deps_changed() } } diff --git a/src/Tools/jEdit/src/main_plugin.scala b/src/Tools/jEdit/src/main_plugin.scala --- a/src/Tools/jEdit/src/main_plugin.scala +++ b/src/Tools/jEdit/src/main_plugin.scala @@ -1,477 +1,472 @@ /* Title: Tools/jEdit/src/main_plugin.scala Author: Makarius Main plumbing for PIDE infrastructure as jEdit plugin. */ package isabelle.jedit import isabelle._ import javax.swing.JOptionPane import java.io.{File => JFile} import org.gjt.sp.jedit.{jEdit, EBMessage, EBPlugin, Buffer, View, PerspectiveManager} import org.gjt.sp.jedit.textarea.JEditTextArea import org.gjt.sp.jedit.syntax.ModeProvider import org.gjt.sp.jedit.msg.{EditorStarted, BufferUpdate, EditPaneUpdate, PropertiesChanged, ViewUpdate} import org.gjt.sp.util.SyntaxUtilities import org.gjt.sp.util.Log object PIDE { /* semantic document content */ def maybe_snapshot(view: View = null): Option[Document.Snapshot] = GUI_Thread.now { val buffer = JEdit_Lib.jedit_view(view).getBuffer Document_Model.get(buffer).map(_.snapshot()) } def maybe_rendering(view: View = null): Option[JEdit_Rendering] = GUI_Thread.now { val text_area = JEdit_Lib.jedit_view(view).getTextArea Document_View.get(text_area).map(_.get_rendering()) } def snapshot(view: View = null): Document.Snapshot = maybe_snapshot(view) getOrElse error("No document model for current buffer") def rendering(view: View = null): JEdit_Rendering = maybe_rendering(view) getOrElse error("No document view for current text area") /* plugin instance */ @volatile var _plugin: Main_Plugin = null def plugin: Main_Plugin = if (_plugin == null) error("Uninitialized Isabelle/jEdit plugin") else _plugin def options: JEdit_Options = plugin.options def resources: JEdit_Resources = plugin.resources def session: Session = plugin.session object editor extends JEdit_Editor } class Main_Plugin extends EBPlugin { /* options */ private var _options: JEdit_Options = null private def init_options(): Unit = _options = new JEdit_Options(Options.init()) def options: JEdit_Options = _options /* resources */ private var _resources: JEdit_Resources = null private def init_resources(): Unit = _resources = JEdit_Resources(options.value) def resources: JEdit_Resources = _resources /* session */ private var _session: Session = null private def init_session(): Unit = _session = new Session(options.value, resources) def session: Session = _session /* misc support */ val completion_history = new Completion.History_Variable val spell_checker = new Spell_Checker_Variable /* global changes */ def options_changed(): Unit = { session.global_options.post(Session.Global_Options(options.value)) delay_load.invoke() } def deps_changed(): Unit = { delay_load.invoke() } /* theory files */ lazy val delay_init: Delay = Delay.last(options.seconds("editor_load_delay"), gui = true) { init_models() } private val delay_load_active = Synchronized(false) private def delay_load_finished(): Unit = delay_load_active.change(_ => false) private def delay_load_activated(): Boolean = delay_load_active.guarded_access(a => Some((!a, true))) private def delay_load_body(): Unit = { val required_files = { val models = Document_Model.get_models() val thy_files = resources.resolve_dependencies(models, Nil) val aux_files = if (options.bool("jedit_auto_resolve")) { - val stable_tip_version = - if (models.valuesIterator.forall(_.is_stable)) { - session.get_state().stable_tip_version - } - else None - stable_tip_version match { - case Some(version) => resources.undefined_blobs(version.nodes) + session.stable_tip_version(models) match { + case Some(version) => resources.undefined_blobs(version) case None => delay_load.invoke(); Nil } } else Nil (thy_files ::: aux_files).filterNot(models.isDefinedAt) } if (required_files.nonEmpty) { try { Isabelle_Thread.fork(name = "resolve_dependencies") { val loaded_files = for { name <- required_files text <- resources.read_file_content(name) } yield (name, text) GUI_Thread.later { try { Document_Model.provide_files(session, loaded_files) delay_init.invoke() } finally { delay_load_finished() } } } } catch { case _: Throwable => delay_load_finished() } } else delay_load_finished() } private lazy val delay_load: Delay = Delay.last(options.seconds("editor_load_delay"), gui = true) { if (JEdit_Options.continuous_checking()) { if (!PerspectiveManager.isPerspectiveEnabled || JEdit_Lib.jedit_buffers().exists(_.isLoading)) delay_load.invoke() else if (delay_load_activated()) delay_load_body() else delay_load.invoke() } } private def file_watcher_action(changed: Set[JFile]): Unit = if (Document_Model.sync_files(changed)) PIDE.editor.invoke_generated() lazy val file_watcher: File_Watcher = File_Watcher(file_watcher_action, options.seconds("editor_load_delay")) /* session phase */ val session_phase_changed: Session.Consumer[Session.Phase] = Session.Consumer("Isabelle/jEdit") { case Session.Terminated(result) if !result.ok => GUI_Thread.later { GUI.error_dialog(jEdit.getActiveView, "Prover process terminated with error", "Isabelle Syslog", GUI.scrollable_text(session.syslog.content())) } case Session.Ready if !shutting_down.value => init_models() if (!JEdit_Options.continuous_checking()) { GUI_Thread.later { val answer = GUI.confirm_dialog(jEdit.getActiveView, "Continuous checking of PIDE document", JOptionPane.YES_NO_OPTION, "Continuous checking is presently disabled:", "editor buffers will remain inactive!", "Enable continuous checking now?") if (answer == 0) JEdit_Options.continuous_checking.set() } } delay_load.invoke() case Session.Shutdown => GUI_Thread.later { delay_load.revoke() delay_init.revoke() PIDE.editor.shutdown() exit_models(JEdit_Lib.jedit_buffers().toList) } case _ => } /* document model and view */ def exit_models(buffers: List[Buffer]): Unit = { GUI_Thread.now { buffers.foreach(buffer => JEdit_Lib.buffer_lock(buffer) { JEdit_Lib.jedit_text_areas(buffer).foreach(Document_View.exit) Document_Model.exit(buffer) }) } } def init_models(): Unit = { GUI_Thread.now { PIDE.editor.flush() for { buffer <- JEdit_Lib.jedit_buffers() if buffer != null && !buffer.getBooleanProperty(Buffer.GZIPPED) } { if (buffer.isLoaded) { JEdit_Lib.buffer_lock(buffer) { val node_name = resources.node_name(buffer) val model = Document_Model.init(session, node_name, buffer) for { text_area <- JEdit_Lib.jedit_text_areas(buffer) if !Document_View.get(text_area).map(_.model).contains(model) } Document_View.init(model, text_area) } } else delay_init.invoke() } PIDE.editor.invoke_generated() } } def init_view(buffer: Buffer, text_area: JEditTextArea): Unit = GUI_Thread.now { JEdit_Lib.buffer_lock(buffer) { Document_Model.get(buffer) match { case Some(model) => Document_View.init(model, text_area) case None => } } } def exit_view(buffer: Buffer, text_area: JEditTextArea): Unit = GUI_Thread.now { JEdit_Lib.buffer_lock(buffer) { Document_View.exit(text_area) } } /* main plugin plumbing */ @volatile private var startup_failure: Option[Throwable] = None @volatile private var startup_notified = false private def init_editor(view: View): Unit = { Keymap_Merge.check_dialog(view) Session_Build.check_dialog(view) } private def init_title(view: View): Unit = { val title = proper_string(Isabelle_System.getenv("ISABELLE_IDENTIFIER")).getOrElse("Isabelle") + "/" + PIDE.resources.session_base.session_name val marker = "\u200B" val old_title = view.getViewConfig.title if (old_title == null || old_title.startsWith(marker)) { view.setUserTitle(marker + title) } } override def handleMessage(message: EBMessage): Unit = { GUI_Thread.assert {} if (startup_failure.isDefined && !startup_notified) { message match { case _: EditorStarted => GUI.error_dialog(null, "Isabelle plugin startup failure", GUI.scrollable_text(Exn.message(startup_failure.get)), "Prover IDE inactive!") startup_notified = true case _ => } } if (startup_failure.isEmpty) { message match { case _: EditorStarted => if (resources.session_errors.nonEmpty) { GUI.warning_dialog(jEdit.getActiveView, "Bad session structure: may cause problems with theory imports", GUI.scrollable_text(cat_lines(resources.session_errors))) } jEdit.propertiesChanged() val view = jEdit.getActiveView() init_editor(view) PIDE.editor.hyperlink_position(true, Document.Snapshot.init, JEdit_Sessions.logic_root(options.value)).foreach(_.follow(view)) case msg: ViewUpdate if msg.getWhat == ViewUpdate.CREATED && msg.getView != null => init_title(msg.getView) case msg: BufferUpdate if msg.getWhat == BufferUpdate.LOAD_STARTED || msg.getWhat == BufferUpdate.CLOSING => if (msg.getBuffer != null) { exit_models(List(msg.getBuffer)) PIDE.editor.invoke_generated() } case msg: BufferUpdate if msg.getWhat == BufferUpdate.PROPERTIES_CHANGED || msg.getWhat == BufferUpdate.LOADED => if (session.is_ready) { delay_init.invoke() delay_load.invoke() } case msg: EditPaneUpdate if msg.getWhat == EditPaneUpdate.BUFFER_CHANGING || msg.getWhat == EditPaneUpdate.BUFFER_CHANGED || msg.getWhat == EditPaneUpdate.CREATED || msg.getWhat == EditPaneUpdate.DESTROYED => val edit_pane = msg.getEditPane val buffer = edit_pane.getBuffer val text_area = edit_pane.getTextArea if (buffer != null && text_area != null) { if (msg.getWhat == EditPaneUpdate.BUFFER_CHANGED || msg.getWhat == EditPaneUpdate.CREATED) { if (session.is_ready) init_view(buffer, text_area) } else { Isabelle.dismissed_popups(text_area.getView) exit_view(buffer, text_area) } if (msg.getWhat == EditPaneUpdate.CREATED) Completion_Popup.Text_Area.init(text_area) if (msg.getWhat == EditPaneUpdate.DESTROYED) Completion_Popup.Text_Area.exit(text_area) } case _: PropertiesChanged => for { view <- JEdit_Lib.jedit_views() edit_pane <- JEdit_Lib.jedit_edit_panes(view) } { val buffer = edit_pane.getBuffer val text_area = edit_pane.getTextArea if (buffer != null && text_area != null) init_view(buffer, text_area) } spell_checker.update(options.value) session.update_options(options.value) case _ => } } } /* mode provider */ private var orig_mode_provider: ModeProvider = null private var pide_mode_provider: ModeProvider = null def init_mode_provider(): Unit = { orig_mode_provider = ModeProvider.instance if (orig_mode_provider.isInstanceOf[ModeProvider]) { pide_mode_provider = new Token_Markup.Mode_Provider(orig_mode_provider) ModeProvider.instance = pide_mode_provider } } def exit_mode_provider(): Unit = { if (ModeProvider.instance == pide_mode_provider) ModeProvider.instance = orig_mode_provider } /* HTTP server */ val http_root: String = "/" + UUID.random().toString val http_server: HTTP.Server = HTTP.server(services = Document_Model.Preview_Service :: HTTP.isabelle_services) /* start and stop */ private val shutting_down = Synchronized(false) override def start(): Unit = { /* strict initialization */ init_options() init_resources() init_session() PIDE._plugin = this /* non-strict initialization */ try { completion_history.load() spell_checker.update(options.value) JEdit_Lib.jedit_views().foreach(init_title) Syntax_Style.set_extender(Syntax_Style.Main_Extender) init_mode_provider() JEdit_Lib.jedit_text_areas().foreach(Completion_Popup.Text_Area.init) http_server.start() startup_failure = None } catch { case exn: Throwable => startup_failure = Some(exn) startup_notified = false Log.log(Log.ERROR, this, exn) } shutting_down.change(_ => false) val view = jEdit.getActiveView() if (view != null) init_editor(view) } override def stop(): Unit = { http_server.stop() Syntax_Style.set_extender(Syntax_Style.Base_Extender) exit_mode_provider() JEdit_Lib.jedit_text_areas().foreach(Completion_Popup.Text_Area.exit) if (startup_failure.isEmpty) { options.value.save_prefs() completion_history.value.save() } exit_models(JEdit_Lib.jedit_buffers().toList) shutting_down.change(_ => true) session.stop() file_watcher.shutdown() PIDE.editor.shutdown() Document_Model.reset() } }