diff --git a/src/HOL/Tools/Nitpick/kodkod.scala b/src/HOL/Tools/Nitpick/kodkod.scala --- a/src/HOL/Tools/Nitpick/kodkod.scala +++ b/src/HOL/Tools/Nitpick/kodkod.scala @@ -1,178 +1,178 @@ /* Title: HOL/Tools/Nitpick/kodkod.scala Author: Makarius Scala interface for Kodkod. */ package isabelle.nitpick import isabelle._ import java.util.concurrent.{TimeUnit, LinkedBlockingQueue, ThreadPoolExecutor} import org.antlr.runtime.{ANTLRInputStream, RecognitionException} import de.tum.in.isabelle.Kodkodi.{Context, KodkodiLexer, KodkodiParser} object Kodkod { /** result **/ sealed case class Result(rc: Int, out: String, err: String) { def ok: Boolean = rc == 0 def check: String = if (ok) out else error(if (err.isEmpty) "Error" else err) def encode: XML.Body = { import XML.Encode._ triple(int, string, string)((rc, out, err)) } } /** execute **/ def execute(source: String, solve_all: Boolean = false, prove: Boolean = false, max_solutions: Int = Integer.MAX_VALUE, cleanup_inst: Boolean = false, timeout: Time = Time.zero, max_threads: Int = 0): Result = { /* executor */ val pool_size = if (max_threads == 0) Isabelle_Thread.max_threads() else max_threads val executor: ThreadPoolExecutor = new ThreadPoolExecutor(pool_size, pool_size, 0L, TimeUnit.MILLISECONDS, new LinkedBlockingQueue[Runnable], new ThreadPoolExecutor.CallerRunsPolicy) val executor_killed = Synchronized(false) def executor_kill(): Unit = executor_killed.change(b => if (b) b else { Isabelle_Thread.fork() { executor.shutdownNow() }; true }) /* system context */ class Exit extends Exception("EXIT") class Exec_Context extends Context { private var rc = 0 private val out = new StringBuilder private val err = new StringBuilder def return_code(i: Int): Unit = synchronized { rc = rc max i} override def output(s: String): Unit = synchronized { Exn.Interrupt.expose() out ++= s out += '\n' } override def error(s: String): Unit = synchronized { Exn.Interrupt.expose() err ++= s err += '\n' } override def exit(i: Int): Unit = synchronized { return_code(i) executor_kill() throw new Exit } def result(): Result = synchronized { Result(rc, out.toString, err.toString) } } val context = new Exec_Context /* main */ try { - val lexer = new KodkodiLexer(new ANTLRInputStream(Bytes(source).stream)) + val lexer = new KodkodiLexer(new ANTLRInputStream(Bytes(source).stream())) val parser = KodkodiParser.create(context, executor, false, solve_all, prove, max_solutions, cleanup_inst, lexer) val timeout_request = if (timeout.is_zero) None else { Some(Event_Timer.request(Time.now() + timeout) { context.error("Ran out of time") context.return_code(2) executor_kill() }) } try { parser.problems() } catch { case exn: RecognitionException => parser.reportError(exn) } - timeout_request.foreach(_.cancel) + timeout_request.foreach(_.cancel()) if (parser.getTokenStream.LA(1) != KodkodiParser.EOF) { context.error("Error: trailing tokens") context.exit(1) } if (lexer.getNumberOfSyntaxErrors + parser.getNumberOfSyntaxErrors > 0) { context.exit(1) } } catch { case _: Exit => case exn: Throwable => val message = exn.getMessage context.error(if (message.isEmpty) exn.toString else "Error: " + message) context.return_code(1) } executor.shutdownNow() context.result() } /** protocol handler **/ def warmup(): String = execute( "solver: \"MiniSat\"\n" + File.read(Path.explode("$KODKODI/examples/weber3.kki"))).check class Handler extends Session.Protocol_Handler { override def init(session: Session): Unit = warmup() } /** scala function **/ object Fun extends Scala.Fun("kodkod") { val here = Scala_Project.here def apply(args: String): String = { val (timeout, (solve_all, (max_solutions, (max_threads, kki)))) = { import XML.Decode._ pair(int, pair(bool, pair(int, pair(int, string))))(YXML.parse_body(args)) } val result = execute(kki, solve_all = solve_all, max_solutions = max_solutions, timeout = Time.ms(timeout), max_threads = max_threads) YXML.string_of_body(result.encode) } } } class Scala_Functions extends Scala.Functions(Kodkod.Fun) diff --git a/src/Pure/Concurrent/delay.scala b/src/Pure/Concurrent/delay.scala --- a/src/Pure/Concurrent/delay.scala +++ b/src/Pure/Concurrent/delay.scala @@ -1,66 +1,66 @@ /* Title: Pure/Concurrent/delay.scala Author: Makarius Delayed events. */ package isabelle object Delay { // delayed event after first invocation def first(delay: => Time, log: Logger = No_Logger, gui: Boolean = false)(event: => Unit): Delay = new Delay(true, delay, log, if (gui) GUI_Thread.later { event } else event ) // delayed event after last invocation def last(delay: => Time, log: Logger = No_Logger, gui: Boolean = false)(event: => Unit): Delay = new Delay(false, delay, log, if (gui) GUI_Thread.later { event } else event ) } final class Delay private(first: Boolean, delay: => Time, log: Logger, event: => Unit) { private var running: Option[Event_Timer.Request] = None private def run: Unit = { val do_run = synchronized { if (running.isDefined) { running = None; true } else false } if (do_run) { try { event } catch { case exn: Throwable if !Exn.is_interrupt(exn) => log(Exn.message(exn)); throw exn } } } def invoke(): Unit = synchronized { val new_run = running match { - case Some(request) => if (first) false else { request.cancel; true } + case Some(request) => if (first) false else { request.cancel(); true } case None => true } if (new_run) running = Some(Event_Timer.request(Time.now() + delay)(run)) } def revoke(): Unit = synchronized { running match { - case Some(request) => request.cancel; running = None + case Some(request) => request.cancel(); running = None case None => } } def postpone(alt_delay: Time): Unit = synchronized { running match { case Some(request) => val alt_time = Time.now() + alt_delay - if (request.time < alt_time && request.cancel) { + if (request.time < alt_time && request.cancel()) { running = Some(Event_Timer.request(alt_time)(run)) } case None => } } } diff --git a/src/Pure/Concurrent/event_timer.scala b/src/Pure/Concurrent/event_timer.scala --- a/src/Pure/Concurrent/event_timer.scala +++ b/src/Pure/Concurrent/event_timer.scala @@ -1,34 +1,34 @@ /* Title: Pure/Concurrent/event_timer.scala Author: Makarius Initiate event after given point in time. Note: events are run as synchronized action within a dedicated thread and should finish quickly without further ado. */ package isabelle import java.util.{Timer, TimerTask, Date => JDate} object Event_Timer { private lazy val event_timer = new Timer("event_timer", true) final class Request private[Event_Timer](val time: Time, val repeat: Option[Time], task: TimerTask) { - def cancel: Boolean = task.cancel + def cancel(): Boolean = task.cancel() } def request(time: Time, repeat: Option[Time] = None)(event: => Unit): Request = { val task = new TimerTask { def run: Unit = event } repeat match { case None => event_timer.schedule(task, new JDate(time.ms)) case Some(rep) => event_timer.schedule(task, new JDate(time.ms), rep.ms) } new Request(time, repeat, task) } } diff --git a/src/Pure/Concurrent/future.scala b/src/Pure/Concurrent/future.scala --- a/src/Pure/Concurrent/future.scala +++ b/src/Pure/Concurrent/future.scala @@ -1,161 +1,161 @@ /* Title: Pure/Concurrent/future.scala Author: Makarius Value-oriented parallel execution via futures and promises. */ package isabelle import java.util.concurrent.Callable /* futures and promises */ object Future { def value[A](x: A): Future[A] = new Value_Future(x) def fork[A](body: => A): Future[A] = new Task_Future[A](body) def promise[A]: Promise[A] = new Promise_Future[A] def thread[A]( name: String = "", group: ThreadGroup = Isabelle_Thread.current_thread_group, pri: Int = Thread.NORM_PRIORITY, daemon: Boolean = false, inherit_locals: Boolean = false, uninterruptible: Boolean = false)(body: => A): Future[A] = { new Thread_Future[A](name, group, pri, daemon, inherit_locals, uninterruptible, body) } } trait Future[A] { def peek: Option[Exn.Result[A]] def is_finished: Boolean = peek.isDefined def get_finished: A = { require(is_finished, "future not finished"); Exn.release(peek.get) } def join_result: Exn.Result[A] def join: A = Exn.release(join_result) def map[B](f: A => B): Future[B] = Future.fork { f(join) } - def cancel: Unit + def cancel(): Unit override def toString: String = peek match { case None => "" case Some(Exn.Exn(_)) => "" case Some(Exn.Res(x)) => x.toString } } trait Promise[A] extends Future[A] { def fulfill_result(res: Exn.Result[A]): Unit def fulfill(x: A): Unit } /* value future */ private class Value_Future[A](x: A) extends Future[A] { val peek: Option[Exn.Result[A]] = Some(Exn.Res(x)) def join_result: Exn.Result[A] = peek.get - def cancel: Unit = {} + def cancel(): Unit = {} } /* task future via thread pool */ private class Task_Future[A](body: => A) extends Future[A] { private sealed abstract class Status private case object Ready extends Status private case class Running(thread: Thread) extends Status private case object Terminated extends Status private case class Finished(result: Exn.Result[A]) extends Status private val status = Synchronized[Status](Ready) def peek: Option[Exn.Result[A]] = status.value match { case Finished(result) => Some(result) case _ => None } private def try_run(): Unit = { val do_run = status.change_result { case Ready => (true, Running(Thread.currentThread)) case st => (false, st) } if (do_run) { val result = Exn.capture(body) status.change(_ => Terminated) - status.change(_ => Finished(if (Thread.interrupted) Exn.Exn(Exn.Interrupt()) else result)) + status.change(_ => Finished(if (Thread.interrupted()) Exn.Exn(Exn.Interrupt()) else result)) } } private val task = Isabelle_Thread.pool.submit(new Callable[Unit] { def call = try_run() }) def join_result: Exn.Result[A] = { try_run() status.guarded_access { case st @ Finished(result) => Some((result, st)) case _ => None } } - def cancel = + def cancel(): Unit = { status.change { case Ready => task.cancel(false); Finished(Exn.Exn(Exn.Interrupt())) - case st @ Running(thread) => thread.interrupt; st + case st @ Running(thread) => thread.interrupt(); st case st => st } } } /* promise future */ private class Promise_Future[A] extends Promise[A] { private val state = Synchronized[Option[Exn.Result[A]]](None) def peek: Option[Exn.Result[A]] = state.value def join_result: Exn.Result[A] = state.guarded_access(st => if (st.isEmpty) None else Some((st.get, st))) def fulfill_result(result: Exn.Result[A]): Unit = state.change(st => if (st.isEmpty) Some(result) else throw new IllegalStateException) def fulfill(x: A): Unit = fulfill_result(Exn.Res(x)) - def cancel: Unit = + def cancel(): Unit = state.change(st => if (st.isEmpty) Some(Exn.Exn(Exn.Interrupt())) else st) } /* thread future */ private class Thread_Future[A]( name: String, group: ThreadGroup, pri: Int, daemon: Boolean, inherit_locals: Boolean, uninterruptible: Boolean, body: => A) extends Future[A] { private val result = Future.promise[A] private val thread = Isabelle_Thread.fork(name = name, group = group, pri = pri, daemon = daemon, inherit_locals = inherit_locals, uninterruptible = uninterruptible) { result.fulfill_result(Exn.capture(body)) } def peek: Option[Exn.Result[A]] = result.peek def join_result: Exn.Result[A] = result.join_result - def cancel: Unit = thread.interrupt + def cancel(): Unit = thread.interrupt() } diff --git a/src/Pure/Concurrent/isabelle_thread.scala b/src/Pure/Concurrent/isabelle_thread.scala --- a/src/Pure/Concurrent/isabelle_thread.scala +++ b/src/Pure/Concurrent/isabelle_thread.scala @@ -1,195 +1,195 @@ /* Title: Pure/Concurrent/isabelle_thread.scala Author: Makarius Isabelle-specific thread management. */ package isabelle import java.util.concurrent.{ThreadPoolExecutor, TimeUnit, LinkedBlockingQueue} object Isabelle_Thread { /* self-thread */ def self: Isabelle_Thread = Thread.currentThread match { case thread: Isabelle_Thread => thread case thread => error("Isabelle-specific thread required: " + thread) } def check_self: Boolean = Thread.currentThread.isInstanceOf[Isabelle_Thread] /* create threads */ private val counter = Counter.make() def make_name(name: String = "", base: String = "thread"): String = { val prefix = "Isabelle." val suffix = if (name.nonEmpty) name else base + counter() if (suffix.startsWith(prefix)) suffix else prefix + suffix } def current_thread_group: ThreadGroup = Thread.currentThread.getThreadGroup lazy val worker_thread_group: ThreadGroup = new ThreadGroup(current_thread_group, "Isabelle worker") def create( main: Runnable, name: String = "", group: ThreadGroup = current_thread_group, pri: Int = Thread.NORM_PRIORITY, daemon: Boolean = false, inherit_locals: Boolean = false): Isabelle_Thread = { new Isabelle_Thread(main, name = make_name(name = name), group = group, pri = pri, daemon = daemon, inherit_locals = inherit_locals) } def fork( name: String = "", group: ThreadGroup = current_thread_group, pri: Int = Thread.NORM_PRIORITY, daemon: Boolean = false, inherit_locals: Boolean = false, uninterruptible: Boolean = false)(body: => Unit): Isabelle_Thread = { val main: Runnable = if (uninterruptible) { () => Isabelle_Thread.uninterruptible { body } } else { () => body } val thread = create(main, name = name, group = group, pri = pri, daemon = daemon, inherit_locals = inherit_locals) thread.start thread } /* thread pool */ def max_threads(): Int = { val m = Value.Int.unapply(System.getProperty("isabelle.threads", "0")) getOrElse 0 if (m > 0) m else (Runtime.getRuntime.availableProcessors max 1) min 8 } lazy val pool: ThreadPoolExecutor = { val n = max_threads() val executor = new ThreadPoolExecutor(n, n, 2500L, TimeUnit.MILLISECONDS, new LinkedBlockingQueue[Runnable]) executor.setThreadFactory( create(_, name = make_name(base = "worker"), group = worker_thread_group)) executor } /* interrupt handlers */ object Interrupt_Handler { def apply(handle: Isabelle_Thread => Unit, name: String = "handler"): Interrupt_Handler = new Interrupt_Handler(handle, name) val interruptible: Interrupt_Handler = Interrupt_Handler(_.raise_interrupt, name = "interruptible") val uninterruptible: Interrupt_Handler = Interrupt_Handler(_.postpone_interrupt, name = "uninterruptible") } class Interrupt_Handler private(handle: Isabelle_Thread => Unit, name: String) extends Function[Isabelle_Thread, Unit] { def apply(thread: Isabelle_Thread): Unit = handle(thread) override def toString: String = name } def interrupt_handler[A](handler: Interrupt_Handler)(body: => A): A = if (handler == null) body else self.interrupt_handler(handler)(body) def interrupt_handler[A](handle: Isabelle_Thread => Unit)(body: => A): A = self.interrupt_handler(Interrupt_Handler(handle))(body) def interruptible[A](body: => A): A = interrupt_handler(Interrupt_Handler.interruptible)(body) def uninterruptible[A](body: => A): A = interrupt_handler(Interrupt_Handler.uninterruptible)(body) def try_uninterruptible[A](body: => A): A = if (check_self) interrupt_handler(Interrupt_Handler.uninterruptible)(body) else body } class Isabelle_Thread private(main: Runnable, name: String, group: ThreadGroup, pri: Int, daemon: Boolean, inherit_locals: Boolean) extends Thread(group, null, name, 0L, inherit_locals) { thread => thread.setPriority(pri) thread.setDaemon(daemon) override def run: Unit = main.run() def is_self: Boolean = Thread.currentThread == thread /* interrupt state */ // synchronized, with concurrent changes private var interrupt_postponed: Boolean = false def clear_interrupt: Boolean = synchronized { val was_interrupted = isInterrupted || interrupt_postponed Exn.Interrupt.dispose() interrupt_postponed = false was_interrupted } def raise_interrupt: Unit = synchronized { interrupt_postponed = false super.interrupt() } def postpone_interrupt: Unit = synchronized { interrupt_postponed = true Exn.Interrupt.dispose() } /* interrupt handler */ // non-synchronized, only changed on self-thread @volatile private var handler = Isabelle_Thread.Interrupt_Handler.interruptible - override def interrupt: Unit = handler(thread) + override def interrupt(): Unit = handler(thread) def interrupt_handler[A](new_handler: Isabelle_Thread.Interrupt_Handler)(body: => A): A = if (new_handler == null) body else { require(is_self, "interrupt handler on other thread") val old_handler = handler handler = new_handler try { - if (clear_interrupt) interrupt + if (clear_interrupt) interrupt() body } finally { handler = old_handler - if (clear_interrupt) interrupt + if (clear_interrupt) interrupt() } } } diff --git a/src/Pure/Concurrent/par_list.scala b/src/Pure/Concurrent/par_list.scala --- a/src/Pure/Concurrent/par_list.scala +++ b/src/Pure/Concurrent/par_list.scala @@ -1,63 +1,63 @@ /* Title: Pure/Concurrent/par_list.scala Author: Makarius Parallel list combinators. */ package isabelle object Par_List { def managed_results[A, B](f: A => B, xs: List[A]): List[Exn.Result[B]] = if (xs.isEmpty || xs.tail.isEmpty) xs.map(x => Exn.capture { f(x) }) else { val state = Synchronized[(List[Future[B]], Boolean)]((Nil, false)) def cancel_other(self: Int = -1): Unit = state.change { case (futures, canceled) => if (!canceled) { for ((future, i) <- futures.iterator.zipWithIndex if i != self) - future.cancel + future.cancel() } (futures, true) } try { state.change(_ => (xs.iterator.zipWithIndex.map({ case (x, self) => Future.fork { Exn.capture { f(x) } match { case Exn.Exn(exn) => cancel_other(self); throw exn case Exn.Res(res) => res } } }).toList, false)) state.value._1.map(_.join_result) } finally { cancel_other() } } def map[A, B](f: A => B, xs: List[A]): List[B] = Exn.release_first(managed_results(f, xs)) def get_some[A, B](f: A => Option[B], xs: List[A]): Option[B] = { class Found(val res: B) extends Exception val results = managed_results( (x: A) => f(x) match { case None => () case Some(y) => throw new Found(y) }, xs) results.collectFirst({ case Exn.Exn(found: Found) => found.res }) match { case None => Exn.release_first(results); None case some => some } } def find_some[A](P: A => Boolean, xs: List[A]): Option[A] = get_some((x: A) => if (P(x)) Some(x) else None, xs) def exists[A](P: A => Boolean, xs: List[A]): Boolean = find_some(P, xs).isDefined def forall[A](P: A => Boolean, xs: List[A]): Boolean = !exists((x: A) => !P(x), xs) } diff --git a/src/Pure/GUI/gui.scala b/src/Pure/GUI/gui.scala --- a/src/Pure/GUI/gui.scala +++ b/src/Pure/GUI/gui.scala @@ -1,356 +1,356 @@ /* Title: Pure/GUI/gui.scala Author: Makarius Basic GUI tools (for AWT/Swing). */ package isabelle import java.awt.{Component, Container, Font, Image, Insets, KeyboardFocusManager, Window, Point, Rectangle, Dimension, GraphicsEnvironment, MouseInfo, Toolkit} import java.awt.font.{FontRenderContext, LineMetrics, TextAttribute, TransformAttribute} import java.awt.geom.AffineTransform import javax.swing.{ImageIcon, JButton, JDialog, JFrame, JLabel, JLayeredPane, JOptionPane, JTextField, JWindow, LookAndFeel, UIManager, SwingUtilities} import scala.swing.{ComboBox, ScrollPane, TextArea} import scala.swing.event.SelectionChanged object GUI { /* Swing look-and-feel */ def init_laf(): Unit = com.formdev.flatlaf.FlatLightLaf.install() - def current_laf(): String = UIManager.getLookAndFeel.getClass.getName() + def current_laf: String = UIManager.getLookAndFeel.getClass.getName() - def is_macos_laf(): Boolean = - Platform.is_macos && UIManager.getSystemLookAndFeelClassName() == current_laf() + def is_macos_laf: Boolean = + Platform.is_macos && UIManager.getSystemLookAndFeelClassName() == current_laf class Look_And_Feel(laf: LookAndFeel) extends Isabelle_System.Service { def info: UIManager.LookAndFeelInfo = new UIManager.LookAndFeelInfo(laf.getName, laf.getClass.getName) } lazy val look_and_feels: List[Look_And_Feel] = Isabelle_System.make_services(classOf[Look_And_Feel]) def init_lafs(): Unit = { val old_lafs = Set( "com.sun.java.swing.plaf.motif.MotifLookAndFeel", "com.sun.java.swing.plaf.windows.WindowsClassicLookAndFeel") val lafs = UIManager.getInstalledLookAndFeels().toList .filterNot(info => old_lafs(info.getClassName)) val more_lafs = look_and_feels.map(_.info) UIManager.setInstalledLookAndFeels((more_lafs ::: lafs).toArray) } /* additional look-and-feels */ /* plain focus traversal, notably for text fields */ def plain_focus_traversal(component: Component): Unit = { val dummy_button = new JButton def apply(id: Int): Unit = component.setFocusTraversalKeys(id, dummy_button.getFocusTraversalKeys(id)) apply(KeyboardFocusManager.FORWARD_TRAVERSAL_KEYS) apply(KeyboardFocusManager.BACKWARD_TRAVERSAL_KEYS) } /* simple dialogs */ def scrollable_text(raw_txt: String, width: Int = 60, height: Int = 20, editable: Boolean = false) : ScrollPane = { val txt = Output.clean_yxml(raw_txt) val text = new TextArea(txt) if (width > 0) text.columns = width if (height > 0 && split_lines(txt).length > height) text.rows = height text.editable = editable new ScrollPane(text) } private def simple_dialog(kind: Int, default_title: String, parent: Component, title: String, message: Iterable[Any]): Unit = { GUI_Thread.now { val java_message = message.iterator.map({ case x: scala.swing.Component => x.peer case x => x }). toArray.asInstanceOf[Array[AnyRef]] JOptionPane.showMessageDialog(parent, java_message, if (title == null) default_title else title, kind) } } def dialog(parent: Component, title: String, message: Any*): Unit = simple_dialog(JOptionPane.PLAIN_MESSAGE, null, parent, title, message) def warning_dialog(parent: Component, title: String, message: Any*): Unit = simple_dialog(JOptionPane.WARNING_MESSAGE, "Warning", parent, title, message) def error_dialog(parent: Component, title: String, message: Any*): Unit = simple_dialog(JOptionPane.ERROR_MESSAGE, "Error", parent, title, message) def confirm_dialog(parent: Component, title: String, option_type: Int, message: Any*): Int = GUI_Thread.now { val java_message = message map { case x: scala.swing.Component => x.peer case x => x } JOptionPane.showConfirmDialog(parent, java_message.toArray.asInstanceOf[Array[AnyRef]], title, option_type, JOptionPane.QUESTION_MESSAGE) } /* zoom box */ private val Zoom_Factor = "([0-9]+)%?".r abstract class Zoom_Box extends ComboBox[String]( List("50%", "70%", "85%", "100%", "125%", "150%", "175%", "200%", "300%", "400%")) { def changed: Unit def factor: Int = parse(selection.item) private def parse(text: String): Int = text match { case Zoom_Factor(s) => val i = Integer.parseInt(s) if (10 <= i && i < 1000) i else 100 case _ => 100 } private def print(i: Int): String = i.toString + "%" def set_item(i: Int): Unit = { peer.getEditor match { case null => case editor => editor.setItem(print(i)) } } makeEditable()(c => new ComboBox.BuiltInEditor(c)(text => print(parse(text)), x => x)) peer.getEditor.getEditorComponent match { case text: JTextField => text.setColumns(4) case _ => } selection.index = 3 listenTo(selection) reactions += { case SelectionChanged(_) => changed } } /* tooltip with multi-line support */ def tooltip_lines(text: String): String = if (text == null || text == "") null else "" + HTML.output(text) + "" /* icon */ def isabelle_icon(): ImageIcon = new ImageIcon(getClass.getClassLoader.getResource("isabelle/isabelle_transparent-32.gif")) def isabelle_icons(): List[ImageIcon] = for (icon <- List("isabelle/isabelle_transparent-32.gif", "isabelle/isabelle_transparent.gif")) yield new ImageIcon(getClass.getClassLoader.getResource(icon)) def isabelle_image(): Image = isabelle_icon().getImage /* location within multi-screen environment */ final case class Screen_Location(point: Point, bounds: Rectangle) { def relative(parent: Component, size: Dimension): Point = { val w = size.width val h = size.height val x0 = parent.getLocationOnScreen.x val y0 = parent.getLocationOnScreen.y val x1 = x0 + parent.getWidth - w val y1 = y0 + parent.getHeight - h val x2 = point.x min (bounds.x + bounds.width - w) val y2 = point.y min (bounds.y + bounds.height - h) val location = new Point((x2 min x1) max x0, (y2 min y1) max y0) SwingUtilities.convertPointFromScreen(location, parent) location } } def screen_location(component: Component, point: Point): Screen_Location = { val screen_point = new Point(point.x, point.y) if (component != null) SwingUtilities.convertPointToScreen(screen_point, component) val ge = GraphicsEnvironment.getLocalGraphicsEnvironment val screen_bounds = (for { device <- ge.getScreenDevices.iterator config <- device.getConfigurations.iterator bounds = config.getBounds } yield bounds).find(_.contains(screen_point)) getOrElse ge.getMaximumWindowBounds Screen_Location(screen_point, screen_bounds) } def mouse_location(): Screen_Location = screen_location(null, MouseInfo.getPointerInfo.getLocation) /* screen size */ sealed case class Screen_Size(bounds: Rectangle, insets: Insets) { def full_screen_bounds: Rectangle = if (Platform.is_linux) { // avoid menu bar and docking areas new Rectangle( bounds.x + insets.left, bounds.y + insets.top, bounds.width - insets.left - insets.right, bounds.height - insets.top - insets.bottom) } else if (Platform.is_macos) { // avoid menu bar, but ignore docking areas new Rectangle( bounds.x, bounds.y + insets.top, bounds.width, bounds.height - insets.top) } else bounds } def screen_size(component: Component): Screen_Size = { val config = component.getGraphicsConfiguration val bounds = config.getBounds val insets = Toolkit.getDefaultToolkit.getScreenInsets(config) Screen_Size(bounds, insets) } /* component hierachy */ def get_parent(component: Component): Option[Container] = component.getParent match { case null => None case parent => Some(parent) } def ancestors(component: Component): Iterator[Container] = new Iterator[Container] { private var next_elem = get_parent(component) def hasNext: Boolean = next_elem.isDefined def next(): Container = next_elem match { case Some(parent) => next_elem = get_parent(parent) parent case None => Iterator.empty.next() } } def parent_window(component: Component): Option[Window] = ancestors(component).collectFirst({ case x: Window => x }) def layered_pane(component: Component): Option[JLayeredPane] = parent_window(component) match { case Some(w: JWindow) => Some(w.getLayeredPane) case Some(w: JFrame) => Some(w.getLayeredPane) case Some(w: JDialog) => Some(w.getLayeredPane) case _ => None } def traverse_components(component: Component, apply: Component => Unit): Unit = { def traverse(comp: Component): Unit = { apply(comp) comp match { case cont: Container => for (i <- 0 until cont.getComponentCount) traverse(cont.getComponent(i)) case _ => } } traverse(component) } /* font operations */ def copy_font(font: Font): Font = if (font == null) null else new Font(font.getFamily, font.getStyle, font.getSize) def line_metrics(font: Font): LineMetrics = font.getLineMetrics("", new FontRenderContext(null, false, false)) def transform_font(font: Font, transform: AffineTransform): Font = font.deriveFont(java.util.Map.of(TextAttribute.TRANSFORM, new TransformAttribute(transform))) def font(family: String = Isabelle_Fonts.sans, size: Int = 1, bold: Boolean = false): Font = new Font(family, if (bold) Font.BOLD else Font.PLAIN, size) def label_font(): Font = (new JLabel).getFont /* Isabelle fonts */ def imitate_font(font: Font, family: String = Isabelle_Fonts.sans, scale: Double = 1.0): Font = { val font1 = new Font(family, font.getStyle, font.getSize) val rel_size = line_metrics(font).getHeight.toDouble / line_metrics(font1).getHeight new Font(family, font.getStyle, (scale * rel_size * font.getSize).toInt) } def imitate_font_css(font: Font, family: String = Isabelle_Fonts.sans, scale: Double = 1.0): String = { val font1 = new Font(family, font.getStyle, font.getSize) val rel_size = line_metrics(font).getHeight.toDouble / line_metrics(font1).getHeight "font-family: " + family + "; font-size: " + (scale * rel_size * 100).toInt + "%;" } def use_isabelle_fonts(): Unit = { val default_font = label_font() val ui = UIManager.getDefaults for (prop <- List( "ToggleButton.font", "CheckBoxMenuItem.font", "Label.font", "Menu.font", "MenuItem.font", "PopupMenu.font", "Table.font", "TableHeader.font", "TextArea.font", "TextField.font", "TextPane.font", "ToolTip.font", "Tree.font")) { val font = ui.get(prop) match { case font: Font => font case _ => default_font } ui.put(prop, GUI.imitate_font(font)) } } } class FlatLightLaf extends GUI.Look_And_Feel(new com.formdev.flatlaf.FlatLightLaf) class FlatDarkLaf extends GUI.Look_And_Feel(new com.formdev.flatlaf.FlatDarkLaf) diff --git a/src/Pure/General/exn.scala b/src/Pure/General/exn.scala --- a/src/Pure/General/exn.scala +++ b/src/Pure/General/exn.scala @@ -1,147 +1,147 @@ /* Title: Pure/General/exn.scala Author: Makarius Support for exceptions (arbitrary throwables). */ package isabelle object Exn { /* user errors */ class User_Error(message: String) extends RuntimeException(message) { override def equals(that: Any): Boolean = that match { case other: User_Error => message == other.getMessage case _ => false } override def hashCode: Int = message.hashCode override def toString: String = "\n" + Output.error_message_text(message) } object ERROR { def apply(message: String): User_Error = new User_Error(message) def unapply(exn: Throwable): Option[String] = user_message(exn) } def error(message: String): Nothing = throw ERROR(message) def cat_message(msgs: String*): String = cat_lines(msgs.iterator.filterNot(_ == "")) def cat_error(msgs: String*): Nothing = error(cat_message(msgs:_*)) /* exceptions as values */ sealed abstract class Result[A] { def user_error: Result[A] = this match { case Exn(ERROR(msg)) => Exn(ERROR(msg)) case _ => this } def map[B](f: A => B): Result[B] = this match { case Res(res) => Res(f(res)) case Exn(exn) => Exn(exn) } } case class Res[A](res: A) extends Result[A] case class Exn[A](exn: Throwable) extends Result[A] def capture[A](e: => A): Result[A] = try { Res(e) } catch { case exn: Throwable => Exn[A](exn) } def release[A](result: Result[A]): A = result match { case Res(x) => x case Exn(exn) => throw exn } def release_first[A](results: List[Result[A]]): List[A] = results.find({ case Exn(exn) => !is_interrupt(exn) case _ => false }) match { case Some(Exn(exn)) => throw exn case _ => results.map(release) } /* interrupts */ def is_interrupt(exn: Throwable): Boolean = { var found_interrupt = false var e = exn while (!found_interrupt && e != null) { found_interrupt |= e.isInstanceOf[InterruptedException] e = e.getCause } found_interrupt } def interruptible_capture[A](e: => A): Result[A] = try { Res(e) } catch { case exn: Throwable if !is_interrupt(exn) => Exn[A](exn) } object Interrupt { object ERROR { def unapply(exn: Throwable): Option[String] = if (is_interrupt(exn)) Some(message(exn)) else user_message(exn) } def apply(): Throwable = new InterruptedException("Interrupt") def unapply(exn: Throwable): Boolean = is_interrupt(exn) - def dispose(): Unit = Thread.interrupted - def expose(): Unit = if (Thread.interrupted) throw apply() - def impose(): Unit = Thread.currentThread.interrupt + def dispose(): Unit = Thread.interrupted() + def expose(): Unit = if (Thread.interrupted()) throw apply() + def impose(): Unit = Thread.currentThread.interrupt() val return_code = 130 } /* POSIX return code */ def return_code(exn: Throwable, rc: Int): Int = if (is_interrupt(exn)) Interrupt.return_code else rc /* message */ def user_message(exn: Throwable): Option[String] = if (exn.getClass == classOf[RuntimeException] || exn.getClass == classOf[User_Error]) { Some(proper_string(exn.getMessage) getOrElse "Error") } else if (exn.isInstanceOf[java.sql.SQLException]) { Some(proper_string(exn.getMessage) getOrElse "SQL error") } else if (exn.isInstanceOf[java.io.IOException]) { val msg = exn.getMessage Some(if (msg == null || msg == "") "I/O error" else "I/O error: " + msg) } else if (exn.isInstanceOf[RuntimeException]) Some(exn.toString) else None def message(exn: Throwable): String = user_message(exn) getOrElse (if (is_interrupt(exn)) "Interrupt" else exn.toString) /* trace */ def trace(exn: Throwable): String = exn.getStackTrace.mkString("\n") } diff --git a/src/Pure/General/file.scala b/src/Pure/General/file.scala --- a/src/Pure/General/file.scala +++ b/src/Pure/General/file.scala @@ -1,328 +1,328 @@ /* Title: Pure/General/file.scala Author: Makarius File-system operations. */ package isabelle import java.io.{BufferedWriter, OutputStreamWriter, FileOutputStream, BufferedOutputStream, OutputStream, InputStream, FileInputStream, BufferedInputStream, BufferedReader, InputStreamReader, File => JFile, IOException} import java.nio.file.{StandardOpenOption, Path => JPath, Files, SimpleFileVisitor, FileVisitOption, FileVisitResult} import java.nio.file.attribute.BasicFileAttributes import java.net.{URL, MalformedURLException} import java.util.zip.{GZIPInputStream, GZIPOutputStream} import java.util.regex.Pattern import java.util.EnumSet import org.tukaani.xz.{XZInputStream, XZOutputStream} import scala.collection.mutable import scala.util.matching.Regex object File { /* standard path (Cygwin or Posix) */ def standard_path(path: Path): String = path.expand.implode def standard_path(platform_path: String): String = if (Platform.is_windows) { val Platform_Root = new Regex("(?i)" + Pattern.quote(Isabelle_System.cygwin_root()) + """(?:\\+|\z)(.*)""") val Drive = new Regex("""([a-zA-Z]):\\*(.*)""") platform_path.replace('/', '\\') match { case Platform_Root(rest) => "/" + rest.replace('\\', '/') case Drive(letter, rest) => "/cygdrive/" + Word.lowercase(letter) + (if (rest == "") "" else "/" + rest.replace('\\', '/')) case path => path.replace('\\', '/') } } else platform_path def standard_path(file: JFile): String = standard_path(file.getPath) def standard_url(name: String): String = try { val url = new URL(name) if (url.getProtocol == "file" && Url.is_wellformed_file(name)) standard_path(Url.parse_file(name)) else name } catch { case _: MalformedURLException => standard_path(name) } /* platform path (Windows or Posix) */ private val Cygdrive = new Regex("/cygdrive/([a-zA-Z])($|/.*)") private val Named_Root = new Regex("//+([^/]*)(.*)") def platform_path(standard_path: String): String = if (Platform.is_windows) { val result_path = new StringBuilder val rest = standard_path match { case Cygdrive(drive, rest) => result_path ++= (Word.uppercase(drive) + ":" + JFile.separator) rest case Named_Root(root, rest) => result_path ++= JFile.separator result_path ++= JFile.separator result_path ++= root rest case path if path.startsWith("/") => result_path ++= Isabelle_System.cygwin_root() path case path => path } for (p <- space_explode('/', rest) if p != "") { val len = result_path.length if (len > 0 && result_path(len - 1) != JFile.separatorChar) result_path += JFile.separatorChar result_path ++= p } result_path.toString } else standard_path def platform_path(path: Path): String = platform_path(standard_path(path)) def platform_file(path: Path): JFile = new JFile(platform_path(path)) /* platform files */ def absolute(file: JFile): JFile = file.toPath.toAbsolutePath.normalize.toFile def absolute_name(file: JFile): String = absolute(file).getPath def canonical(file: JFile): JFile = file.getCanonicalFile def canonical_name(file: JFile): String = canonical(file).getPath def path(file: JFile): Path = Path.explode(standard_path(file)) def pwd(): Path = path(Path.current.absolute_file) /* relative paths */ def relative_path(base: Path, other: Path): Option[Path] = { val base_path = base.file.toPath val other_path = other.file.toPath if (other_path.startsWith(base_path)) Some(path(base_path.relativize(other_path).toFile)) else None } /* bash path */ def bash_path(path: Path): String = Bash.string(standard_path(path)) def bash_path(file: JFile): String = Bash.string(standard_path(file)) def bash_platform_path(path: Path): String = Bash.string(platform_path(path)) /* directory entries */ def check_dir(path: Path): Path = if (path.is_dir) path else error("No such directory: " + path) def check_file(path: Path): Path = if (path.is_file) path else error("No such file: " + path) /* directory content */ def read_dir(dir: Path): List[String] = { if (!dir.is_dir) error("No such directory: " + dir.toString) val files = dir.file.listFiles if (files == null) Nil else files.toList.map(_.getName).sorted } def get_dir(dir: Path): String = read_dir(dir).filter(name => (dir + Path.basic(name)).is_dir) match { case List(entry) => entry case dirs => error("Exactly one directory entry expected: " + commas_quote(dirs.sorted)) } def find_files( start: JFile, pred: JFile => Boolean = _ => true, include_dirs: Boolean = false, follow_links: Boolean = false): List[JFile] = { val result = new mutable.ListBuffer[JFile] def check(file: JFile): Unit = if (pred(file)) result += file if (start.isFile) check(start) else if (start.isDirectory) { val options = if (follow_links) EnumSet.of(FileVisitOption.FOLLOW_LINKS) else EnumSet.noneOf(classOf[FileVisitOption]) Files.walkFileTree(start.toPath, options, Integer.MAX_VALUE, new SimpleFileVisitor[JPath] { override def preVisitDirectory(path: JPath, attrs: BasicFileAttributes): FileVisitResult = { if (include_dirs) check(path.toFile) FileVisitResult.CONTINUE } override def visitFile(path: JPath, attrs: BasicFileAttributes): FileVisitResult = { val file = path.toFile if (include_dirs || !file.isDirectory) check(file) FileVisitResult.CONTINUE } } ) } result.toList } /* read */ def read(file: JFile): String = Bytes.read(file).text def read(path: Path): String = read(path.file) def read_stream(reader: BufferedReader): String = { val output = new StringBuilder(100) var c = -1 while ({ c = reader.read; c != -1 }) output += c.toChar - reader.close + reader.close() output.toString } def read_stream(stream: InputStream): String = read_stream(new BufferedReader(new InputStreamReader(stream, UTF8.charset))) def read_gzip(file: JFile): String = read_stream(new GZIPInputStream(new BufferedInputStream(new FileInputStream(file)))) def read_gzip(path: Path): String = read_gzip(path.file) def read_xz(file: JFile): String = read_stream(new XZInputStream(new BufferedInputStream(new FileInputStream(file)))) def read_xz(path: Path): String = read_xz(path.file) /* read lines */ def read_line(reader: BufferedReader): Option[String] = { val line = try { reader.readLine} catch { case _: IOException => null } Option(line).map(Library.trim_line) } def read_lines(reader: BufferedReader, progress: String => Unit): List[String] = { val result = new mutable.ListBuffer[String] var line: Option[String] = None while ({ line = read_line(reader); line.isDefined }) { progress(line.get) result += line.get } - reader.close + reader.close() result.toList } /* write */ def writer(file: JFile): BufferedWriter = new BufferedWriter(new OutputStreamWriter(new FileOutputStream(file), UTF8.charset)) def write_file( file: JFile, text: CharSequence, make_stream: OutputStream => OutputStream): Unit = { val stream = make_stream(new FileOutputStream(file)) using(new BufferedWriter(new OutputStreamWriter(stream, UTF8.charset)))(_.append(text)) } def write(file: JFile, text: CharSequence): Unit = write_file(file, text, s => s) def write(path: Path, text: CharSequence): Unit = write(path.file, text) def write_gzip(file: JFile, text: CharSequence): Unit = write_file(file, text, (s: OutputStream) => new GZIPOutputStream(new BufferedOutputStream(s))) def write_gzip(path: Path, text: CharSequence): Unit = write_gzip(path.file, text) def write_xz(file: JFile, text: CharSequence, options: XZ.Options): Unit = File.write_file(file, text, s => new XZOutputStream(new BufferedOutputStream(s), options)) def write_xz(file: JFile, text: CharSequence): Unit = write_xz(file, text, XZ.options()) def write_xz(path: Path, text: CharSequence, options: XZ.Options): Unit = write_xz(path.file, text, options) def write_xz(path: Path, text: CharSequence): Unit = write_xz(path, text, XZ.options()) def write_backup(path: Path, text: CharSequence): Unit = { if (path.is_file) Isabelle_System.move_file(path, path.backup) write(path, text) } def write_backup2(path: Path, text: CharSequence): Unit = { if (path.is_file) Isabelle_System.move_file(path, path.backup2) write(path, text) } /* change */ def change(file: JFile, f: String => String): Unit = write(file, f(read(file))) def change(path: Path, f: String => String): Unit = write(path, f(read(path))) /* append */ def append(file: JFile, text: CharSequence): Unit = Files.write(file.toPath, UTF8.bytes(text.toString), StandardOpenOption.APPEND, StandardOpenOption.CREATE) def append(path: Path, text: CharSequence): Unit = append(path.file, text) /* eq */ def eq(file1: JFile, file2: JFile): Boolean = try { Files.isSameFile(file1.toPath, file2.toPath) } catch { case ERROR(_) => false } def eq(path1: Path, path2: Path): Boolean = eq(path1.file, path2.file) /* eq_content */ def eq_content(file1: JFile, file2: JFile): Boolean = if (eq(file1, file2)) true else if (file1.length != file2.length) false else Bytes.read(file1) == Bytes.read(file2) def eq_content(path1: Path, path2: Path): Boolean = eq_content(path1.file, path2.file) /* permissions */ def is_executable(path: Path): Boolean = { if (Platform.is_windows) Isabelle_System.bash("test -x " + bash_path(path)).check.ok else path.file.canExecute } def set_executable(path: Path, flag: Boolean): Unit = { if (Platform.is_windows && flag) Isabelle_System.chmod("a+x", path) else if (Platform.is_windows) Isabelle_System.chmod("a-x", path) else path.file.setExecutable(flag, false) } } diff --git a/src/Pure/General/file_watcher.scala b/src/Pure/General/file_watcher.scala --- a/src/Pure/General/file_watcher.scala +++ b/src/Pure/General/file_watcher.scala @@ -1,135 +1,135 @@ /* Title: Pure/General/file_watcher.scala Author: Makarius Watcher for file-system events. */ package isabelle import java.io.{File => JFile} import java.nio.file.FileSystems import java.nio.file.{WatchKey, WatchEvent, Path => JPath} import java.nio.file.StandardWatchEventKinds.{ENTRY_CREATE, ENTRY_DELETE, ENTRY_MODIFY} import scala.jdk.CollectionConverters._ class File_Watcher private[File_Watcher] // dummy template { def register(dir: JFile): Unit = {} def register_parent(file: JFile): Unit = {} def deregister(dir: JFile): Unit = {} def purge(retain: Set[JFile]): Unit = {} def shutdown(): Unit = {} } object File_Watcher { val none: File_Watcher = new File_Watcher { override def toString: String = "File_Watcher.none" } def apply(handle: Set[JFile] => Unit, delay: => Time = Time.seconds(0.5)): File_Watcher = if (Platform.is_windows) none else new Impl(handle, delay) /* proper implementation */ sealed case class State( dirs: Map[JFile, WatchKey] = Map.empty, changed: Set[JFile] = Set.empty) class Impl private[File_Watcher](handle: Set[JFile] => Unit, delay: Time) extends File_Watcher { private val state = Synchronized(File_Watcher.State()) private val watcher = FileSystems.getDefault.newWatchService() override def toString: String = state.value.dirs.keySet.mkString("File_Watcher(", ", ", ")") /* registered directories */ override def register(dir: JFile): Unit = state.change(st => st.dirs.get(dir) match { case Some(key) if key.isValid => st case _ => val key = dir.toPath.register(watcher, ENTRY_CREATE, ENTRY_DELETE, ENTRY_MODIFY) st.copy(dirs = st.dirs + (dir -> key)) }) override def register_parent(file: JFile): Unit = { val dir = file.getParentFile if (dir != null && dir.isDirectory) register(dir) } override def deregister(dir: JFile): Unit = state.change(st => st.dirs.get(dir) match { case None => st case Some(key) => - key.cancel + key.cancel() st.copy(dirs = st.dirs - dir) }) override def purge(retain: Set[JFile]): Unit = state.change(st => st.copy(dirs = st.dirs -- - (for ((dir, key) <- st.dirs.iterator if !retain(dir)) yield { key.cancel; dir }))) + (for ((dir, key) <- st.dirs.iterator if !retain(dir)) yield { key.cancel(); dir }))) /* changed directory entries */ private val delay_changed = Delay.last(delay) { val changed = state.change_result(st => (st.changed, st.copy(changed = Set.empty))) handle(changed) } private val watcher_thread = Isabelle_Thread.fork(name = "file_watcher", daemon = true) { try { while (true) { val key = watcher.take val has_changed = state.change_result(st => { val (remove, changed) = st.dirs.collectFirst({ case (dir, key1) if key == key1 => dir }) match { case Some(dir) => val events: Iterable[WatchEvent[JPath]] = key.pollEvents.asInstanceOf[java.util.List[WatchEvent[JPath]]].asScala val remove = if (key.reset) None else Some(dir) val changed = events.iterator.foldLeft(Set.empty[JFile]) { case (set, event) => set + dir.toPath.resolve(event.context).toFile } (remove, changed) case None => key.pollEvents key.reset (None, Set.empty[JFile]) } (changed.nonEmpty, st.copy(dirs = st.dirs -- remove, changed = st.changed ++ changed)) }) if (has_changed) delay_changed.invoke() } } catch { case Exn.Interrupt() => } } /* shutdown */ override def shutdown(): Unit = { - watcher_thread.interrupt + watcher_thread.interrupt() watcher_thread.join - delay_changed.revoke + delay_changed.revoke() } } } diff --git a/src/Pure/General/http.scala b/src/Pure/General/http.scala --- a/src/Pure/General/http.scala +++ b/src/Pure/General/http.scala @@ -1,166 +1,166 @@ /* Title: Pure/General/http.scala Author: Makarius HTTP server support. */ package isabelle import java.net.{InetAddress, InetSocketAddress, URI} import com.sun.net.httpserver.{HttpExchange, HttpHandler, HttpServer} import scala.collection.immutable.SortedMap object HTTP { /** server **/ /* response */ object Response { def apply( bytes: Bytes = Bytes.empty, content_type: String = "application/octet-stream"): Response = new Response(bytes, content_type) val empty: Response = apply() def text(s: String): Response = apply(Bytes(s), "text/plain; charset=utf-8") def html(s: String): Response = apply(Bytes(s), "text/html; charset=utf-8") } class Response private[HTTP](val bytes: Bytes, val content_type: String) { override def toString: String = bytes.toString } /* exchange */ class Exchange private[HTTP](val http_exchange: HttpExchange) { def request_method: String = http_exchange.getRequestMethod def request_uri: URI = http_exchange.getRequestURI def read_request(): Bytes = using(http_exchange.getRequestBody)(Bytes.read_stream(_)) def write_response(code: Int, response: Response): Unit = { http_exchange.getResponseHeaders().set("Content-Type", response.content_type) http_exchange.sendResponseHeaders(code, response.bytes.length.toLong) using(http_exchange.getResponseBody)(response.bytes.write_stream(_)) } } /* handler */ class Handler private[HTTP](val root: String, val handler: HttpHandler) { override def toString: String = root } def handler(root: String, body: Exchange => Unit): Handler = new Handler(root, new HttpHandler { def handle(x: HttpExchange): Unit = body(new Exchange(x)) }) /* particular methods */ sealed case class Arg(method: String, uri: URI, request: Bytes) { def decode_properties: Properties.T = space_explode('&', request.text).map(s => space_explode('=', s) match { case List(a, b) => Url.decode(a) -> Url.decode(b) case _ => error("Malformed key-value pair in HTTP/POST: " + quote(s)) }) } def method(name: String, root: String, body: Arg => Option[Response]): Handler = handler(root, http => { val request = http.read_request() if (http.request_method == name) { val arg = Arg(name, http.request_uri, request) Exn.capture(body(arg)) match { case Exn.Res(Some(response)) => http.write_response(200, response) case Exn.Res(None) => http.write_response(404, Response.empty) case Exn.Exn(ERROR(msg)) => http.write_response(500, Response.text(Output.error_message_text(msg))) case Exn.Exn(exn) => throw exn } } else http.write_response(400, Response.empty) }) def get(root: String, body: Arg => Option[Response]): Handler = method("GET", root, body) def post(root: String, body: Arg => Option[Response]): Handler = method("POST", root, body) /* server */ class Server private[HTTP](val http_server: HttpServer) { def += (handler: Handler): Unit = http_server.createContext(handler.root, handler.handler) def -= (handler: Handler): Unit = http_server.removeContext(handler.root) - def start: Unit = http_server.start - def stop: Unit = http_server.stop(0) + def start(): Unit = http_server.start + def stop(): Unit = http_server.stop(0) def address: InetSocketAddress = http_server.getAddress def url: String = "http://" + address.getHostName + ":" + address.getPort override def toString: String = url } def server(handlers: List[Handler] = isabelle_resources): Server = { val http_server = HttpServer.create(new InetSocketAddress(isabelle.Server.localhost, 0), 0) http_server.setExecutor(null) val server = new Server(http_server) for (handler <- handlers) server += handler server } /** Isabelle resources **/ lazy val isabelle_resources: List[Handler] = List(welcome, fonts()) /* welcome */ val welcome: Handler = get("/", arg => if (arg.uri.toString == "/") { val id = Isabelle_System.isabelle_id() Some(Response.text("Welcome to Isabelle/" + id + ": " + Distribution.version)) } else None) /* fonts */ private lazy val html_fonts: List[Isabelle_Fonts.Entry] = Isabelle_Fonts.fonts(hidden = true) def fonts(root: String = "/fonts"): Handler = get(root, arg => { val uri_name = arg.uri.toString if (uri_name == root) { Some(Response.text(cat_lines(html_fonts.map(entry => entry.path.file_name)))) } else { html_fonts.collectFirst( { case entry if uri_name == root + "/" + entry.path.file_name => Response(entry.bytes) }) } }) } diff --git a/src/Pure/General/mailman.scala b/src/Pure/General/mailman.scala --- a/src/Pure/General/mailman.scala +++ b/src/Pure/General/mailman.scala @@ -1,71 +1,71 @@ /* Title: Pure/General/mailman.scala Author: Makarius Support for Mailman list servers. */ package isabelle import java.net.URL object Mailman { /* mailing list archives */ def archive(url: URL, name: String = ""): Archive = { val text = Url.read(url) val hrefs = """href="([^"]+\.txt(?:\.gz)?)"""".r.findAllMatchIn(text).map(_.group(1)).toList val title = """The ([^</>]*) Archives""".r.findFirstMatchIn(text).map(_.group(1)) val list_url = Url(Library.take_suffix[Char](_ == '/', Url.trim_index(url).toString.toList)._1.mkString + "/") val list_name = (proper_string(name) orElse title).getOrElse(error("Failed to determine mailing list name")) new Archive(list_url, list_name, hrefs) } class Archive private[Mailman](val list_url: URL, val list_name: String, hrefs: List[String]) { override def toString: String = list_name def download(target_dir: Path, progress: Progress = new Progress): List[Path] = { val dir = target_dir + Path.basic(list_name) Isabelle_System.make_directory(dir) hrefs.flatMap(name => { val path = dir + Path.basic(name) val url = new URL(list_url, name) val connection = url.openConnection try { val length = connection.getContentLengthLong val timestamp = connection.getLastModified val file = path.file if (file.isFile && file.length == length && file.lastModified == timestamp) None else { progress.echo("Getting " + url) val bytes = using(connection.getInputStream)(Bytes.read_stream(_, hint = length.toInt max 1024)) Bytes.write(file, bytes) file.setLastModified(timestamp) Some(path) } } - finally { connection.getInputStream.close } + finally { connection.getInputStream.close() } }) } } /* Isabelle mailing lists */ def isabelle_users: Archive = archive(Url("https://lists.cam.ac.uk/pipermail/cl-isabelle-users"), name = "isabelle-users") def isabelle_dev: Archive = archive(Url("https://mailmanbroy.informatik.tu-muenchen.de/pipermail/isabelle-dev")) } diff --git a/src/Pure/General/scan.scala b/src/Pure/General/scan.scala --- a/src/Pure/General/scan.scala +++ b/src/Pure/General/scan.scala @@ -1,512 +1,512 @@ /* Title: Pure/General/scan.scala Author: Makarius Efficient scanning of keywords and tokens. */ package isabelle import scala.annotation.tailrec import scala.collection.IndexedSeq import scala.util.matching.Regex import scala.util.parsing.input.{OffsetPosition, Position => InputPosition, Reader, CharSequenceReader, PagedSeq} import scala.util.parsing.combinator.RegexParsers import java.io.{File => JFile, BufferedInputStream, FileInputStream, InputStream} import java.net.URL object Scan { /** context of partial line-oriented scans **/ abstract class Line_Context case object Finished extends Line_Context case class Quoted(quote: String) extends Line_Context case object Verbatim extends Line_Context case class Cartouche(depth: Int) extends Line_Context case class Comment_Prefix(symbol: Symbol.Symbol) extends Line_Context case class Cartouche_Comment(depth: Int) extends Line_Context case class Comment(depth: Int) extends Line_Context /** parser combinators **/ object Parsers extends Parsers trait Parsers extends RegexParsers { override val whiteSpace: Regex = "".r /* optional termination */ def opt_term[T](p: => Parser[T]): Parser[Option[T]] = p ^^ (x => Some(x)) | """\z""".r ^^ (_ => None) /* repeated symbols */ def repeated(pred: Symbol.Symbol => Boolean, min_count: Int, max_count: Int): Parser[String] = new Parser[String] { def apply(in: Input) = { val start = in.offset val end = in.source.length val matcher = new Symbol.Matcher(in.source) var i = start var count = 0 var finished = false while (!finished && i < end && count < max_count) { val n = matcher(i, end) val sym = in.source.subSequence(i, i + n).toString if (pred(sym)) { i += n; count += 1 } else finished = true } if (count < min_count) Failure("bad input", in) else Success(in.source.subSequence(start, i).toString, in.drop(i - start)) } }.named("repeated") def one(pred: Symbol.Symbol => Boolean): Parser[String] = repeated(pred, 1, 1) def maybe(pred: Symbol.Symbol => Boolean): Parser[String] = repeated(pred, 0, 1) def many(pred: Symbol.Symbol => Boolean): Parser[String] = repeated(pred, 0, Integer.MAX_VALUE) def many1(pred: Symbol.Symbol => Boolean): Parser[String] = repeated(pred, 1, Integer.MAX_VALUE) /* character */ def character(pred: Char => Boolean): Symbol.Symbol => Boolean = (s: Symbol. Symbol) => s.length == 1 && pred(s.charAt(0)) /* quoted strings */ private def quoted_body(quote: Symbol.Symbol): Parser[String] = { rep(many1(sym => sym != quote && sym != "\\") | "\\" + quote | "\\\\" | ("""\\\d\d\d""".r ^? { case x if x.substring(1, 4).toInt <= 255 => x })) ^^ (_.mkString) } def quoted(quote: Symbol.Symbol): Parser[String] = { quote ~ quoted_body(quote) ~ quote ^^ { case x ~ y ~ z => x + y + z } }.named("quoted") def quoted_content(quote: Symbol.Symbol, source: String): String = { require(parseAll(quoted(quote), source).successful, "no quoted text") val body = source.substring(1, source.length - 1) if (body.exists(_ == '\\')) { val content = rep(many1(sym => sym != quote && sym != "\\") | "\\" ~> (quote | "\\" | """\d\d\d""".r ^^ { case x => x.toInt.toChar.toString })) parseAll(content ^^ (_.mkString), body).get } else body } def quoted_line(quote: Symbol.Symbol, ctxt: Line_Context): Parser[(String, Line_Context)] = { ctxt match { case Finished => quote ~ quoted_body(quote) ~ opt_term(quote) ^^ { case x ~ y ~ Some(z) => (x + y + z, Finished) case x ~ y ~ None => (x + y, Quoted(quote)) } case Quoted(q) if q == quote => quoted_body(quote) ~ opt_term(quote) ^^ { case x ~ Some(y) => (x + y, Finished) case x ~ None => (x, ctxt) } case _ => failure("") } }.named("quoted_line") def recover_quoted(quote: Symbol.Symbol): Parser[String] = quote ~ quoted_body(quote) ^^ { case x ~ y => x + y } /* verbatim text */ private def verbatim_body: Parser[String] = rep(many1(sym => sym != "*") | """\*(?!\})""".r) ^^ (_.mkString) def verbatim: Parser[String] = { "{*" ~ verbatim_body ~ "*}" ^^ { case x ~ y ~ z => x + y + z } }.named("verbatim") def verbatim_content(source: String): String = { require(parseAll(verbatim, source).successful, "no verbatim text") source.substring(2, source.length - 2) } def verbatim_line(ctxt: Line_Context): Parser[(String, Line_Context)] = { ctxt match { case Finished => "{*" ~ verbatim_body ~ opt_term("*}") ^^ { case x ~ y ~ Some(z) => (x + y + z, Finished) case x ~ y ~ None => (x + y, Verbatim) } case Verbatim => verbatim_body ~ opt_term("*}") ^^ { case x ~ Some(y) => (x + y, Finished) case x ~ None => (x, Verbatim) } case _ => failure("") } }.named("verbatim_line") val recover_verbatim: Parser[String] = "{*" ~ verbatim_body ^^ { case x ~ y => x + y } /* nested text cartouches */ def cartouche_depth(depth: Int): Parser[(String, Int)] = new Parser[(String, Int)] { require(depth >= 0, "bad cartouche depth") def apply(in: Input) = { val start = in.offset val end = in.source.length val matcher = new Symbol.Matcher(in.source) var i = start var d = depth var finished = false while (!finished && i < end) { val n = matcher(i, end) val sym = in.source.subSequence(i, i + n).toString if (Symbol.is_open(sym)) { i += n; d += 1 } else if (Symbol.is_close(sym) && d > 0) { i += n; d -= 1; if (d == 0) finished = true } else if (d > 0) i += n else finished = true } if (i == start) Failure("bad input", in) else Success((in.source.subSequence(start, i).toString, d), in.drop(i - start)) } }.named("cartouche_depth") def cartouche: Parser[String] = cartouche_depth(0) ^? { case (x, d) if d == 0 => x } def cartouche_line(ctxt: Line_Context): Parser[(String, Line_Context)] = { def cartouche_context(d: Int): Line_Context = if (d == 0) Finished else Cartouche(d) ctxt match { case Finished => cartouche_depth(0) ^^ { case (c, d) => (c, cartouche_context(d)) } case Cartouche(depth) => cartouche_depth(depth) ^^ { case (c, d) => (c, cartouche_context(d)) } case _ => failure("") } } val recover_cartouche: Parser[String] = cartouche_depth(0) ^^ (_._1) def cartouche_content(source: String): String = { def err(): Nothing = error("Malformed text cartouche: " + quote(source)) val source1 = Library.try_unprefix(Symbol.open_decoded, source) orElse Library.try_unprefix(Symbol.open, source) getOrElse err() Library.try_unsuffix(Symbol.close_decoded, source1) orElse Library.try_unsuffix(Symbol.close, source1) getOrElse err() } /* nested comments */ private def comment_depth(depth: Int): Parser[(String, Int)] = new Parser[(String, Int)] { require(depth >= 0, "bad comment depth") val comment_text: Parser[List[String]] = rep1(many1(sym => sym != "*" && sym != "(") | """\*(?!\))|\((?!\*)""".r) def apply(in: Input) = { var rest = in def try_parse[A](p: Parser[A]): Boolean = { parse(p ^^^ (()), rest) match { case Success(_, next) => { rest = next; true } case _ => false } } var d = depth var finished = false while (!finished) { if (try_parse("(*")) d += 1 else if (d > 0 && try_parse("*)")) { d -= 1; if (d == 0) finished = true } else if (d == 0 || !try_parse(comment_text)) finished = true } if (in.offset < rest.offset) Success((in.source.subSequence(in.offset, rest.offset).toString, d), rest) else Failure("comment expected", in) } }.named("comment_depth") def comment: Parser[String] = comment_depth(0) ^? { case (x, d) if d == 0 => x } def comment_line(ctxt: Line_Context): Parser[(String, Line_Context)] = { val depth = ctxt match { case Finished => 0 case Comment(d) => d case _ => -1 } if (depth >= 0) comment_depth(depth) ^^ { case (x, 0) => (x, Finished) case (x, d) => (x, Comment(d)) } else failure("") } val recover_comment: Parser[String] = comment_depth(0) ^^ (_._1) def comment_content(source: String): String = { require(parseAll(comment, source).successful, "no comment") source.substring(2, source.length - 2) } /* keyword */ def literal(lexicon: Lexicon): Parser[String] = new Parser[String] { def apply(in: Input) = { val result = lexicon.scan(in) if (result.isEmpty) Failure("keyword expected", in) else Success(result, in.drop(result.length)) } }.named("keyword") } /** Lexicon -- position tree **/ object Lexicon { /* representation */ private sealed case class Tree(branches: Map[Char, (String, Tree)]) private val empty_tree = Tree(Map()) val empty: Lexicon = new Lexicon(empty_tree) def apply(elems: String*): Lexicon = empty ++ elems } final class Lexicon private(rep: Lexicon.Tree) { /* auxiliary operations */ private def dest(tree: Lexicon.Tree, result: List[String]): List[String] = tree.branches.toList.foldLeft(result) { case (res, (_, (s, tr))) => if (s.isEmpty) dest(tr, res) else dest(tr, s :: res) } private def lookup(str: CharSequence): Option[(Boolean, Lexicon.Tree)] = { val len = str.length @tailrec def look(tree: Lexicon.Tree, tip: Boolean, i: Int): Option[(Boolean, Lexicon.Tree)] = { if (i < len) { tree.branches.get(str.charAt(i)) match { case Some((s, tr)) => look(tr, s.nonEmpty, i + 1) case None => None } } else Some(tip, tree) } look(rep, false, 0) } def completions(str: CharSequence): List[String] = lookup(str) match { case Some((true, tree)) => dest(tree, List(str.toString)) case Some((false, tree)) => dest(tree, Nil) case None => Nil } /* pseudo Set methods */ def raw_iterator: Iterator[String] = dest(rep, Nil).iterator def iterator: Iterator[String] = dest(rep, Nil).sorted.iterator override def toString: String = iterator.mkString("Lexicon(", ", ", ")") def is_empty: Boolean = rep.branches.isEmpty def contains(elem: String): Boolean = lookup(elem) match { case Some((tip, _)) => tip case _ => false } /* build lexicon */ def + (elem: String): Lexicon = if (contains(elem)) this else { val len = elem.length def extend(tree: Lexicon.Tree, i: Int): Lexicon.Tree = if (i < len) { val c = elem.charAt(i) val end = (i + 1 == len) tree.branches.get(c) match { case Some((s, tr)) => Lexicon.Tree(tree.branches + (c -> (if (end) elem else s, extend(tr, i + 1)))) case None => Lexicon.Tree(tree.branches + (c -> (if (end) elem else "", extend(Lexicon.empty_tree, i + 1)))) } } else tree new Lexicon(extend(rep, 0)) } def ++ (elems: IterableOnce[String]): Lexicon = elems.iterator.foldLeft(this)(_ + _) def ++ (other: Lexicon): Lexicon = if (this eq other) this else if (is_empty) other else this ++ other.raw_iterator def -- (remove: Iterable[String]): Lexicon = if (remove.exists(contains)) Lexicon.empty ++ iterator.filterNot(a => remove.exists(b => a == b)) else this /* scan */ def scan(in: Reader[Char]): String = { val source = in.source val offset = in.offset val len = source.length - offset @tailrec def scan_tree(tree: Lexicon.Tree, result: String, i: Int): String = { if (i < len) { tree.branches.get(source.charAt(offset + i)) match { case Some((s, tr)) => scan_tree(tr, if (s.isEmpty) result else s, i + 1) case None => result } } else result } scan_tree(rep, "", 0) } } /** read stream without decoding: efficient length operation **/ private class Restricted_Seq(seq: IndexedSeq[Char], start: Int, end: Int) extends CharSequence { def charAt(i: Int): Char = if (0 <= i && i < length) seq(start + i) else throw new IndexOutOfBoundsException def length: Int = end - start // avoid expensive seq.length def subSequence(i: Int, j: Int): CharSequence = if (0 <= i && i <= j && j <= length) new Restricted_Seq(seq, start + i, start + j) else throw new IndexOutOfBoundsException override def toString: String = { val buf = new StringBuilder(length) for (offset <- start until end) buf.append(seq(offset)) buf.toString } } abstract class Byte_Reader extends Reader[Char] with AutoCloseable private def make_byte_reader(stream: InputStream, stream_length: Int): Byte_Reader = { val buffered_stream = new BufferedInputStream(stream) val seq = new PagedSeq( (buf: Array[Char], offset: Int, length: Int) => { var i = 0 var c = 0 var eof = false while (!eof && i < length) { c = buffered_stream.read if (c == -1) eof = true else { buf(offset + i) = c.toChar; i += 1 } } if (i > 0) i else -1 }) val restricted_seq = new Restricted_Seq(seq, 0, stream_length) class Paged_Reader(override val offset: Int) extends Byte_Reader { override lazy val source: CharSequence = restricted_seq def first: Char = if (seq.isDefinedAt(offset)) seq(offset) else '\u001a' def rest: Paged_Reader = if (seq.isDefinedAt(offset)) new Paged_Reader(offset + 1) else this def pos: InputPosition = new OffsetPosition(source, offset) def atEnd: Boolean = !seq.isDefinedAt(offset) override def drop(n: Int): Paged_Reader = new Paged_Reader(offset + n) - def close: Unit = buffered_stream.close + def close(): Unit = buffered_stream.close() } new Paged_Reader(0) } def byte_reader(file: JFile): Byte_Reader = make_byte_reader(new FileInputStream(file), file.length.toInt) def byte_reader(url: URL): Byte_Reader = { val connection = url.openConnection val stream = connection.getInputStream val stream_length = connection.getContentLength make_byte_reader(stream, stream_length) } def reader_is_utf8(reader: Reader[Char]): Boolean = reader.isInstanceOf[Byte_Reader] def reader_decode_utf8(is_utf8: Boolean, s: String): String = if (is_utf8) UTF8.decode_permissive(s) else s def reader_decode_utf8(reader: Reader[Char], s: String): String = reader_decode_utf8(reader_is_utf8(reader), s) /* plain text reader */ def char_reader(text: CharSequence): CharSequenceReader = new CharSequenceReader(text) } diff --git a/src/Pure/General/sql.scala b/src/Pure/General/sql.scala --- a/src/Pure/General/sql.scala +++ b/src/Pure/General/sql.scala @@ -1,514 +1,514 @@ /* Title: Pure/General/sql.scala Author: Makarius Support for SQL databases: SQLite and PostgreSQL. */ package isabelle import java.time.OffsetDateTime import java.sql.{DriverManager, Connection, PreparedStatement, ResultSet} import scala.collection.mutable object SQL { /** SQL language **/ type Source = String /* concrete syntax */ def escape_char(c: Char): String = c match { case '\u0000' => "\\0" case '\'' => "\\'" case '\"' => "\\\"" case '\b' => "\\b" case '\n' => "\\n" case '\r' => "\\r" case '\t' => "\\t" case '\u001a' => "\\Z" case '\\' => "\\\\" case _ => c.toString } def string(s: String): Source = s.iterator.map(escape_char).mkString("'", "", "'") def ident(s: String): Source = Long_Name.implode(Long_Name.explode(s).map(a => quote(a.replace("\"", "\"\"")))) def enclose(s: Source): Source = "(" + s + ")" def enclosure(ss: Iterable[Source]): Source = ss.mkString("(", ", ", ")") def select(columns: List[Column] = Nil, distinct: Boolean = false): Source = "SELECT " + (if (distinct) "DISTINCT " else "") + (if (columns.isEmpty) "*" else commas(columns.map(_.ident))) + " FROM " val join_outer: Source = " LEFT OUTER JOIN " val join_inner: Source = " INNER JOIN " def join(outer: Boolean): Source = if (outer) join_outer else join_inner def member(x: Source, set: Iterable[String]): Source = set.iterator.map(a => x + " = " + SQL.string(a)).mkString("(", " OR ", ")") /* types */ object Type extends Enumeration { val Boolean = Value("BOOLEAN") val Int = Value("INTEGER") val Long = Value("BIGINT") val Double = Value("DOUBLE PRECISION") val String = Value("TEXT") val Bytes = Value("BLOB") val Date = Value("TIMESTAMP WITH TIME ZONE") } def sql_type_default(T: Type.Value): Source = T.toString def sql_type_sqlite(T: Type.Value): Source = if (T == Type.Boolean) "INTEGER" else if (T == Type.Date) "TEXT" else sql_type_default(T) def sql_type_postgresql(T: Type.Value): Source = if (T == Type.Bytes) "BYTEA" else sql_type_default(T) /* columns */ object Column { def bool(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = Column(name, Type.Boolean, strict, primary_key) def int(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = Column(name, Type.Int, strict, primary_key) def long(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = Column(name, Type.Long, strict, primary_key) def double(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = Column(name, Type.Double, strict, primary_key) def string(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = Column(name, Type.String, strict, primary_key) def bytes(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = Column(name, Type.Bytes, strict, primary_key) def date(name: String, strict: Boolean = false, primary_key: Boolean = false): Column = Column(name, Type.Date, strict, primary_key) } sealed case class Column( name: String, T: Type.Value, strict: Boolean = false, primary_key: Boolean = false, expr: SQL.Source = "") { def make_primary_key: Column = copy(primary_key = true) def apply(table: Table): Column = Column(Long_Name.qualify(table.name, name), T, strict = strict, primary_key = primary_key) def ident: Source = if (expr == "") SQL.ident(name) else enclose(expr) + " AS " + SQL.ident(name) def decl(sql_type: Type.Value => Source): Source = ident + " " + sql_type(T) + (if (strict || primary_key) " NOT NULL" else "") def defined: String = ident + " IS NOT NULL" def undefined: String = ident + " IS NULL" def equal(s: String): Source = ident + " = " + string(s) def where_equal(s: String): Source = "WHERE " + equal(s) override def toString: Source = ident } /* tables */ sealed case class Table(name: String, columns: List[Column], body: Source = "") { private val columns_index: Map[String, Int] = columns.iterator.map(_.name).zipWithIndex.toMap Library.duplicates(columns.map(_.name)) match { case Nil => case bad => error("Duplicate column names " + commas_quote(bad) + " for table " + quote(name)) } def ident: Source = SQL.ident(name) def query: Source = if (body == "") error("Missing SQL body for table " + quote(name)) else SQL.enclose(body) def query_named: Source = query + " AS " + SQL.ident(name) def create(strict: Boolean = false, sql_type: Type.Value => Source): Source = { val primary_key = columns.filter(_.primary_key).map(_.name) match { case Nil => Nil case keys => List("PRIMARY KEY " + enclosure(keys)) } "CREATE TABLE " + (if (strict) "" else "IF NOT EXISTS ") + ident + " " + enclosure(columns.map(_.decl(sql_type)) ::: primary_key) } def create_index(index_name: String, index_columns: List[Column], strict: Boolean = false, unique: Boolean = false): Source = "CREATE " + (if (unique) "UNIQUE " else "") + "INDEX " + (if (strict) "" else "IF NOT EXISTS ") + SQL.ident(index_name) + " ON " + ident + " " + enclosure(index_columns.map(_.name)) def insert_cmd(cmd: Source, sql: Source = ""): Source = cmd + " INTO " + ident + " VALUES " + enclosure(columns.map(_ => "?")) + (if (sql == "") "" else " " + sql) def insert(sql: Source = ""): Source = insert_cmd("INSERT", sql) def delete(sql: Source = ""): Source = "DELETE FROM " + ident + (if (sql == "") "" else " " + sql) def select( select_columns: List[Column] = Nil, sql: Source = "", distinct: Boolean = false): Source = SQL.select(select_columns, distinct = distinct) + ident + (if (sql == "") "" else " " + sql) override def toString: Source = ident } /** SQL database operations **/ /* statements */ class Statement private[SQL](val db: Database, val rep: PreparedStatement) extends AutoCloseable { stmt => object bool { def update(i: Int, x: Boolean): Unit = rep.setBoolean(i, x) def update(i: Int, x: Option[Boolean]): Unit = { if (x.isDefined) update(i, x.get) else rep.setNull(i, java.sql.Types.BOOLEAN) } } object int { def update(i: Int, x: Int): Unit = rep.setInt(i, x) def update(i: Int, x: Option[Int]): Unit = { if (x.isDefined) update(i, x.get) else rep.setNull(i, java.sql.Types.INTEGER) } } object long { def update(i: Int, x: Long): Unit = rep.setLong(i, x) def update(i: Int, x: Option[Long]): Unit = { if (x.isDefined) update(i, x.get) else rep.setNull(i, java.sql.Types.BIGINT) } } object double { def update(i: Int, x: Double): Unit = rep.setDouble(i, x) def update(i: Int, x: Option[Double]): Unit = { if (x.isDefined) update(i, x.get) else rep.setNull(i, java.sql.Types.DOUBLE) } } object string { def update(i: Int, x: String): Unit = rep.setString(i, x) def update(i: Int, x: Option[String]): Unit = update(i, x.orNull) } object bytes { def update(i: Int, bytes: Bytes): Unit = { if (bytes == null) rep.setBytes(i, null) else rep.setBinaryStream(i, bytes.stream(), bytes.length) } def update(i: Int, bytes: Option[Bytes]): Unit = update(i, bytes.orNull) } object date { def update(i: Int, date: Date): Unit = db.update_date(stmt, i, date) def update(i: Int, date: Option[Date]): Unit = update(i, date.orNull) } def execute(): Boolean = rep.execute() def execute_query(): Result = new Result(this, rep.executeQuery()) - def close(): Unit = rep.close + def close(): Unit = rep.close() } /* results */ class Result private[SQL](val stmt: Statement, val rep: ResultSet) { res => def next(): Boolean = rep.next() def iterator[A](get: Result => A): Iterator[A] = new Iterator[A] { private var _next: Boolean = res.next() def hasNext: Boolean = _next def next(): A = { val x = get(res); _next = res.next(); x } } def bool(column: Column): Boolean = rep.getBoolean(column.name) def int(column: Column): Int = rep.getInt(column.name) def long(column: Column): Long = rep.getLong(column.name) def double(column: Column): Double = rep.getDouble(column.name) def string(column: Column): String = { val s = rep.getString(column.name) if (s == null) "" else s } def bytes(column: Column): Bytes = { val bs = rep.getBytes(column.name) if (bs == null) Bytes.empty else Bytes(bs) } def date(column: Column): Date = stmt.db.date(res, column) def timing(c1: Column, c2: Column, c3: Column): Timing = Timing(Time.ms(long(c1)), Time.ms(long(c2)), Time.ms(long(c3))) def get[A](column: Column, f: Column => A): Option[A] = { val x = f(column) if (rep.wasNull) None else Some(x) } def get_bool(column: Column): Option[Boolean] = get(column, bool) def get_int(column: Column): Option[Int] = get(column, int) def get_long(column: Column): Option[Long] = get(column, long) def get_double(column: Column): Option[Double] = get(column, double) def get_string(column: Column): Option[String] = get(column, string) def get_bytes(column: Column): Option[Bytes] = get(column, bytes) def get_date(column: Column): Option[Date] = get(column, date) } /* database */ trait Database extends AutoCloseable { db => /* types */ def sql_type(T: Type.Value): Source /* connection */ def connection: Connection - def close(): Unit = connection.close + def close(): Unit = connection.close() def transaction[A](body: => A): A = { val auto_commit = connection.getAutoCommit try { connection.setAutoCommit(false) val savepoint = connection.setSavepoint try { val result = body connection.commit result } catch { case exn: Throwable => connection.rollback(savepoint); throw exn } } finally { connection.setAutoCommit(auto_commit) } } /* statements and results */ def statement(sql: Source): Statement = new Statement(db, connection.prepareStatement(sql)) def using_statement[A](sql: Source)(f: Statement => A): A = using(statement(sql))(f) def update_date(stmt: Statement, i: Int, date: Date): Unit def date(res: Result, column: Column): Date def insert_permissive(table: Table, sql: Source = ""): Source /* tables and views */ def tables: List[String] = { val result = new mutable.ListBuffer[String] val rs = connection.getMetaData.getTables(null, null, "%", null) while (rs.next) { result += rs.getString(3) } result.toList } def create_table(table: Table, strict: Boolean = false, sql: Source = ""): Unit = using_statement( table.create(strict, sql_type) + (if (sql == "") "" else " " + sql))(_.execute()) def create_index(table: Table, name: String, columns: List[Column], strict: Boolean = false, unique: Boolean = false): Unit = using_statement(table.create_index(name, columns, strict, unique))(_.execute()) def create_view(table: Table, strict: Boolean = false): Unit = { if (strict || !tables.contains(table.name)) { val sql = "CREATE VIEW " + table + " AS " + { table.query; table.body } using_statement(sql)(_.execute()) } } } } /** SQLite **/ object SQLite { // see https://www.sqlite.org/lang_datefunc.html val date_format: Date.Format = Date.Format("uuuu-MM-dd HH:mm:ss.SSS x") lazy val init_jdbc: Unit = { val lib_path = Path.explode("$ISABELLE_SQLITE_HOME/" + Platform.jvm_platform) val lib_name = File.find_files(lib_path.file) match { case List(file) => file.getName case _ => error("Exactly one file expected in directory " + lib_path.expand) } System.setProperty("org.sqlite.lib.path", File.platform_path(lib_path)) System.setProperty("org.sqlite.lib.name", lib_name) Class.forName("org.sqlite.JDBC") } def open_database(path: Path): Database = { init_jdbc val path0 = path.expand val s0 = File.platform_path(path0) val s1 = if (Platform.is_windows) s0.replace('\\', '/') else s0 val connection = DriverManager.getConnection("jdbc:sqlite:" + s1) new Database(path0.toString, connection) } class Database private[SQLite](name: String, val connection: Connection) extends SQL.Database { override def toString: String = name def sql_type(T: SQL.Type.Value): SQL.Source = SQL.sql_type_sqlite(T) def update_date(stmt: SQL.Statement, i: Int, date: Date): Unit = if (date == null) stmt.string(i) = (null: String) else stmt.string(i) = date_format(date) def date(res: SQL.Result, column: SQL.Column): Date = date_format.parse(res.string(column)) def insert_permissive(table: SQL.Table, sql: SQL.Source = ""): SQL.Source = table.insert_cmd("INSERT OR IGNORE", sql = sql) def rebuild: Unit = using_statement("VACUUM")(_.execute()) } } /** PostgreSQL **/ object PostgreSQL { val default_port = 5432 lazy val init_jdbc: Unit = Class.forName("org.postgresql.Driver") def open_database( user: String, password: String, database: String = "", host: String = "", port: Int = 0, ssh: Option[SSH.Session] = None, ssh_close: Boolean = false): Database = { init_jdbc if (user == "") error("Undefined database user") val db_host = proper_string(host) getOrElse "localhost" val db_port = if (port > 0 && port != default_port) ":" + port else "" val db_name = "/" + (proper_string(database) getOrElse user) val (url, name, port_forwarding) = ssh match { case None => val spec = db_host + db_port + db_name val url = "jdbc:postgresql://" + spec val name = user + "@" + spec (url, name, None) case Some(ssh) => val fw = ssh.port_forwarding(remote_host = db_host, remote_port = if (port > 0) port else default_port, ssh_close = ssh_close) val url = "jdbc:postgresql://localhost:" + fw.local_port + db_name val name = user + "@" + fw + db_name + " via ssh " + ssh (url, name, Some(fw)) } try { val connection = DriverManager.getConnection(url, user, password) new Database(name, connection, port_forwarding) } - catch { case exn: Throwable => port_forwarding.foreach(_.close); throw exn } + catch { case exn: Throwable => port_forwarding.foreach(_.close()); throw exn } } class Database private[PostgreSQL]( name: String, val connection: Connection, port_forwarding: Option[SSH.Port_Forwarding]) extends SQL.Database { override def toString: String = name def sql_type(T: SQL.Type.Value): SQL.Source = SQL.sql_type_postgresql(T) // see https://jdbc.postgresql.org/documentation/head/8-date-time.html def update_date(stmt: SQL.Statement, i: Int, date: Date): Unit = if (date == null) stmt.rep.setObject(i, null) else stmt.rep.setObject(i, OffsetDateTime.from(date.to(Date.timezone_utc).rep)) def date(res: SQL.Result, column: SQL.Column): Date = { val obj = res.rep.getObject(column.name, classOf[OffsetDateTime]) if (obj == null) null else Date.instant(obj.toInstant) } def insert_permissive(table: SQL.Table, sql: SQL.Source = ""): SQL.Source = table.insert_cmd("INSERT", sql = sql + (if (sql == "") "" else " ") + "ON CONFLICT DO NOTHING") - override def close(): Unit = { super.close; port_forwarding.foreach(_.close) } + override def close(): Unit = { super.close(); port_forwarding.foreach(_.close()) } } } diff --git a/src/Pure/General/ssh.scala b/src/Pure/General/ssh.scala --- a/src/Pure/General/ssh.scala +++ b/src/Pure/General/ssh.scala @@ -1,513 +1,513 @@ /* Title: Pure/General/ssh.scala Author: Makarius SSH client based on JSch (see also http://www.jcraft.com/jsch/examples). */ package isabelle import java.io.{InputStream, OutputStream, ByteArrayOutputStream} import scala.collection.mutable import scala.util.matching.Regex import com.jcraft.jsch.{JSch, Logger => JSch_Logger, Session => JSch_Session, SftpException, OpenSSHConfig, UserInfo, Channel => JSch_Channel, ChannelExec, ChannelSftp, SftpATTRS, JSchException} object SSH { /* target machine: user@host syntax */ object Target { val User_Host: Regex = "^([^@]+)@(.+)$".r def parse(s: String): (String, String) = s match { case User_Host(user, host) => (user, host) case _ => ("", s) } def unapplySeq(s: String): Option[List[String]] = parse(s) match { case (_, "") => None case (user, host) => Some(List(user, host)) } } val default_port = 22 def make_port(port: Int): Int = if (port > 0) port else default_port def port_suffix(port: Int): String = if (port == default_port) "" else ":" + port def user_prefix(user: String): String = proper_string(user) match { case None => "" case Some(name) => name + "@" } def connect_timeout(options: Options): Int = options.seconds("ssh_connect_timeout").ms.toInt def alive_interval(options: Options): Int = options.seconds("ssh_alive_interval").ms.toInt def alive_count_max(options: Options): Int = options.int("ssh_alive_count_max") /* init context */ def init_context(options: Options): Context = { val config_dir = Path.explode(options.string("ssh_config_dir")) if (!config_dir.is_dir) error("Bad ssh config directory: " + config_dir) val jsch = new JSch val config_file = Path.explode(options.string("ssh_config_file")) if (config_file.is_file) jsch.setConfigRepository(OpenSSHConfig.parseFile(File.platform_path(config_file))) val known_hosts = config_dir + Path.explode("known_hosts") if (!known_hosts.is_file) known_hosts.file.createNewFile jsch.setKnownHosts(File.platform_path(known_hosts)) val identity_files = space_explode(':', options.string("ssh_identity_files")).map(Path.explode) for (identity_file <- identity_files if identity_file.is_file) { try { jsch.addIdentity(File.platform_path(identity_file)) } catch { case exn: JSchException => error("Error in ssh identity file " + identity_file + ": " + exn.getMessage) } } new Context(options, jsch) } def open_session(options: Options, host: String, user: String = "", port: Int = 0, actual_host: String = "", proxy_host: String = "", proxy_user: String = "", proxy_port: Int = 0, permissive: Boolean = false): Session = init_context(options).open_session( host = host, user = user, port = port, actual_host = actual_host, proxy_host = proxy_host, proxy_user = proxy_user, proxy_port = proxy_port, permissive = permissive) class Context private[SSH](val options: Options, val jsch: JSch) { def update_options(new_options: Options): Context = new Context(new_options, jsch) private def connect_session(host: String, user: String = "", port: Int = 0, host_key_permissive: Boolean = false, nominal_host: String = "", nominal_user: String = "", on_close: () => Unit = () => ()): Session = { val session = jsch.getSession(proper_string(user).orNull, host, make_port(port)) session.setUserInfo(No_User_Info) session.setServerAliveInterval(alive_interval(options)) session.setServerAliveCountMax(alive_count_max(options)) session.setConfig("MaxAuthTries", "3") if (host_key_permissive) session.setConfig("StrictHostKeyChecking", "no") if (nominal_host != "") session.setHostKeyAlias(nominal_host) if (options.bool("ssh_compression")) { session.setConfig("compression.s2c", "zlib@openssh.com,zlib,none") session.setConfig("compression.c2s", "zlib@openssh.com,zlib,none") session.setConfig("compression_level", "9") } session.connect(connect_timeout(options)) new Session(options, session, on_close, proper_string(nominal_host) getOrElse host, proper_string(nominal_user) getOrElse user) } def open_session( host: String, user: String = "", port: Int = 0, actual_host: String = "", proxy_host: String = "", proxy_user: String = "", proxy_port: Int = 0, permissive: Boolean = false): Session = { val connect_host = proper_string(actual_host) getOrElse host if (proxy_host == "") connect_session(host = connect_host, user = user, port = port) else { val proxy = connect_session(host = proxy_host, port = proxy_port, user = proxy_user) val fw = try { proxy.port_forwarding(remote_host = connect_host, remote_port = make_port(port)) } - catch { case exn: Throwable => proxy.close; throw exn } + catch { case exn: Throwable => proxy.close(); throw exn } try { connect_session(host = fw.local_host, port = fw.local_port, host_key_permissive = permissive, nominal_host = host, nominal_user = user, user = user, - on_close = () => { fw.close; proxy.close }) + on_close = () => { fw.close(); proxy.close() }) } - catch { case exn: Throwable => fw.close; proxy.close; throw exn } + catch { case exn: Throwable => fw.close(); proxy.close(); throw exn } } } } /* logging */ def logging(verbose: Boolean = true, debug: Boolean = false): Unit = { JSch.setLogger(if (verbose) new Logger(debug) else null) } private class Logger(debug: Boolean) extends JSch_Logger { def isEnabled(level: Int): Boolean = level != JSch_Logger.DEBUG || debug def log(level: Int, msg: String): Unit = { level match { case JSch_Logger.ERROR | JSch_Logger.FATAL => Output.error_message(msg) case JSch_Logger.WARN => Output.warning(msg) case _ => Output.writeln(msg) } } } /* user info */ object No_User_Info extends UserInfo { def getPassphrase: String = null def getPassword: String = null def promptPassword(msg: String): Boolean = false def promptPassphrase(msg: String): Boolean = false def promptYesNo(msg: String): Boolean = false def showMessage(msg: String): Unit = Output.writeln(msg) } /* port forwarding */ object Port_Forwarding { def open(ssh: Session, ssh_close: Boolean, local_host: String, local_port: Int, remote_host: String, remote_port: Int): Port_Forwarding = { val port = ssh.session.setPortForwardingL(local_host, local_port, remote_host, remote_port) new Port_Forwarding(ssh, ssh_close, local_host, port, remote_host, remote_port) } } class Port_Forwarding private[SSH]( ssh: SSH.Session, ssh_close: Boolean, val local_host: String, val local_port: Int, val remote_host: String, val remote_port: Int) extends AutoCloseable { override def toString: String = local_host + ":" + local_port + ":" + remote_host + ":" + remote_port def close(): Unit = { ssh.session.delPortForwardingL(local_host, local_port) if (ssh_close) ssh.close() } } /* Sftp channel */ type Attrs = SftpATTRS sealed case class Dir_Entry(name: String, is_dir: Boolean) { def is_file: Boolean = !is_dir } /* exec channel */ private val exec_wait_delay = Time.seconds(0.3) class Exec private[SSH](session: Session, channel: ChannelExec) extends AutoCloseable { override def toString: String = "exec " + session.toString def close(): Unit = channel.disconnect val exit_status: Future[Int] = Future.thread("ssh_wait") { while (!channel.isClosed) exec_wait_delay.sleep channel.getExitStatus } val stdin: OutputStream = channel.getOutputStream val stdout: InputStream = channel.getInputStream val stderr: InputStream = channel.getErrStream // connect after preparing streams channel.connect(connect_timeout(session.options)) def result( progress_stdout: String => Unit = (_: String) => (), progress_stderr: String => Unit = (_: String) => (), strict: Boolean = true): Process_Result = { - stdin.close + stdin.close() def read_lines(stream: InputStream, progress: String => Unit): List[String] = { val result = new mutable.ListBuffer[String] val line_buffer = new ByteArrayOutputStream(100) def line_flush(): Unit = { val line = Library.trim_line(line_buffer.toString(UTF8.charset_name)) progress(line) result += line line_buffer.reset } var c = 0 var finished = false while (!finished) { while ({ c = stream.read; c != -1 && c != 10 }) line_buffer.write(c) if (c == 10) line_flush() else if (channel.isClosed) { if (line_buffer.size > 0) line_flush() finished = true } else exec_wait_delay.sleep } result.toList } val out_lines = Future.thread("ssh_stdout") { read_lines(stdout, progress_stdout) } val err_lines = Future.thread("ssh_stderr") { read_lines(stderr, progress_stderr) } def terminate(): Unit = { - close + close() out_lines.join err_lines.join exit_status.join } val rc = try { exit_status.join } catch { case Exn.Interrupt() => terminate(); Exn.Interrupt.return_code } - close + close() if (strict && rc == Exn.Interrupt.return_code) throw Exn.Interrupt() Process_Result(rc, out_lines.join, err_lines.join) } } /* session */ class Session private[SSH]( val options: Options, val session: JSch_Session, on_close: () => Unit, val nominal_host: String, val nominal_user: String) extends System with AutoCloseable { def update_options(new_options: Options): Session = new Session(new_options, session, on_close, nominal_host, nominal_user) def host: String = if (session.getHost == null) "" else session.getHost override def hg_url: String = "ssh://" + user_prefix(nominal_user) + nominal_host + "/" override def toString: String = user_prefix(session.getUserName) + host + port_suffix(session.getPort) + (if (session.isConnected) "" else " (disconnected)") /* port forwarding */ def port_forwarding( remote_port: Int, remote_host: String = "localhost", local_port: Int = 0, local_host: String = "localhost", ssh_close: Boolean = false): Port_Forwarding = Port_Forwarding.open(this, ssh_close, local_host, local_port, remote_host, remote_port) /* sftp channel */ val sftp: ChannelSftp = session.openChannel("sftp").asInstanceOf[ChannelSftp] sftp.connect(connect_timeout(options)) def close(): Unit = { sftp.disconnect; session.disconnect; on_close() } val settings: Map[String, String] = { val home = sftp.getHome Map("HOME" -> home, "USER_HOME" -> home) } override def expand_path(path: Path): Path = path.expand_env(settings) def remote_path(path: Path): String = expand_path(path).implode override def bash_path(path: Path): String = Bash.string(remote_path(path)) def chmod(permissions: Int, path: Path): Unit = sftp.chmod(permissions, remote_path(path)) def mv(path1: Path, path2: Path): Unit = sftp.rename(remote_path(path1), remote_path(path2)) def rm(path: Path): Unit = sftp.rm(remote_path(path)) def mkdir(path: Path): Unit = sftp.mkdir(remote_path(path)) def rmdir(path: Path): Unit = sftp.rmdir(remote_path(path)) private def test_entry(path: Path, as_dir: Boolean): Boolean = try { val is_dir = sftp.stat(remote_path(path)).isDir if (as_dir) is_dir else !is_dir } catch { case _: SftpException => false } override def is_dir(path: Path): Boolean = test_entry(path, true) override def is_file(path: Path): Boolean = test_entry(path, false) def is_link(path: Path): Boolean = try { sftp.lstat(remote_path(path)).isLink } catch { case _: SftpException => false } override def make_directory(path: Path): Path = { if (!is_dir(path)) { execute( "perl -e \"use File::Path make_path; make_path('" + remote_path(path) + "');\"") if (!is_dir(path)) error("Failed to create directory: " + quote(remote_path(path))) } path } def read_dir(path: Path): List[Dir_Entry] = { if (!is_dir(path)) error("No such directory: " + path.toString) val dir_name = remote_path(path) val dir = sftp.ls(dir_name) (for { i <- (0 until dir.size).iterator a = dir.get(i).asInstanceOf[AnyRef] name = Untyped.get[String](a, "filename") attrs = Untyped.get[Attrs](a, "attrs") if name != "." && name != ".." } yield { Dir_Entry(name, if (attrs.isLink) { try { sftp.stat(dir_name + "/" + name).isDir } catch { case _: SftpException => false } } else attrs.isDir) }).toList.sortBy(_.name) } def find_files( start: Path, pred: Path => Boolean = _ => true, include_dirs: Boolean = false, follow_links: Boolean = false): List[Path] = { val result = new mutable.ListBuffer[Path] def check(path: Path): Unit = { if (pred(path)) result += path } def find(dir: Path): Unit = { if (include_dirs) check(dir) if (follow_links || !is_link(dir)) { for (entry <- read_dir(dir)) { val path = dir + Path.basic(entry.name) if (entry.is_file) check(path) else find(path) } } } if (is_file(start)) check(start) else find(start) result.toList } def open_input(path: Path): InputStream = sftp.get(remote_path(path)) def open_output(path: Path): OutputStream = sftp.put(remote_path(path)) def read_file(path: Path, local_path: Path): Unit = sftp.get(remote_path(path), File.platform_path(local_path)) def read_bytes(path: Path): Bytes = using(open_input(path))(Bytes.read_stream(_)) def read(path: Path): String = using(open_input(path))(File.read_stream) def write_file(path: Path, local_path: Path): Unit = sftp.put(File.platform_path(local_path), remote_path(path)) def write_bytes(path: Path, bytes: Bytes): Unit = using(open_output(path))(bytes.write_stream) def write(path: Path, text: String): Unit = using(open_output(path))(stream => Bytes(text).write_stream(stream)) /* exec channel */ def exec(command: String): Exec = { val channel = session.openChannel("exec").asInstanceOf[ChannelExec] channel.setCommand("export USER_HOME=\"$HOME\"\n" + command) new Exec(this, channel) } override def execute(command: String, progress_stdout: String => Unit = (_: String) => (), progress_stderr: String => Unit = (_: String) => (), strict: Boolean = true): Process_Result = exec(command).result(progress_stdout, progress_stderr, strict) override def isabelle_platform: Isabelle_Platform = Isabelle_Platform(ssh = Some(this)) /* tmp dirs */ def rm_tree(dir: Path): Unit = rm_tree(remote_path(dir)) def rm_tree(remote_dir: String): Unit = execute("rm -r -f " + Bash.string(remote_dir)).check def tmp_dir(): String = execute("mktemp -d -t tmp.XXXXXXXXXX").check.out def with_tmp_dir[A](body: Path => A): A = { val remote_dir = tmp_dir() try { body(Path.explode(remote_dir)) } finally { rm_tree(remote_dir) } } } /* system operations */ trait System { def hg_url: String = "" def expand_path(path: Path): Path = path.expand def bash_path(path: Path): String = File.bash_path(path) def is_dir(path: Path): Boolean = path.is_dir def is_file(path: Path): Boolean = path.is_file def make_directory(path: Path): Path = Isabelle_System.make_directory(path) def execute(command: String, progress_stdout: String => Unit = (_: String) => (), progress_stderr: String => Unit = (_: String) => (), strict: Boolean = true): Process_Result = Isabelle_System.bash(command, progress_stdout = progress_stdout, progress_stderr = progress_stderr, strict = strict) def isabelle_platform: Isabelle_Platform = Isabelle_Platform() } object Local extends System } diff --git a/src/Pure/General/url.scala b/src/Pure/General/url.scala --- a/src/Pure/General/url.scala +++ b/src/Pure/General/url.scala @@ -1,110 +1,110 @@ /* Title: Pure/General/url.scala Author: Makarius Basic URL operations. */ package isabelle import java.io.{File => JFile} import java.nio.file.{Paths, FileSystemNotFoundException} import java.net.{URI, URISyntaxException, URL, MalformedURLException, URLDecoder, URLEncoder} import java.util.Locale import java.util.zip.GZIPInputStream object Url { /* special characters */ def escape_special(c: Char): String = if ("!#$&'()*+,/:;=?@[]".contains(c)) { String.format(Locale.ROOT, "%%%02X", Integer.valueOf(c.toInt)) } else c.toString def escape_special(s: String): String = s.iterator.map(escape_special).mkString def escape_name(name: String): String = name.iterator.map({ case '\'' => "%27" case c => c.toString }).mkString /* make and check URLs */ def apply(name: String): URL = { try { new URL(name) } catch { case _: MalformedURLException => error("Malformed URL " + quote(name)) } } def is_wellformed(name: String): Boolean = try { Url(name); true } catch { case ERROR(_) => false } def is_readable(name: String): Boolean = - try { Url(name).openStream.close; true } + try { Url(name).openStream.close(); true } catch { case ERROR(_) => false } /* trim index */ def trim_index(url: URL): URL = { Library.try_unprefix("/index.html", url.toString) match { case Some(u) => Url(u) case None => Library.try_unprefix("/index.php", url.toString) match { case Some(u) => Url(u) case None => url } } } /* strings */ def decode(s: String): String = URLDecoder.decode(s, UTF8.charset_name) def encode(s: String): String = URLEncoder.encode(s, UTF8.charset_name) /* read */ private def read(url: URL, gzip: Boolean): String = using(url.openStream)(stream => File.read_stream(if (gzip) new GZIPInputStream(stream) else stream)) def read(url: URL): String = read(url, false) def read_gzip(url: URL): String = read(url, true) def read(name: String): String = read(Url(name), false) def read_gzip(name: String): String = read(Url(name), true) def read_bytes(url: URL): Bytes = { val connection = url.openConnection val length = connection.getContentLength using(connection.getInputStream)(Bytes.read_stream(_, hint = length)) } /* file URIs */ def print_file(file: JFile): String = File.absolute(file).toPath.toUri.toString def print_file_name(name: String): String = print_file(new JFile(name)) def parse_file(uri: String): JFile = Paths.get(new URI(uri)).toFile def is_wellformed_file(uri: String): Boolean = try { parse_file(uri); true } catch { case _: URISyntaxException | _: IllegalArgumentException | _: FileSystemNotFoundException => false } def absolute_file(uri: String): JFile = File.absolute(parse_file(uri)) def absolute_file_name(uri: String): String = absolute_file(uri).getPath def canonical_file(uri: String): JFile = File.canonical(parse_file(uri)) def canonical_file_name(uri: String): String = canonical_file(uri).getPath } diff --git a/src/Pure/Isar/outer_syntax.scala b/src/Pure/Isar/outer_syntax.scala --- a/src/Pure/Isar/outer_syntax.scala +++ b/src/Pure/Isar/outer_syntax.scala @@ -1,211 +1,211 @@ /* Title: Pure/Isar/outer_syntax.scala Author: Makarius Isabelle/Isar outer syntax. */ package isabelle import scala.collection.mutable object Outer_Syntax { /* syntax */ val empty: Outer_Syntax = new Outer_Syntax() def merge(syns: List[Outer_Syntax]): Outer_Syntax = syns.foldLeft(empty)(_ ++ _) /* string literals */ def quote_string(str: String): String = { val result = new StringBuilder(str.length + 10) result += '"' for (s <- Symbol.iterator(str)) { if (s.length == 1) { val c = s(0) if (c < 32 && c != YXML.X && c != YXML.Y || c == '\\' || c == '"') { result += '\\' if (c < 10) result += '0' if (c < 100) result += '0' result ++= c.asInstanceOf[Int].toString } else result += c } else result ++= s } result += '"' result.toString } } final class Outer_Syntax private( val keywords: Keyword.Keywords = Keyword.Keywords.empty, val rev_abbrevs: Thy_Header.Abbrevs = Nil, val language_context: Completion.Language_Context = Completion.Language_Context.outer, val has_tokens: Boolean = true) { /** syntax content **/ override def toString: String = keywords.toString /* keywords */ def + (name: String, kind: String = "", load_command: String = ""): Outer_Syntax = new Outer_Syntax( keywords + (name, kind, load_command), rev_abbrevs, language_context, has_tokens = true) def add_keywords(keywords: Thy_Header.Keywords): Outer_Syntax = keywords.foldLeft(this) { case (syntax, (name, spec)) => syntax + (Symbol.decode(name), spec.kind, spec.load_command) + (Symbol.encode(name), spec.kind, spec.load_command) } /* abbrevs */ def abbrevs: Thy_Header.Abbrevs = rev_abbrevs.reverse def add_abbrevs(new_abbrevs: Thy_Header.Abbrevs): Outer_Syntax = if (new_abbrevs.isEmpty) this else { val rev_abbrevs1 = Library.distinct(new_abbrevs) reverse_::: rev_abbrevs new Outer_Syntax(keywords, rev_abbrevs1, language_context, has_tokens) } /* completion */ private lazy val completion: Completion = { val completion_keywords = (keywords.minor.iterator ++ keywords.major.iterator).toList val completion_abbrevs = completion_keywords.flatMap(name => (if (Keyword.theory_block.contains(keywords.kinds(name))) List((name, name + "\nbegin\n\u0007\nend")) else Nil) ::: (if (Completion.Word_Parsers.is_word(name)) List((name, name)) else Nil)) ::: abbrevs.flatMap( { case (a, b) => val a1 = Symbol.decode(a) val a2 = Symbol.encode(a) val b1 = Symbol.decode(b) List((a1, b1), (a2, b1)) }) Completion.make(completion_keywords, completion_abbrevs) } def complete( history: Completion.History, unicode: Boolean, explicit: Boolean, start: Text.Offset, text: CharSequence, caret: Int, context: Completion.Language_Context): Option[Completion.Result] = { completion.complete(history, unicode, explicit, start, text, caret, context) } /* build */ def + (header: Document.Node.Header): Outer_Syntax = add_keywords(header.keywords).add_abbrevs(header.abbrevs) def ++ (other: Outer_Syntax): Outer_Syntax = if (this eq other) this else if (this eq Outer_Syntax.empty) other else { val keywords1 = keywords ++ other.keywords val rev_abbrevs1 = Library.merge(rev_abbrevs, other.rev_abbrevs) if ((keywords eq keywords1) && (rev_abbrevs eq rev_abbrevs1)) this else new Outer_Syntax(keywords1, rev_abbrevs1, language_context, has_tokens) } /* load commands */ def load_command(name: String): Option[String] = keywords.load_commands.get(name) def has_load_commands(text: String): Boolean = keywords.load_commands.exists({ case (cmd, _) => text.containsSlice(cmd) }) /* language context */ def set_language_context(context: Completion.Language_Context): Outer_Syntax = new Outer_Syntax(keywords, rev_abbrevs, context, has_tokens) def no_tokens: Outer_Syntax = { require(keywords.is_empty, "bad syntax keywords") new Outer_Syntax( rev_abbrevs = rev_abbrevs, language_context = language_context, has_tokens = false) } /** parsing **/ /* command spans */ def parse_spans(toks: List[Token]): List[Command_Span.Span] = { val result = new mutable.ListBuffer[Command_Span.Span] val content = new mutable.ListBuffer[Token] val ignored = new mutable.ListBuffer[Token] def ship(content: List[Token]): Unit = { val kind = if (content.forall(_.is_ignored)) Command_Span.Ignored_Span else if (content.exists(_.is_error)) Command_Span.Malformed_Span else content.find(_.is_command) match { case None => Command_Span.Malformed_Span case Some(cmd) => val name = cmd.source val offset = content.takeWhile(_ != cmd).foldLeft(0) { case (i, tok) => i + Symbol.length(tok.source) } val end_offset = offset + Symbol.length(name) val range = Text.Range(offset, end_offset) + 1 Command_Span.Command_Span(name, Position.Range(range)) } result += Command_Span.Span(kind, content) } def flush(): Unit = { - if (content.nonEmpty) { ship(content.toList); content.clear } - if (ignored.nonEmpty) { ship(ignored.toList); ignored.clear } + if (content.nonEmpty) { ship(content.toList); content.clear() } + if (ignored.nonEmpty) { ship(ignored.toList); ignored.clear() } } for (tok <- toks) { if (tok.is_ignored) ignored += tok else if (keywords.is_before_command(tok) || tok.is_command && (!content.exists(keywords.is_before_command) || content.exists(_.is_command))) { flush(); content += tok } - else { content ++= ignored; ignored.clear; content += tok } + else { content ++= ignored; ignored.clear(); content += tok } } flush() result.toList } def parse_spans(input: CharSequence): List[Command_Span.Span] = parse_spans(Token.explode(keywords, input)) } diff --git a/src/Pure/ML/ml_console.scala b/src/Pure/ML/ml_console.scala --- a/src/Pure/ML/ml_console.scala +++ b/src/Pure/ML/ml_console.scala @@ -1,84 +1,84 @@ /* Title: Pure/ML/ml_console.scala Author: Makarius The raw ML process in interactive mode. */ package isabelle object ML_Console { /* command line entry point */ def main(args: Array[String]): Unit = { Command_Line.tool { var dirs: List[Path] = Nil var include_sessions: List[String] = Nil var logic = Isabelle_System.getenv("ISABELLE_LOGIC") var modes: List[String] = Nil var no_build = false var options = Options.init() var raw_ml_system = false val getopts = Getopts(""" Usage: isabelle console [OPTIONS] Options are: -d DIR include session directory -i NAME include session in name-space of theories -l NAME logic session name (default ISABELLE_LOGIC=""" + quote(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) -r bootstrap from raw Poly/ML Build a logic session image and run the raw Isabelle ML process in interactive mode, with line editor ISABELLE_LINE_EDITOR=""" + quote(Isabelle_System.getenv("ISABELLE_LINE_EDITOR")) + """. """, "d:" -> (arg => dirs = dirs ::: List(Path.explode(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), "r" -> (_ => raw_ml_system = true)) val more_args = getopts(args) if (more_args.nonEmpty) getopts.usage() val sessions_structure = Sessions.load_structure(options, dirs = dirs) val store = Sessions.store(options) // build logic if (!no_build && !raw_ml_system) { val progress = new Console_Progress() val rc = progress.interrupt_handler { Build.build_logic(options, logic, build_heap = true, progress = progress, dirs = dirs) } if (rc != 0) sys.exit(rc) } // process loop val process = ML_Process(options, sessions_structure, store, logic = logic, args = List("-i"), redirect = true, modes = if (raw_ml_system) Nil else modes ::: List("ASCII"), raw_ml_system = raw_ml_system, session_base = if (raw_ml_system) None else Some(Sessions.base_info( options, logic, dirs = dirs, include_sessions = include_sessions).check.base)) - POSIX_Interrupt.handler { process.interrupt } { + POSIX_Interrupt.handler { process.interrupt() } { new TTY_Loop(process.stdin, process.stdout).join val rc = process.join if (rc != 0) sys.exit(rc) } } } } diff --git a/src/Pure/ML/ml_statistics.scala b/src/Pure/ML/ml_statistics.scala --- a/src/Pure/ML/ml_statistics.scala +++ b/src/Pure/ML/ml_statistics.scala @@ -1,324 +1,324 @@ /* Title: Pure/ML/ml_statistics.scala Author: Makarius ML runtime statistics. */ package isabelle import scala.annotation.tailrec import scala.collection.mutable import scala.collection.immutable.{SortedSet, SortedMap} import scala.swing.{Frame, Component} import org.jfree.data.xy.{XYSeries, XYSeriesCollection} import org.jfree.chart.{JFreeChart, ChartPanel, ChartFactory} import org.jfree.chart.plot.PlotOrientation object ML_Statistics { /* properties */ val Now = new Properties.Double("now") def now(props: Properties.T): Double = Now.unapply(props).get /* memory status */ val Heap_Size = new Properties.Long("size_heap") val Heap_Free = new Properties.Long("size_heap_free_last_GC") val GC_Percent = new Properties.Int("GC_percent") sealed case class Memory_Status(heap_size: Long, heap_free: Long, gc_percent: Int) { def heap_used: Long = (heap_size - heap_free) max 0 def heap_used_fraction: Double = if (heap_size == 0) 0.0 else heap_used.toDouble / heap_size def gc_progress: Option[Double] = if (1 <= gc_percent && gc_percent <= 100) Some((gc_percent - 1) * 0.01) else None } def memory_status(props: Properties.T): Memory_Status = { val heap_size = Heap_Size.get(props) val heap_free = Heap_Free.get(props) val gc_percent = GC_Percent.get(props) Memory_Status(heap_size, heap_free, gc_percent) } /* monitor process */ def monitor(pid: Long, stats_dir: String = "", delay: Time = Time.seconds(0.5), consume: Properties.T => Unit = Console.println): Unit = { def progress_stdout(line: String): Unit = { val props = Library.space_explode(',', line).flatMap((entry: String) => Library.space_explode('=', entry) match { case List(a, b) => Some((a, b)) case _ => None }) if (props.nonEmpty) consume(props) } val env_prefix = if (stats_dir.isEmpty) "" else "export POLYSTATSDIR=" + Bash.string(stats_dir) + "\n" Bash.process(env_prefix + "exec \"$POLYML_EXE\" -q --use src/Pure/ML/ml_statistics.ML --eval " + Bash.string("ML_Statistics.monitor " + ML_Syntax.print_long(pid) + " " + ML_Syntax.print_double(delay.seconds)), cwd = Path.explode("~~").file) .result(progress_stdout = progress_stdout, strict = false).check } /* protocol handler */ class Handler extends Session.Protocol_Handler { private var session: Session = null private var monitoring: Future[Unit] = Future.value(()) override def init(session: Session): Unit = synchronized { this.session = session } override def exit(): Unit = synchronized { session = null - monitoring.cancel + monitoring.cancel() } private def consume(props: Properties.T): Unit = synchronized { if (session != null) { val props1 = (session.cache.props(props ::: Java_Statistics.jvm_statistics())) session.runtime_statistics.post(Session.Runtime_Statistics(props1)) } } private def ml_statistics(msg: Prover.Protocol_Output): Boolean = synchronized { msg.properties match { case Markup.ML_Statistics(pid, stats_dir) => monitoring = Future.thread("ML_statistics") { monitor(pid, stats_dir = stats_dir, consume = consume) } true case _ => false } } override val functions = List(Markup.ML_Statistics.name -> ml_statistics) } /* memory fields (mega bytes) */ def mem_print(x: Long): Option[String] = if (x == 0L) None else Some(x.toString + " M") def mem_scale(x: Long): Long = x / 1024 / 1024 def mem_field_scale(name: String, x: Double): Double = if (heap_fields._2.contains(name) || program_fields._2.contains(name)) mem_scale(x.toLong).toDouble else x val CODE_SIZE = "size_code" val STACK_SIZE = "size_stacks" val HEAP_SIZE = "size_heap" /* standard fields */ type Fields = (String, List[String]) val tasks_fields: Fields = ("Future tasks", List("tasks_ready", "tasks_pending", "tasks_running", "tasks_passive", "tasks_urgent", "tasks_total")) val workers_fields: Fields = ("Worker threads", List("workers_total", "workers_active", "workers_waiting")) val GC_fields: Fields = ("GCs", List("partial_GCs", "full_GCs", "share_passes")) val heap_fields: Fields = ("Heap", List(HEAP_SIZE, "size_allocation", "size_allocation_free", "size_heap_free_last_full_GC", "size_heap_free_last_GC")) val program_fields: Fields = ("Program", List("size_code", "size_stacks")) val threads_fields: Fields = ("Threads", List("threads_total", "threads_in_ML", "threads_wait_condvar", "threads_wait_IO", "threads_wait_mutex", "threads_wait_signal")) val time_fields: Fields = ("Time", List("time_elapsed", "time_elapsed_GC", "time_CPU", "time_GC")) val speed_fields: Fields = ("Speed", List("speed_CPU", "speed_GC")) private val time_speed = Map("time_CPU" -> "speed_CPU", "time_GC" -> "speed_GC") val java_heap_fields: Fields = ("Java heap", List("java_heap_size", "java_heap_used")) val java_thread_fields: Fields = ("Java threads", List("java_threads_total", "java_workers_total", "java_workers_active")) val main_fields: List[Fields] = List(heap_fields, tasks_fields, workers_fields) val other_fields: List[Fields] = List(threads_fields, GC_fields, program_fields, time_fields, speed_fields, java_heap_fields, java_thread_fields) val all_fields: List[Fields] = main_fields ::: other_fields /* content interpretation */ final case class Entry(time: Double, data: Map[String, Double]) { def get(field: String): Double = data.getOrElse(field, 0.0) } val empty: ML_Statistics = apply(Nil) def apply(ml_statistics0: List[Properties.T], heading: String = "", domain: String => Boolean = (key: String) => true): ML_Statistics = { require(ml_statistics0.forall(props => Now.unapply(props).isDefined), "missing \"now\" field") val ml_statistics = ml_statistics0.sortBy(now) val time_start = if (ml_statistics.isEmpty) 0.0 else now(ml_statistics.head) val duration = if (ml_statistics.isEmpty) 0.0 else now(ml_statistics.last) - time_start val fields = SortedSet.empty[String] ++ (for { props <- ml_statistics.iterator (x, _) <- props.iterator if x != Now.name && domain(x) } yield x) val content = { var last_edge = Map.empty[String, (Double, Double, Double)] val result = new mutable.ListBuffer[ML_Statistics.Entry] for (props <- ml_statistics) { val time = now(props) - time_start // rising edges -- relative speed val speeds = (for { (key, value) <- props.iterator key1 <- time_speed.get(key) if domain(key1) } yield { val (x0, y0, s0) = last_edge.getOrElse(key, (0.0, 0.0, 0.0)) val x1 = time val y1 = java.lang.Double.parseDouble(value) val s1 = if (x1 == x0) 0.0 else (y1 - y0) / (x1 - x0) if (y1 > y0) { last_edge += (key -> (x1, y1, s1)) (key1, s1.toString) } else (key1, s0.toString) }).toList val data = SortedMap.empty[String, Double] ++ (for { (x, y) <- props.iterator ++ speeds.iterator if x != Now.name && domain(x) z = java.lang.Double.parseDouble(y) if z != 0.0 } yield { (x.intern, mem_field_scale(x, z)) }) result += ML_Statistics.Entry(time, data) } result.toList } new ML_Statistics(heading, fields, content, time_start, duration) } } final class ML_Statistics private( val heading: String, val fields: Set[String], val content: List[ML_Statistics.Entry], val time_start: Double, val duration: Double) { /* content */ def maximum(field: String): Double = content.foldLeft(0.0) { case (m, e) => m max e.get(field) } def average(field: String): Double = { @tailrec def sum(t0: Double, list: List[ML_Statistics.Entry], acc: Double): Double = list match { case Nil => acc case e :: es => val t = e.time sum(t, es, (t - t0) * e.get(field) + acc) } content match { case Nil => 0.0 case List(e) => e.get(field) case e :: es => sum(e.time, es, 0.0) / duration } } /* charts */ def update_data(data: XYSeriesCollection, selected_fields: List[String]): Unit = { data.removeAllSeries for (field <- selected_fields) { val series = new XYSeries(field) content.foreach(entry => series.add(entry.time, entry.get(field))) data.addSeries(series) } } def chart(title: String, selected_fields: List[String]): JFreeChart = { val data = new XYSeriesCollection update_data(data, selected_fields) ChartFactory.createXYLineChart(title, "time", "value", data, PlotOrientation.VERTICAL, true, true, true) } def chart(fields: ML_Statistics.Fields): JFreeChart = chart(fields._1, fields._2) def show_frames(fields: List[ML_Statistics.Fields] = ML_Statistics.main_fields): Unit = fields.map(chart).foreach(c => GUI_Thread.later { new Frame { iconImage = GUI.isabelle_image() title = heading contents = Component.wrap(new ChartPanel(c)) visible = true } }) } diff --git a/src/Pure/PIDE/headless.scala b/src/Pure/PIDE/headless.scala --- a/src/Pure/PIDE/headless.scala +++ b/src/Pure/PIDE/headless.scala @@ -1,663 +1,663 @@ /* Title: Pure/PIDE/headless.scala Author: Makarius Headless PIDE session and resources from file-system. */ package isabelle import java.io.{File => JFile} import scala.annotation.tailrec import scala.collection.mutable object Headless { /** session **/ private def stable_snapshot( state: Document.State, version: Document.Version, name: Document.Node.Name): Document.Snapshot = { val snapshot = state.snapshot(name) assert(version.id == snapshot.version.id) snapshot } class Use_Theories_Result private[Headless]( val state: Document.State, val version: Document.Version, val nodes: List[(Document.Node.Name, Document_Status.Node_Status)], val nodes_committed: List[(Document.Node.Name, Document_Status.Node_Status)]) { def nodes_pending: List[(Document.Node.Name, Document_Status.Node_Status)] = { val committed = nodes_committed.iterator.map(_._1).toSet nodes.filter(p => !committed(p._1)) } def snapshot(name: Document.Node.Name): Document.Snapshot = stable_snapshot(state, version, name) def ok: Boolean = (nodes.iterator ++ nodes_committed.iterator).forall({ case (_, st) => st.ok }) } class Session private[Headless]( session_name: String, _session_options: => Options, override val resources: Resources) extends isabelle.Session(_session_options, resources) { session => private def loaded_theory(name: Document.Node.Name): Boolean = resources.session_base.loaded_theory(name.theory) /* options */ override def consolidate_delay: Time = session_options.seconds("headless_consolidate_delay") override def prune_delay: Time = session_options.seconds("headless_prune_delay") def default_check_delay: Time = session_options.seconds("headless_check_delay") def default_check_limit: Int = session_options.int("headless_check_limit") def default_nodes_status_delay: Time = session_options.seconds("headless_nodes_status_delay") def default_watchdog_timeout: Time = session_options.seconds("headless_watchdog_timeout") def default_commit_cleanup_delay: Time = session_options.seconds("headless_commit_cleanup_delay") /* temporary directory */ val tmp_dir: JFile = Isabelle_System.tmp_dir("server_session") val tmp_dir_name: String = File.path(tmp_dir).implode def master_directory(master_dir: String): String = proper_string(master_dir) getOrElse tmp_dir_name override def toString: String = session_name override def stop(): Process_Result = { try { super.stop() } finally { Isabelle_System.rm_tree(tmp_dir) } } /* theories */ private object Load_State { def finished: Load_State = Load_State(Nil, Nil, 0) def count_file(name: Document.Node.Name): Long = if (loaded_theory(name)) 0 else name.path.file.length } private case class Load_State( pending: List[Document.Node.Name], rest: List[Document.Node.Name], load_limit: Long) { def next( dep_graph: Document.Node.Name.Graph[Unit], finished: Document.Node.Name => Boolean): (List[Document.Node.Name], Load_State) = { def load_requirements(pending1: List[Document.Node.Name], rest1: List[Document.Node.Name]) : (List[Document.Node.Name], Load_State) = { val load_theories = dep_graph.all_preds_rev(pending1).filterNot(finished) (load_theories, Load_State(pending1, rest1, load_limit)) } if (!pending.forall(finished)) (Nil, this) else if (rest.isEmpty) (Nil, Load_State.finished) else if (load_limit == 0) load_requirements(rest, Nil) else { val reachable = dep_graph.reachable_limit(load_limit, Load_State.count_file, dep_graph.imm_preds, rest) val (pending1, rest1) = rest.partition(reachable) load_requirements(pending1, rest1) } } } private sealed case class Use_Theories_State( dep_graph: Document.Node.Name.Graph[Unit], load_state: Load_State, watchdog_timeout: Time, commit: Option[(Document.Snapshot, Document_Status.Node_Status) => Unit], last_update: Time = Time.now(), nodes_status: Document_Status.Nodes_Status = Document_Status.Nodes_Status.empty, already_committed: Map[Document.Node.Name, Document_Status.Node_Status] = Map.empty, result: Option[Exn.Result[Use_Theories_Result]] = None) { def update(new_nodes_status: Document_Status.Nodes_Status): Use_Theories_State = copy(last_update = Time.now(), nodes_status = new_nodes_status) def watchdog: Boolean = watchdog_timeout > Time.zero && Time.now() - last_update > watchdog_timeout def finished_result: Boolean = result.isDefined def join_result: Option[(Exn.Result[Use_Theories_Result], Use_Theories_State)] = if (finished_result) Some((result.get, this)) else None def cancel_result: Use_Theories_State = if (finished_result) this else copy(result = Some(Exn.Exn(Exn.Interrupt()))) def clean_theories: (List[Document.Node.Name], Use_Theories_State) = { @tailrec def frontier(base: List[Document.Node.Name], front: Set[Document.Node.Name]) : Set[Document.Node.Name] = { val add = base.filter(name => dep_graph.imm_succs(name).forall(front)) if (add.isEmpty) front else { val preds = add.map(dep_graph.imm_preds) val base1 = preds.tail.foldLeft(preds.head)(_ ++ _).toList.filter(already_committed.keySet) frontier(base1, front ++ add) } } if (already_committed.isEmpty) (Nil, this) else { val base = (for { (name, (_, (_, succs))) <- dep_graph.iterator if succs.isEmpty && already_committed.isDefinedAt(name) } yield name).toList val clean = frontier(base, Set.empty) if (clean.isEmpty) (Nil, this) else { (dep_graph.topological_order.filter(clean), copy(dep_graph = dep_graph.exclude(clean))) } } } def check(state: Document.State, version: Document.Version, beyond_limit: Boolean) : (List[Document.Node.Name], Use_Theories_State) = { val already_committed1 = commit match { case None => already_committed case Some(commit_fn) => dep_graph.topological_order.foldLeft(already_committed) { case (committed, name) => def parents_committed: Boolean = version.nodes(name).header.imports.forall(parent => loaded_theory(parent) || committed.isDefinedAt(parent)) if (!committed.isDefinedAt(name) && parents_committed && state.node_consolidated(version, name)) { val snapshot = stable_snapshot(state, version, name) val status = Document_Status.Node_Status.make(state, version, name) commit_fn(snapshot, status) committed + (name -> status) } else committed } } def finished_theory(name: Document.Node.Name): Boolean = loaded_theory(name) || (if (commit.isDefined) already_committed1.isDefinedAt(name) else state.node_consolidated(version, name)) val result1 = if (!finished_result && (beyond_limit || watchdog || dep_graph.keys_iterator.forall(name => finished_theory(name) || nodes_status.quasi_consolidated(name)))) { val nodes = (for { name <- dep_graph.keys_iterator if !loaded_theory(name) } yield { (name -> Document_Status.Node_Status.make(state, version, name)) }).toList val nodes_committed = (for { name <- dep_graph.keys_iterator status <- already_committed1.get(name) } yield (name -> status)).toList Some(Exn.Res(new Use_Theories_Result(state, version, nodes, nodes_committed))) } else result val (load_theories, load_state1) = load_state.next(dep_graph, finished_theory) (load_theories, copy(already_committed = already_committed1, result = result1, load_state = load_state1)) } } def use_theories( theories: List[String], qualifier: String = Sessions.DRAFT, master_dir: String = "", unicode_symbols: Boolean = false, check_delay: Time = default_check_delay, check_limit: Int = default_check_limit, watchdog_timeout: Time = default_watchdog_timeout, nodes_status_delay: Time = default_nodes_status_delay, id: UUID.T = UUID.random(), // commit: must not block, must not fail commit: Option[(Document.Snapshot, Document_Status.Node_Status) => Unit] = None, commit_cleanup_delay: Time = default_commit_cleanup_delay, progress: Progress = new Progress): Use_Theories_Result = { val dependencies = { val import_names = theories.map(thy => resources.import_name(qualifier, master_directory(master_dir), thy) -> Position.none) resources.dependencies(import_names, progress = progress).check_errors } val dep_theories = dependencies.theories val dep_theories_set = dep_theories.toSet val dep_files = for (path <- dependencies.loaded_files) yield Document.Node.Name(resources.append("", path)) val use_theories_state = { val dep_graph = dependencies.theory_graph val maximals = dep_graph.maximals val rest = if (maximals.isEmpty || maximals.tail.isEmpty) maximals else { val depth = dep_graph.node_depth(Load_State.count_file) maximals.sortBy(node => - depth(node)) } val load_limit = if (commit.isDefined) (session_options.real("headless_load_limit") * 1024 * 1024).round else 0 val load_state = Load_State(Nil, rest, load_limit) Synchronized(Use_Theories_State(dep_graph, load_state, watchdog_timeout, commit)) } def check_state(beyond_limit: Boolean = false): Unit = { val state = session.get_state() for { version <- state.stable_tip_version load_theories = use_theories_state.change_result(_.check(state, version, beyond_limit)) if load_theories.nonEmpty } resources.load_theories(session, id, load_theories, dep_files, unicode_symbols, progress) } val check_progress = { var check_count = 0 Event_Timer.request(Time.now(), repeat = Some(check_delay)) { if (progress.stopped) use_theories_state.change(_.cancel_result) else { check_count += 1 check_state(check_limit > 0 && check_count > check_limit) } } } val consumer = { val delay_nodes_status = Delay.first(nodes_status_delay max Time.zero) { progress.nodes_status(use_theories_state.value.nodes_status) } val delay_commit_clean = Delay.first(commit_cleanup_delay max Time.zero) { val clean_theories = use_theories_state.change_result(_.clean_theories) if (clean_theories.nonEmpty) { progress.echo("Removing " + clean_theories.length + " theories ...") resources.clean_theories(session, id, clean_theories) } } Session.Consumer[Session.Commands_Changed](getClass.getName) { case changed => if (changed.nodes.exists(dep_theories_set)) { val snapshot = session.snapshot() val state = snapshot.state val version = snapshot.version val theory_progress = use_theories_state.change_result(st => { val domain = if (st.nodes_status.is_empty) dep_theories_set else changed.nodes.iterator.filter(dep_theories_set).toSet val (nodes_status_changed, nodes_status1) = st.nodes_status.update(resources, state, version, domain = Some(domain), trim = changed.assignment) if (nodes_status_delay >= Time.zero && nodes_status_changed) { - delay_nodes_status.invoke + delay_nodes_status.invoke() } val theory_progress = (for { (name, node_status) <- nodes_status1.present.iterator if changed.nodes.contains(name) && !st.already_committed.isDefinedAt(name) p1 = node_status.percentage if p1 > 0 && Some(p1) != st.nodes_status.get(name).map(_.percentage) } yield Progress.Theory(name.theory, percentage = Some(p1))).toList (theory_progress, st.update(nodes_status1)) }) theory_progress.foreach(progress.theory) check_state() if (commit.isDefined && commit_cleanup_delay > Time.zero) { if (use_theories_state.value.finished_result) - delay_commit_clean.revoke - else delay_commit_clean.invoke + delay_commit_clean.revoke() + else delay_commit_clean.invoke() } } } } try { session.commands_changed += consumer check_state() use_theories_state.guarded_access(_.join_result) - check_progress.cancel + check_progress.cancel() } finally { session.commands_changed -= consumer resources.unload_theories(session, id, dep_theories) } Exn.release(use_theories_state.guarded_access(_.join_result)) } def purge_theories( theories: List[String], qualifier: String = Sessions.DRAFT, master_dir: String = "", all: Boolean = false): (List[Document.Node.Name], List[Document.Node.Name]) = { val nodes = if (all) None else Some(theories.map(resources.import_name(qualifier, master_directory(master_dir), _))) resources.purge_theories(session, nodes) } } /** resources **/ object Resources { def apply(options: Options, base_info: Sessions.Base_Info, log: Logger = No_Logger): Resources = new Resources(options, base_info, log = log) def make( options: Options, session_name: String, session_dirs: List[Path] = Nil, include_sessions: List[String] = Nil, progress: Progress = new Progress, log: Logger = No_Logger): Resources = { val base_info = Sessions.base_info(options, session_name, dirs = session_dirs, include_sessions = include_sessions, progress = progress) apply(options, base_info, log = log) } final class Theory private[Headless]( val node_name: Document.Node.Name, val node_header: Document.Node.Header, val text: String, val node_required: Boolean) { override def toString: String = node_name.toString def node_perspective: Document.Node.Perspective_Text = Document.Node.Perspective(node_required, Text.Perspective.empty, Document.Node.Overlays.empty) def make_edits(text_edits: List[Text.Edit]): List[Document.Edit_Text] = List(node_name -> Document.Node.Deps(node_header), node_name -> Document.Node.Edits(text_edits), node_name -> node_perspective) def node_edits(old: Option[Theory]): List[Document.Edit_Text] = { val (text_edits, old_required) = if (old.isEmpty) (Text.Edit.inserts(0, text), false) else (Text.Edit.replace(0, old.get.text, text), old.get.node_required) if (text_edits.isEmpty && node_required == old_required) Nil else make_edits(text_edits) } def purge_edits: List[Document.Edit_Text] = make_edits(Text.Edit.removes(0, text)) def required(required: Boolean): Theory = if (required == node_required) this else new Theory(node_name, node_header, text, required) } sealed case class State( blobs: Map[Document.Node.Name, Document.Blob] = Map.empty, theories: Map[Document.Node.Name, Theory] = Map.empty, required: Multi_Map[Document.Node.Name, UUID.T] = Multi_Map.empty) { /* blobs */ def doc_blobs: Document.Blobs = Document.Blobs(blobs) def update_blobs(names: List[Document.Node.Name]): (Document.Blobs, State) = { val new_blobs = names.flatMap(name => { val bytes = Bytes.read(name.path) def new_blob: Document.Blob = { val text = bytes.text Document.Blob(bytes, text, Symbol.Text_Chunk(text), changed = true) } blobs.get(name) match { case Some(blob) => if (blob.bytes == bytes) None else Some(name -> new_blob) case None => Some(name -> new_blob) } }) val blobs1 = new_blobs.foldLeft(blobs)(_ + _) val blobs2 = new_blobs.foldLeft(blobs) { case (map, (a, b)) => map + (a -> b.unchanged) } (Document.Blobs(blobs1), copy(blobs = blobs2)) } def blob_edits(name: Document.Node.Name, old_blob: Option[Document.Blob]) : List[Document.Edit_Text] = { val blob = blobs.getOrElse(name, error("Missing blob " + quote(name.toString))) val text_edits = old_blob match { case None => List(Text.Edit.insert(0, blob.source)) case Some(blob0) => Text.Edit.replace(0, blob0.source, blob.source) } if (text_edits.isEmpty) Nil else List(name -> Document.Node.Blob(blob), name -> Document.Node.Edits(text_edits)) } /* theories */ lazy val theory_graph: Document.Node.Name.Graph[Unit] = Document.Node.Name.make_graph( for ((name, theory) <- theories.toList) yield ((name, ()), theory.node_header.imports.filter(theories.isDefinedAt))) def is_required(name: Document.Node.Name): Boolean = required.isDefinedAt(name) def insert_required(id: UUID.T, names: List[Document.Node.Name]): State = copy(required = names.foldLeft(required)(_.insert(_, id))) def remove_required(id: UUID.T, names: List[Document.Node.Name]): State = copy(required = names.foldLeft(required)(_.remove(_, id))) def update_theories(update: List[(Document.Node.Name, Theory)]): State = copy(theories = update.foldLeft(theories) { case (thys, (name, thy)) => thys.get(name) match { case Some(thy1) if thy1 == thy => thys case _ => thys + (name -> thy) } }) def remove_theories(remove: List[Document.Node.Name]): State = { require(remove.forall(name => !is_required(name)), "attempt to remove required nodes") copy(theories = theories -- remove) } def unload_theories(session: Session, id: UUID.T, theories: List[Document.Node.Name]) : (List[Document.Edit_Text], State) = { val st1 = remove_required(id, theories) val theory_edits = for { node_name <- theories theory <- st1.theories.get(node_name) } yield { val theory1 = theory.required(st1.is_required(node_name)) val edits = theory1.node_edits(Some(theory)) (edits, (node_name, theory1)) } (theory_edits.flatMap(_._1), st1.update_theories(theory_edits.map(_._2))) } def purge_theories(session: Session, nodes: Option[List[Document.Node.Name]]) : ((List[Document.Node.Name], List[Document.Node.Name], List[Document.Edit_Text]), State) = { val all_nodes = theory_graph.topological_order val purge = nodes.getOrElse(all_nodes).filterNot(is_required).toSet val retain = theory_graph.all_preds(all_nodes.filterNot(purge)).toSet val (retained, purged) = all_nodes.partition(retain) val purge_edits = purged.flatMap(name => theories(name).purge_edits) ((purged, retained, purge_edits), remove_theories(purged)) } } } class Resources private[Headless]( val options: Options, val session_base_info: Sessions.Base_Info, log: Logger = No_Logger) extends isabelle.Resources( session_base_info.sessions_structure, session_base_info.check.base, log = log) { resources => val store: Sessions.Store = Sessions.store(options) /* session */ def start_session(print_mode: List[String] = Nil, progress: Progress = new Progress): Session = { val session = new Session(session_base_info.session, options, resources) progress.echo("Starting session " + session_base_info.session + " ...") Isabelle_Process(session, options, session_base_info.sessions_structure, store, - logic = session_base_info.session, modes = print_mode).await_startup + logic = session_base_info.session, modes = print_mode).await_startup() session } /* theories */ private val state = Synchronized(Resources.State()) def load_theories( session: Session, id: UUID.T, theories: List[Document.Node.Name], files: List[Document.Node.Name], unicode_symbols: Boolean, progress: Progress): Unit = { val loaded_theories = for (node_name <- theories) yield { val path = node_name.path if (!node_name.is_theory) error("Not a theory file: " + path) progress.expose_interrupt() val text0 = File.read(path) val text = if (unicode_symbols) Symbol.decode(text0) else text0 val node_header = resources.check_thy(node_name, Scan.char_reader(text)) new Resources.Theory(node_name, node_header, text, true) } val loaded = loaded_theories.length if (loaded > 1) progress.echo("Loading " + loaded + " theories ...") state.change(st => { val (doc_blobs1, st1) = st.insert_required(id, theories).update_blobs(files) val theory_edits = for (theory <- loaded_theories) yield { val node_name = theory.node_name val theory1 = theory.required(st1.is_required(node_name)) val edits = theory1.node_edits(st1.theories.get(node_name)) (edits, (node_name, theory1)) } val file_edits = for { node_name <- files if doc_blobs1.changed(node_name) } yield st1.blob_edits(node_name, st.blobs.get(node_name)) session.update(doc_blobs1, theory_edits.flatMap(_._1) ::: file_edits.flatten) st1.update_theories(theory_edits.map(_._2)) }) } def unload_theories(session: Session, id: UUID.T, theories: List[Document.Node.Name]): Unit = { state.change(st => { val (edits, st1) = st.unload_theories(session, id, theories) session.update(st.doc_blobs, edits) st1 }) } def clean_theories(session: Session, id: UUID.T, theories: List[Document.Node.Name]): Unit = { state.change(st => { val (edits1, st1) = st.unload_theories(session, id, theories) val ((_, _, edits2), st2) = st1.purge_theories(session, None) session.update(st.doc_blobs, edits1 ::: edits2) st2 }) } def purge_theories(session: Session, nodes: Option[List[Document.Node.Name]]) : (List[Document.Node.Name], List[Document.Node.Name]) = { state.change_result(st => { val ((purged, retained, _), st1) = st.purge_theories(session, nodes) ((purged, retained), st1) }) } } } diff --git a/src/Pure/PIDE/prover.scala b/src/Pure/PIDE/prover.scala --- a/src/Pure/PIDE/prover.scala +++ b/src/Pure/PIDE/prover.scala @@ -1,368 +1,368 @@ /* Title: Pure/PIDE/prover.scala Author: Makarius Options: :folding=explicit: Prover process wrapping. */ package isabelle import java.io.{InputStream, OutputStream, BufferedOutputStream, IOException} object Prover { /* messages */ sealed abstract class Message type Receiver = Message => Unit class Input(val name: String, val args: List[String]) extends Message { override def toString: String = XML.Elem(Markup(Markup.PROVER_COMMAND, List((Markup.NAME, name))), args.flatMap(s => List(XML.newline, XML.elem(Markup.PROVER_ARG, YXML.parse_body(s))))).toString } class Output(val message: XML.Elem) extends Message { def kind: String = message.markup.name def properties: Properties.T = message.markup.properties def body: XML.Body = message.body def is_init: Boolean = kind == Markup.INIT def is_exit: Boolean = kind == Markup.EXIT def is_stdout: Boolean = kind == Markup.STDOUT def is_stderr: Boolean = kind == Markup.STDERR def is_system: Boolean = kind == Markup.SYSTEM def is_status: Boolean = kind == Markup.STATUS def is_report: Boolean = kind == Markup.REPORT def is_syslog: Boolean = is_init || is_exit || is_system || is_stderr override def toString: String = { val res = if (is_status || is_report) message.body.map(_.toString).mkString else Pretty.string_of(message.body, metric = Symbol.Metric) if (properties.isEmpty) kind.toString + " [[" + res + "]]" else kind.toString + " " + (for ((x, y) <- properties) yield x + "=" + y).mkString("{", ",", "}") + " [[" + res + "]]" } } class Protocol_Output(props: Properties.T, val bytes: Bytes) extends Output(XML.Elem(Markup(Markup.PROTOCOL, props), Nil)) { lazy val text: String = bytes.text } } class Prover( receiver: Prover.Receiver, cache: XML.Cache, channel: System_Channel, process: Bash.Process) extends Protocol { /** receiver output **/ private def system_output(text: String): Unit = { receiver(new Prover.Output(XML.Elem(Markup(Markup.SYSTEM, Nil), List(XML.Text(text))))) } private def protocol_output(props: Properties.T, bytes: Bytes): Unit = { receiver(new Prover.Protocol_Output(cache.props(props), bytes)) } private def output(kind: String, props: Properties.T, body: XML.Body): Unit = { val main = XML.Elem(Markup(kind, props), Protocol_Message.clean_reports(body)) val reports = Protocol_Message.reports(props, body) for (msg <- main :: reports) receiver(new Prover.Output(cache.elem(msg))) } private def exit_message(result: Process_Result): Unit = { output(Markup.EXIT, Markup.Process_Result(result), List(XML.Text(result.print_return_code))) } /** process manager **/ private val process_result: Future[Process_Result] = Future.thread("process_result") { val rc = process.join val timing = process.get_timing Process_Result(rc, timing = timing) } private def terminate_process(): Unit = { - try { process.terminate } + try { process.terminate() } catch { case exn @ ERROR(_) => system_output("Failed to terminate prover process: " + exn.getMessage) } } private val process_manager = Isabelle_Thread.fork(name = "process_manager") { val stdout = physical_output(false) val (startup_failed, startup_errors) = { var finished: Option[Boolean] = None val result = new StringBuilder(100) while (finished.isEmpty && (process.stderr.ready || !process_result.is_finished)) { while (finished.isEmpty && process.stderr.ready) { try { val c = process.stderr.read if (c == 2) finished = Some(true) else result += c.toChar } catch { case _: IOException => finished = Some(false) } } Time.seconds(0.05).sleep } (finished.isEmpty || !finished.get, result.toString.trim) } if (startup_errors != "") system_output(startup_errors) if (startup_failed) { terminate_process() process_result.join stdout.join exit_message(Process_Result(127)) } else { val (command_stream, message_stream) = channel.rendezvous() command_input_init(command_stream) val stderr = physical_output(true) val message = message_output(message_stream) val result = process_result.join system_output("process terminated") command_input_close() for (thread <- List(stdout, stderr, message)) thread.join system_output("process_manager terminated") exit_message(result) } channel.shutdown() } /* management methods */ def join(): Unit = process_manager.join() def terminate(): Unit = { system_output("Terminating prover process") command_input_close() var count = 10 while (!process_result.is_finished && count > 0) { Time.seconds(0.1).sleep count -= 1 } if (!process_result.is_finished) terminate_process() } /** process streams **/ /* command input */ private var command_input: Option[Consumer_Thread[List[Bytes]]] = None - private def command_input_close(): Unit = command_input.foreach(_.shutdown) + private def command_input_close(): Unit = command_input.foreach(_.shutdown()) private def command_input_init(raw_stream: OutputStream): Unit = { val name = "command_input" val stream = new BufferedOutputStream(raw_stream) command_input = Some( Consumer_Thread.fork(name)( consume = { case chunks => try { Bytes(chunks.map(_.length).mkString("", ",", "\n")).write_stream(stream) chunks.foreach(_.write_stream(stream)) stream.flush true } catch { case e: IOException => system_output(name + ": " + e.getMessage); false } }, - finish = { case () => stream.close; system_output(name + " terminated") } + finish = { case () => stream.close(); system_output(name + " terminated") } ) ) } /* physical output */ private def physical_output(err: Boolean): Thread = { val (name, reader, markup) = if (err) ("standard_error", process.stderr, Markup.STDERR) else ("standard_output", process.stdout, Markup.STDOUT) Isabelle_Thread.fork(name = name) { try { var result = new StringBuilder(100) var finished = false while (!finished) { //{{{ var c = -1 var done = false while (!done && (result.isEmpty || reader.ready)) { c = reader.read if (c >= 0) result.append(c.asInstanceOf[Char]) else done = true } if (result.nonEmpty) { output(markup, Nil, List(XML.Text(Symbol.decode(result.toString)))) - result.clear + result.clear() } else { - reader.close + reader.close() finished = true } //}}} } } catch { case e: IOException => system_output(name + ": " + e.getMessage) } system_output(name + " terminated") } } /* message output */ private def message_output(stream: InputStream): Thread = { class EOF extends Exception class Protocol_Error(msg: String) extends Exception(msg) val name = "message_output" Isabelle_Thread.fork(name = name) { val default_buffer = new Array[Byte](65536) var c = -1 def read_int(): Int = //{{{ { var n = 0 c = stream.read if (c == -1) throw new EOF while (48 <= c && c <= 57) { n = 10 * n + (c - 48) c = stream.read } if (c != 10) throw new Protocol_Error("malformed header: expected integer followed by newline") else n } //}}} def read_chunk_bytes(): (Array[Byte], Int) = //{{{ { val n = read_int() val buf = if (n <= default_buffer.length) default_buffer else new Array[Byte](n) var i = 0 var m = 0 do { m = stream.read(buf, i, n - i) if (m != -1) i += m } while (m != -1 && n > i) if (i != n) throw new Protocol_Error("bad chunk (unexpected EOF after " + i + " of " + n + " bytes)") (buf, n) } //}}} def read_chunk(): XML.Body = { val (buf, n) = read_chunk_bytes() YXML.parse_body_failsafe(UTF8.decode_chars(Symbol.decode, buf, 0, n)) } try { do { try { val header = read_chunk() header match { case List(XML.Elem(Markup(name, props), Nil)) => val kind = name.intern if (kind == Markup.PROTOCOL) { val (buf, n) = read_chunk_bytes() protocol_output(props, Bytes(buf, 0, n)) } else { val body = read_chunk() output(kind, props, body) } case _ => read_chunk() throw new Protocol_Error("bad header: " + header.toString) } } catch { case _: EOF => } } while (c != -1) } catch { case e: IOException => system_output("Cannot read message:\n" + e.getMessage) case e: Protocol_Error => system_output("Malformed message:\n" + e.getMessage) } - stream.close + stream.close() system_output(name + " terminated") } } /** protocol commands **/ var trace: Boolean = false def protocol_command_raw(name: String, args: List[Bytes]): Unit = command_input match { case Some(thread) if thread.is_active => if (trace) { val payload = args.foldLeft(0) { case (n, b) => n + b.length } Output.writeln( "protocol_command " + name + ", args = " + args.length + ", payload = " + payload) } thread.send(Bytes(name) :: args) case _ => error("Inactive prover input thread for command " + quote(name)) } def protocol_command_args(name: String, args: List[String]): Unit = { receiver(new Prover.Input(name, args)) protocol_command_raw(name, args.map(Bytes(_))) } def protocol_command(name: String, args: String*): Unit = protocol_command_args(name, args.toList) } 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,774 +1,774 @@ /* 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 Build_Theories(id: String, master_dir: Path, theories: List[(Options, List[Path])]) 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 */ private[Session] class Syslog(limit: Int) { private var queue = Queue.empty[XML.Elem] private var length = 0 def += (msg: XML.Elem): Unit = synchronized { queue = queue.enqueue(msg) length += 1 if (length > limit) queue = queue.dequeue._2 } def content: String = synchronized { cat_lines(queue.iterator.map(XML.content)) + (if (length > limit) "\n(A total of " + length + " messages...)" else "") } } /* protocol handlers */ type Protocol_Function = Prover.Protocol_Output => Boolean abstract class Protocol_Handler extends Isabelle_System.Service { def init(session: Session): Unit = {} def exit(): Unit = {} def functions: List[(String, Protocol_Function)] = Nil def prover_options(options: Options): Options = options } } class Session(_session_options: => Options, val resources: Resources) extends Document.Session { session => val cache: XML.Cache = XML.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(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 */ private val syslog = new Session.Syslog(syslog_limit) def syslog_content(): String = syslog.content /* 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(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(new_nodes: Set[Document.Node.Name] = Set.empty): Unit = { val active = state.change_result(st => (st.isDefined, st.map(nodes => if (nodes.isEmpty) new_nodes else nodes ++ new_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 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 export = Export.make_entry("", args, msg.bytes, cache) change_command(_.add_export(id, (args.serial, export))) 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.PROPERTY) => 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.PROPERTY) => 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.PROPERTY) => 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 += 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 + 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(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 @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(name: String, args: String*): Unit = manager.send(Protocol_Command(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/Pure/System/bash.scala b/src/Pure/System/bash.scala --- a/src/Pure/System/bash.scala +++ b/src/Pure/System/bash.scala @@ -1,235 +1,235 @@ /* Title: Pure/System/bash.scala Author: Makarius GNU bash processes, with propagation of interrupts. */ package isabelle import java.io.{File => JFile, BufferedReader, InputStreamReader, BufferedWriter, OutputStreamWriter} import scala.annotation.tailrec object Bash { /* concrete syntax */ private def bash_chr(c: Byte): String = { val ch = c.toChar ch match { case '\t' => "$'\\t'" case '\n' => "$'\\n'" case '\f' => "$'\\f'" case '\r' => "$'\\r'" case _ => if (Symbol.is_ascii_letter(ch) || Symbol.is_ascii_digit(ch) || "+,-./:_".contains(ch)) Symbol.ascii(ch) else if (c < 0) "$'\\x" + Integer.toHexString(256 + c) + "'" else if (c < 16) "$'\\x0" + Integer.toHexString(c) + "'" else if (c < 32 || c >= 127) "$'\\x" + Integer.toHexString(c) + "'" else "\\" + ch } } def string(s: String): String = if (s == "") "\"\"" else UTF8.bytes(s).iterator.map(bash_chr).mkString def strings(ss: List[String]): String = ss.iterator.map(Bash.string).mkString(" ") /* process and result */ type Watchdog = (Time, Process => Boolean) def process(script: String, cwd: JFile = null, env: Map[String, String] = Isabelle_System.settings(), redirect: Boolean = false, cleanup: () => Unit = () => ()): Process = new Process(script, cwd, env, redirect, cleanup) class Process private[Bash]( script: String, cwd: JFile, env: Map[String, String], redirect: Boolean, cleanup: () => Unit) { private val timing_file = Isabelle_System.tmp_file("bash_timing") private val timing = Synchronized[Option[Timing]](None) def get_timing: Timing = timing.value getOrElse Timing.zero private val script_file = Isabelle_System.tmp_file("bash_script") File.write(script_file, script) private val proc = Isabelle_System.process( List(File.platform_path(Path.variable("ISABELLE_BASH_PROCESS")), "-", File.standard_path(timing_file), "bash", File.standard_path(script_file)), cwd = cwd, env = env, redirect = redirect) // channels val stdin: BufferedWriter = new BufferedWriter(new OutputStreamWriter(proc.getOutputStream, UTF8.charset)) val stdout: BufferedReader = new BufferedReader(new InputStreamReader(proc.getInputStream, UTF8.charset)) val stderr: BufferedReader = new BufferedReader(new InputStreamReader(proc.getErrorStream, UTF8.charset)) // signals private val pid = stdout.readLine @tailrec private def kill(signal: String, count: Int = 1): Boolean = { count <= 0 || { Isabelle_System.kill(signal, pid) val running = Isabelle_System.kill("0", pid)._2 == 0 if (running) { Time.seconds(0.1).sleep kill(signal, count - 1) } else false } } def terminate(): Unit = Isabelle_Thread.try_uninterruptible { kill("INT", count = 7) && kill("TERM", count = 3) && kill("KILL") - proc.destroy + proc.destroy() do_cleanup() } def interrupt(): Unit = Isabelle_Thread.try_uninterruptible { Isabelle_System.kill("INT", pid) } // JVM shutdown hook private val shutdown_hook = Isabelle_Thread.create(() => terminate()) try { Runtime.getRuntime.addShutdownHook(shutdown_hook) } catch { case _: IllegalStateException => } // cleanup private def do_cleanup(): Unit = { try { Runtime.getRuntime.removeShutdownHook(shutdown_hook) } catch { case _: IllegalStateException => } script_file.delete timing.change { case None => if (timing_file.isFile) { val t = Word.explode(File.read(timing_file)) match { case List(Value.Long(elapsed), Value.Long(cpu)) => Timing(Time.ms(elapsed), Time.ms(cpu), Time.zero) case _ => Timing.zero } timing_file.delete Some(t) } else Some(Timing.zero) case some => some } cleanup() } // join def join: Int = { val rc = proc.waitFor do_cleanup() rc } // result def result( progress_stdout: String => Unit = (_: String) => (), progress_stderr: String => Unit = (_: String) => (), watchdog: Option[Watchdog] = None, strict: Boolean = true): Process_Result = { - stdin.close + stdin.close() val out_lines = Future.thread("bash_stdout") { File.read_lines(stdout, progress_stdout) } val err_lines = Future.thread("bash_stderr") { File.read_lines(stderr, progress_stderr) } val watchdog_thread = for ((time, check) <- watchdog) yield { Future.thread("bash_watchdog") { while (proc.isAlive) { time.sleep if (check(this)) interrupt() } } } val rc = try { join } catch { case Exn.Interrupt() => terminate(); Exn.Interrupt.return_code } - watchdog_thread.foreach(_.cancel) + watchdog_thread.foreach(_.cancel()) if (strict && rc == Exn.Interrupt.return_code) throw Exn.Interrupt() Process_Result(rc, out_lines.join, err_lines.join, get_timing) } } /* Scala function */ object Process extends Scala.Fun("bash_process") { val here = Scala_Project.here def apply(script: String): String = { val result = Exn.capture { Isabelle_System.bash(script) } val is_interrupt = result match { case Exn.Res(res) => res.rc == Exn.Interrupt.return_code case Exn.Exn(exn) => Exn.is_interrupt(exn) } result match { case _ if is_interrupt => "" case Exn.Exn(exn) => Exn.message(exn) case Exn.Res(res) => Library.cat_strings0( res.rc.toString :: res.timing.elapsed.ms.toString :: res.timing.cpu.ms.toString :: res.out_lines.length.toString :: res.out_lines ::: res.err_lines) } } } } diff --git a/src/Pure/System/isabelle_process.scala b/src/Pure/System/isabelle_process.scala --- a/src/Pure/System/isabelle_process.scala +++ b/src/Pure/System/isabelle_process.scala @@ -1,81 +1,81 @@ /* Title: Pure/System/isabelle_process.scala Author: Makarius Isabelle process wrapper. */ package isabelle import java.io.{File => JFile} object Isabelle_Process { def apply( session: Session, options: Options, sessions_structure: Sessions.Structure, store: Sessions.Store, logic: String = "", raw_ml_system: Boolean = false, use_prelude: List[String] = Nil, eval_main: String = "", modes: List[String] = Nil, cwd: JFile = null, env: Map[String, String] = Isabelle_System.settings()): Isabelle_Process = { val channel = System_Channel() val process = try { val channel_options = options.string.update("system_channel_address", channel.address). string.update("system_channel_password", channel.password) ML_Process(channel_options, sessions_structure, store, logic = logic, raw_ml_system = raw_ml_system, use_prelude = use_prelude, eval_main = eval_main, modes = modes, cwd = cwd, env = env) } catch { case exn @ ERROR(_) => channel.shutdown(); throw exn } - process.stdin.close + process.stdin.close() new Isabelle_Process(session, channel, process) } } class Isabelle_Process private(session: Session, channel: System_Channel, process: Bash.Process) { private val startup = Future.promise[String] private val terminated = Future.promise[Process_Result] session.phase_changed += Session.Consumer(getClass.getName) { case Session.Ready => startup.fulfill("") case Session.Terminated(result) => if (!result.ok && !startup.is_finished) { val syslog = session.syslog_content() val err = "Session startup failed" + (if (syslog.isEmpty) "" else ":\n" + syslog) startup.fulfill(err) } terminated.fulfill(result) case _ => } session.start(receiver => new Prover(receiver, session.cache, channel, process)) def await_startup(): Isabelle_Process = startup.join match { case "" => this case err => session.stop(); error(err) } def await_shutdown(): Process_Result = { val result = terminated.join session.stop() result } - def terminate: Unit = process.terminate + def terminate(): Unit = process.terminate() } diff --git a/src/Pure/System/isabelle_system.scala b/src/Pure/System/isabelle_system.scala --- a/src/Pure/System/isabelle_system.scala +++ b/src/Pure/System/isabelle_system.scala @@ -1,568 +1,568 @@ /* Title: Pure/System/isabelle_system.scala Author: Makarius Fundamental Isabelle system environment: quasi-static module with optional init operation. */ package isabelle import java.io.{File => JFile, IOException} import java.nio.file.{Path => JPath, Files, SimpleFileVisitor, FileVisitResult, StandardCopyOption, FileSystemException} import java.nio.file.attribute.BasicFileAttributes object Isabelle_System { /** bootstrap information **/ def jdk_home(): String = { val java_home = System.getProperty("java.home", "") val home = new JFile(java_home) val parent = home.getParent if (home.getName == "jre" && parent != null && (new JFile(new JFile(parent, "bin"), "javac")).exists) parent else java_home } def bootstrap_directory( preference: String, envar: String, property: String, description: String): String = { val value = proper_string(preference) orElse // explicit argument proper_string(System.getenv(envar)) orElse // e.g. inherited from running isabelle tool proper_string(System.getProperty(property)) getOrElse // e.g. via JVM application boot process error("Unknown " + description + " directory") if ((new JFile(value)).isDirectory) value else error("Bad " + description + " directory " + quote(value)) } /** implicit settings environment **/ abstract class Service @volatile private var _settings: Option[Map[String, String]] = None @volatile private var _services: Option[List[Class[Service]]] = None def settings(): Map[String, String] = { if (_settings.isEmpty) init() // unsynchronized check _settings.get } def services(): List[Class[Service]] = { if (_services.isEmpty) init() // unsynchronized check _services.get } def make_services[C](c: Class[C]): List[C] = for { c1 <- services() if Library.is_subclass(c1, c) } yield c1.getDeclaredConstructor().newInstance().asInstanceOf[C] def init(isabelle_root: String = "", cygwin_root: String = ""): Unit = synchronized { if (_settings.isEmpty || _services.isEmpty) { val isabelle_root1 = bootstrap_directory(isabelle_root, "ISABELLE_ROOT", "isabelle.root", "Isabelle root") val cygwin_root1 = if (Platform.is_windows) bootstrap_directory(cygwin_root, "CYGWIN_ROOT", "cygwin.root", "Cygwin root") else "" if (Platform.is_windows) Cygwin.init(isabelle_root1, cygwin_root1) def set_cygwin_root(): Unit = { if (Platform.is_windows) _settings = Some(_settings.getOrElse(Map.empty) + ("CYGWIN_ROOT" -> cygwin_root1)) } set_cygwin_root() def default(env: Map[String, String], entry: (String, String)): Map[String, String] = if (env.isDefinedAt(entry._1) || entry._2 == "") env else env + entry val env = { val temp_windows = { val temp = if (Platform.is_windows) System.getenv("TEMP") else null if (temp != null && temp.contains('\\')) temp else "" } val user_home = System.getProperty("user.home", "") val isabelle_app = System.getProperty("isabelle.app", "") default( default( default(sys.env + ("ISABELLE_JDK_HOME" -> File.standard_path(jdk_home())), "TEMP_WINDOWS" -> temp_windows), "HOME" -> user_home), "ISABELLE_APP" -> "true") } val settings = { val dump = JFile.createTempFile("settings", null) dump.deleteOnExit try { val cmd1 = if (Platform.is_windows) List(cygwin_root1 + "\\bin\\bash", "-l", File.standard_path(isabelle_root1 + "\\bin\\isabelle")) else List(isabelle_root1 + "/bin/isabelle") val cmd = cmd1 ::: List("getenv", "-d", dump.toString) val (output, rc) = process_output(process(cmd, env = env, redirect = true)) if (rc != 0) error(output) val entries = (for (entry <- Library.split_strings0(File.read(dump)) if entry != "") yield { val i = entry.indexOf('=') if (i <= 0) entry -> "" else entry.substring(0, i) -> entry.substring(i + 1) }).toMap entries + ("PATH" -> entries("PATH_JVM")) - "PATH_JVM" } finally { dump.delete } } _settings = Some(settings) set_cygwin_root() val variable = "ISABELLE_SCALA_SERVICES" val services = for (name <- space_explode(':', settings.getOrElse(variable, getenv_error(variable)))) yield { def err(msg: String): Nothing = error("Bad entry " + quote(name) + " in " + variable + "\n" + msg) try { Class.forName(name).asInstanceOf[Class[Service]] } catch { case _: ClassNotFoundException => err("Class not found") case exn: Throwable => err(Exn.message(exn)) } } _services = Some(services) } } /* getenv */ private def getenv_error(name: String): Nothing = error("Undefined Isabelle environment variable: " + quote(name)) def getenv(name: String, env: Map[String, String] = settings()): String = env.getOrElse(name, "") def getenv_strict(name: String, env: Map[String, String] = settings()): String = proper_string(getenv(name, env)) getOrElse error("Undefined Isabelle environment variable: " + quote(name)) def cygwin_root(): String = getenv_strict("CYGWIN_ROOT") def isabelle_id(): String = proper_string(getenv("ISABELLE_ID")) getOrElse Mercurial.repository(Path.explode("~~")).parent() /** file-system operations **/ /* scala functions */ private def apply_paths(arg: String, fun: List[Path] => Unit): String = { fun(Library.split_strings0(arg).map(Path.explode)); "" } private def apply_paths1(arg: String, fun: Path => Unit): String = apply_paths(arg, { case List(path) => fun(path) }) private def apply_paths2(arg: String, fun: (Path, Path) => Unit): String = apply_paths(arg, { case List(path1, path2) => fun(path1, path2) }) private def apply_paths3(arg: String, fun: (Path, Path, Path) => Unit): String = apply_paths(arg, { case List(path1, path2, path3) => fun(path1, path2, path3) }) /* permissions */ def chmod(arg: String, path: Path): Unit = bash("chmod " + arg + " " + File.bash_path(path)).check def chown(arg: String, path: Path): Unit = bash("chown " + arg + " " + File.bash_path(path)).check /* directories */ def make_directory(path: Path): Path = { if (!path.is_dir) { try { Files.createDirectories(path.file.toPath) } catch { case ERROR(_) => error("Failed to create directory: " + path.absolute) } } path } def new_directory(path: Path): Path = if (path.is_dir) error("Directory already exists: " + path.absolute) else make_directory(path) def copy_dir(dir1: Path, dir2: Path): Unit = { val res = bash("cp -a " + File.bash_path(dir1) + " " + File.bash_path(dir2)) if (!res.ok) { cat_error("Failed to copy directory " + dir1.absolute + " to " + dir2.absolute, res.err) } } object Make_Directory extends Scala.Fun("make_directory") { val here = Scala_Project.here def apply(arg: String): String = apply_paths1(arg, make_directory) } object Copy_Dir extends Scala.Fun("copy_dir") { val here = Scala_Project.here def apply(arg: String): String = apply_paths2(arg, copy_dir) } /* copy files */ def copy_file(src: JFile, dst: JFile): Unit = { val target = if (dst.isDirectory) new JFile(dst, src.getName) else dst if (!File.eq(src, target)) { try { Files.copy(src.toPath, target.toPath, StandardCopyOption.COPY_ATTRIBUTES, StandardCopyOption.REPLACE_EXISTING) } catch { case ERROR(msg) => cat_error("Failed top copy file " + File.path(src).absolute + " to " + File.path(dst).absolute, msg) } } } def copy_file(src: Path, dst: Path): Unit = copy_file(src.file, dst.file) def copy_file_base(base_dir: Path, src: Path, target_dir: Path): Unit = { val src1 = src.expand val src1_dir = src1.dir if (!src1.starts_basic) error("Illegal path specification " + src1 + " beyond base directory") copy_file(base_dir + src1, Isabelle_System.make_directory(target_dir + src1_dir)) } object Copy_File extends Scala.Fun("copy_file") { val here = Scala_Project.here def apply(arg: String): String = apply_paths2(arg, copy_file) } object Copy_File_Base extends Scala.Fun("copy_file_base") { val here = Scala_Project.here def apply(arg: String): String = apply_paths3(arg, copy_file_base) } /* move files */ def move_file(src: JFile, dst: JFile): Unit = { val target = if (dst.isDirectory) new JFile(dst, src.getName) else dst if (!File.eq(src, target)) Files.move(src.toPath, target.toPath, StandardCopyOption.REPLACE_EXISTING) } def move_file(src: Path, dst: Path): Unit = move_file(src.file, dst.file) /* symbolic link */ def symlink(src: Path, dst: Path, force: Boolean = false): Unit = { val src_file = src.file val dst_file = dst.file val target = if (dst_file.isDirectory) new JFile(dst_file, src_file.getName) else dst_file if (force) target.delete try { Files.createSymbolicLink(target.toPath, src_file.toPath) } catch { case _: UnsupportedOperationException if Platform.is_windows => Cygwin.link(File.standard_path(src), target) case _: FileSystemException if Platform.is_windows => Cygwin.link(File.standard_path(src), target) } } /* tmp files */ def isabelle_tmp_prefix(): JFile = { val path = Path.explode("$ISABELLE_TMP_PREFIX") path.file.mkdirs // low-level mkdirs to avoid recursion via Isabelle environment File.platform_file(path) } def tmp_file(name: String, ext: String = "", base_dir: JFile = isabelle_tmp_prefix()): JFile = { val suffix = if (ext == "") "" else "." + ext val file = Files.createTempFile(base_dir.toPath, name, suffix).toFile file.deleteOnExit file } def with_tmp_file[A](name: String, ext: String = "")(body: Path => A): A = { val file = tmp_file(name, ext) try { body(File.path(file)) } finally { file.delete } } /* tmp dirs */ def rm_tree(root: JFile): Unit = { root.delete if (root.isDirectory) { Files.walkFileTree(root.toPath, new SimpleFileVisitor[JPath] { override def visitFile(file: JPath, attrs: BasicFileAttributes): FileVisitResult = { try { Files.deleteIfExists(file) } catch { case _: IOException => } FileVisitResult.CONTINUE } override def postVisitDirectory(dir: JPath, e: IOException): FileVisitResult = { if (e == null) { try { Files.deleteIfExists(dir) } catch { case _: IOException => } FileVisitResult.CONTINUE } else throw e } } ) } } def rm_tree(root: Path): Unit = rm_tree(root.file) object Rm_Tree extends Scala.Fun("rm_tree") { val here = Scala_Project.here def apply(arg: String): String = apply_paths1(arg, rm_tree) } def tmp_dir(name: String, base_dir: JFile = isabelle_tmp_prefix()): JFile = { val dir = Files.createTempDirectory(base_dir.toPath, name).toFile dir.deleteOnExit dir } def with_tmp_dir[A](name: String)(body: Path => A): A = { val dir = tmp_dir(name) try { body(File.path(dir)) } finally { rm_tree(dir) } } /* quasi-atomic update of directory */ def update_directory(dir: Path, f: Path => Unit): Unit = { val new_dir = dir.ext("new") val old_dir = dir.ext("old") rm_tree(new_dir) rm_tree(old_dir) f(new_dir) if (dir.is_dir) move_file(dir, old_dir) move_file(new_dir, dir) rm_tree(old_dir) } /** external processes **/ /* raw process */ def process(command_line: List[String], cwd: JFile = null, env: Map[String, String] = settings(), redirect: Boolean = false): Process = { val proc = new ProcessBuilder proc.command(command_line:_*) // fragile on Windows if (cwd != null) proc.directory(cwd) if (env != null) { - proc.environment.clear + proc.environment.clear() for ((x, y) <- env) proc.environment.put(x, y) } proc.redirectErrorStream(redirect) proc.start } def process_output(proc: Process): (String, Int) = { - proc.getOutputStream.close + proc.getOutputStream.close() val output = File.read_stream(proc.getInputStream) val rc = try { proc.waitFor } finally { - proc.getInputStream.close - proc.getErrorStream.close - proc.destroy + proc.getInputStream.close() + proc.getErrorStream.close() + proc.destroy() Exn.Interrupt.dispose() } (output, rc) } def kill(signal: String, group_pid: String): (String, Int) = { val bash = if (Platform.is_windows) List(cygwin_root() + "\\bin\\bash.exe") else List("/usr/bin/env", "bash") process_output(process(bash ::: List("-c", "kill -" + signal + " -" + group_pid))) } /* GNU bash */ def bash(script: String, cwd: JFile = null, env: Map[String, String] = settings(), redirect: Boolean = false, progress_stdout: String => Unit = (_: String) => (), progress_stderr: String => Unit = (_: String) => (), watchdog: Option[Bash.Watchdog] = None, strict: Boolean = true, cleanup: () => Unit = () => ()): Process_Result = { Bash.process(script, cwd = cwd, env = env, redirect = redirect, cleanup = cleanup). result(progress_stdout = progress_stdout, progress_stderr = progress_stderr, watchdog = watchdog, strict = strict) } private lazy val gnutar_check: Boolean = try { bash("tar --version").check.out.containsSlice("GNU tar") || error("") } catch { case ERROR(_) => false } def gnutar( args: String, dir: Path = Path.current, original_owner: Boolean = false, redirect: Boolean = false): Process_Result = { val options = (if (dir.is_current) "" else "-C " + File.bash_path(dir) + " ") + (if (original_owner) "" else "--owner=root --group=staff ") if (gnutar_check) bash("tar " + options + args, redirect = redirect) else error("Expected to find GNU tar executable") } def require_command(cmds: String*): Unit = { for (cmd <- cmds) { if (!bash(Bash.string(cmd) + " --version").ok) error("Missing command: " + quote(cmd)) } } def hostname(): String = bash("hostname -s").check.out def open(arg: String): Unit = bash("exec \"$ISABELLE_OPEN\" " + Bash.string(arg) + " >/dev/null 2>/dev/null &") def pdf_viewer(arg: Path): Unit = bash("exec \"$PDF_VIEWER\" " + File.bash_path(arg) + " >/dev/null 2>/dev/null &") def open_external_file(name: String): Boolean = { val ext = Library.take_suffix((c: Char) => c != '.', name.toList)._2.mkString val external = ext.nonEmpty && Library.space_explode(':', getenv("ISABELLE_EXTERNAL_FILES")).contains(ext) if (external) { if (ext == "pdf" && Path.is_wellformed(name)) pdf_viewer(Path.explode(name)) else open(name) } external } def export_isabelle_identifier(isabelle_identifier: String): String = if (isabelle_identifier == "") "" else "export ISABELLE_IDENTIFIER=" + Bash.string(isabelle_identifier) + "\n" /** Isabelle resources **/ /* repository clone with Admin */ def admin(): Boolean = Path.explode("~~/Admin").is_dir /* components */ def components(): List[Path] = Path.split(getenv_strict("ISABELLE_COMPONENTS")) /* default logic */ def default_logic(args: String*): String = { args.find(_ != "") match { case Some(logic) => logic case None => getenv_strict("ISABELLE_LOGIC") } } /* download file */ def download(url_name: String, file: Path, progress: Progress = new Progress): Unit = { val url = Url(url_name) progress.echo("Getting " + quote(url_name)) val bytes = try { Url.read_bytes(url) } catch { case ERROR(msg) => cat_error("Failed to download " + quote(url_name), msg) } Bytes.write(file, bytes) } object Download extends Scala.Fun("download") { val here = Scala_Project.here def apply(arg: String): String = Library.split_strings0(arg) match { case List(url, file) => download(url, Path.explode(file)); "" } } } diff --git a/src/Pure/System/isabelle_tool.scala b/src/Pure/System/isabelle_tool.scala --- a/src/Pure/System/isabelle_tool.scala +++ b/src/Pure/System/isabelle_tool.scala @@ -1,228 +1,228 @@ /* Title: Pure/System/isabelle_tool.scala Author: Makarius Author: Lars Hupel Isabelle system tools: external executables or internal Scala functions. */ package isabelle import java.net.URLClassLoader import scala.reflect.runtime.universe import scala.tools.reflect.{ToolBox, ToolBoxError} object Isabelle_Tool { /* Scala source tools */ abstract class Body extends Function[List[String], Unit] private def compile(path: Path): Body = { def err(msg: String): Nothing = cat_error(msg, "The error(s) above occurred in Isabelle/Scala tool " + path) val source = File.read(path) val class_loader = new URLClassLoader(Array(), getClass.getClassLoader) val tool_box = universe.runtimeMirror(class_loader).mkToolBox() try { val symbol = tool_box.parse(source) match { case tree: universe.ModuleDef => tool_box.define(tree) case _ => err("Source does not describe a module (Scala object)") } tool_box.compile(universe.Ident(symbol))() match { case body: Body => body case _ => err("Ill-typed source: Isabelle_Tool.Body expected") } } catch { case e: ToolBoxError => if (tool_box.frontEnd.hasErrors) { val infos = tool_box.frontEnd.infos.toList val msgs = infos.map(info => "Error in line " + info.pos.line + ":\n" + info.msg) err(msgs.mkString("\n")) } else err(e.toString) } } /* external tools */ private def dirs(): List[Path] = Path.split(Isabelle_System.getenv_strict("ISABELLE_TOOLS")) private def is_external(dir: Path, file_name: String): Boolean = { val file = (dir + Path.explode(file_name)).file try { file.isFile && file.canRead && (file_name.endsWith(".scala") || file.canExecute) && !file_name.endsWith("~") && !file_name.endsWith(".orig") } catch { case _: SecurityException => false } } private def find_external(name: String): Option[List[String] => Unit] = dirs().collectFirst({ case dir if is_external(dir, name + ".scala") => compile(dir + Path.explode(name + ".scala")) case dir if is_external(dir, name) => (args: List[String]) => { val tool = dir + Path.explode(name) val result = Isabelle_System.bash(File.bash_path(tool) + " " + Bash.strings(args)) sys.exit(result.print_stdout.rc) } }) /* internal tools */ private lazy val internal_tools: List[Isabelle_Tool] = Isabelle_System.make_services(classOf[Isabelle_Scala_Tools]).flatMap(_.tools) private def find_internal(name: String): Option[List[String] => Unit] = internal_tools.collectFirst({ case tool if tool.name == name => args => Command_Line.tool { tool.body(args) } }) /* list tools */ abstract class Entry { def name: String def position: Properties.T def description: String def print: String = description match { case "" => name case descr => name + " - " + descr } } sealed case class External(name: String, path: Path) extends Entry { def position: Properties.T = Position.File(path.absolute.implode) def description: String = { val Pattern = """.*\bDESCRIPTION: *(.*)""".r split_lines(File.read(path)).collectFirst({ case Pattern(s) => s }) getOrElse "" } } def external_tools(): List[External] = { for { dir <- dirs() if dir.is_dir file_name <- File.read_dir(dir) if is_external(dir, file_name) } yield { val path = dir + Path.explode(file_name) val name = Library.perhaps_unsuffix(".scala", file_name) External(name, path) } } def isabelle_tools(): List[Entry] = (external_tools() ::: internal_tools).sortBy(_.name) object Isabelle_Tools extends Scala.Fun("isabelle_tools") { val here = Scala_Project.here def apply(arg: String): String = if (arg.nonEmpty) error("Bad argument: " + quote(arg)) else { val result = isabelle_tools().map(entry => (entry.name, entry.position)) val body = { import XML.Encode._; list(pair(string, properties))(result) } YXML.string_of_body(body) } } /* command line entry point */ def main(args: Array[String]): Unit = { Command_Line.tool { args.toList match { case Nil | List("-?") => val tool_descriptions = isabelle_tools().map(_.print) Getopts(""" Usage: isabelle TOOL [ARGS ...] Start Isabelle TOOL with ARGS; pass "-?" for tool-specific help. -Available tools:""" + tool_descriptions.mkString("\n ", "\n ", "\n")).usage +Available tools:""" + tool_descriptions.mkString("\n ", "\n ", "\n")).usage() case tool_name :: tool_args => find_external(tool_name) orElse find_internal(tool_name) match { case Some(tool) => tool(tool_args) case None => error("Unknown Isabelle tool: " + quote(tool_name)) } } } } } sealed case class Isabelle_Tool( name: String, description: String, here: Scala_Project.Here, body: List[String] => Unit) extends Isabelle_Tool.Entry { def position: Position.T = here.position } class Isabelle_Scala_Tools(val tools: Isabelle_Tool*) extends Isabelle_System.Service class Tools extends Isabelle_Scala_Tools( Build.isabelle_tool, Build_Docker.isabelle_tool, Build_Job.isabelle_tool, Doc.isabelle_tool, Dump.isabelle_tool, Export.isabelle_tool, ML_Process.isabelle_tool, Mercurial.isabelle_tool, Mkroot.isabelle_tool, Options.isabelle_tool, Phabricator.isabelle_tool1, Phabricator.isabelle_tool2, Phabricator.isabelle_tool3, Phabricator.isabelle_tool4, Presentation.isabelle_tool, Profiling_Report.isabelle_tool, Server.isabelle_tool, Sessions.isabelle_tool, Scala_Project.isabelle_tool, Update.isabelle_tool, Update_Cartouches.isabelle_tool, Update_Comments.isabelle_tool, Update_Header.isabelle_tool, Update_Then.isabelle_tool, Update_Theorems.isabelle_tool, isabelle.vscode.TextMate_Grammar.isabelle_tool, isabelle.vscode.Language_Server.isabelle_tool) class Admin_Tools extends Isabelle_Scala_Tools( Build_CSDP.isabelle_tool, Build_Cygwin.isabelle_tool, Build_Doc.isabelle_tool, Build_E.isabelle_tool, Build_Fonts.isabelle_tool, Build_JDK.isabelle_tool, Build_PolyML.isabelle_tool1, Build_PolyML.isabelle_tool2, Build_SPASS.isabelle_tool, Build_SQLite.isabelle_tool, Build_Status.isabelle_tool, Build_Vampire.isabelle_tool, Build_VeriT.isabelle_tool, Build_Zipperposition.isabelle_tool, Check_Sources.isabelle_tool, Components.isabelle_tool, isabelle.vscode.Build_VSCode.isabelle_tool) diff --git a/src/Pure/System/posix_interrupt.scala b/src/Pure/System/posix_interrupt.scala --- a/src/Pure/System/posix_interrupt.scala +++ b/src/Pure/System/posix_interrupt.scala @@ -1,29 +1,29 @@ /* Title: Pure/System/posix_interrupt.scala Author: Makarius Support for POSIX interrupts (bypassed on Windows). */ package isabelle import sun.misc.{Signal, SignalHandler} object POSIX_Interrupt { def handler[A](h: => Unit)(e: => A): A = { val SIGINT = new Signal("INT") val new_handler = new SignalHandler { def handle(s: Signal): Unit = h } val old_handler = Signal.handle(SIGINT, new_handler) try { e } finally { Signal.handle(SIGINT, old_handler) } } def exception[A](e: => A): A = { val thread = Thread.currentThread - handler { thread.interrupt } { e } + handler { thread.interrupt() } { e } } } diff --git a/src/Pure/System/progress.scala b/src/Pure/System/progress.scala --- a/src/Pure/System/progress.scala +++ b/src/Pure/System/progress.scala @@ -1,87 +1,87 @@ /* Title: Pure/System/progress.scala Author: Makarius Progress context for system processes. */ package isabelle import java.io.{File => JFile} object Progress { sealed case class Theory(theory: String, session: String = "", percentage: Option[Int] = None) { def message: String = print_session + print_theory + print_percentage def print_session: String = if (session == "") "" else session + ": " def print_theory: String = "theory " + theory def print_percentage: String = percentage match { case None => "" case Some(p) => " " + p + "%" } } } class Progress { def echo(msg: String): Unit = {} def echo_if(cond: Boolean, msg: String): Unit = { if (cond) echo(msg) } def theory(theory: Progress.Theory): Unit = {} def nodes_status(nodes_status: Document_Status.Nodes_Status): Unit = {} def echo_warning(msg: String): Unit = echo(Output.warning_text(msg)) def echo_error_message(msg: String): Unit = echo(Output.error_message_text(msg)) def timeit[A](message: String = "", enabled: Boolean = true)(e: => A): A = Timing.timeit(message, enabled, echo)(e) @volatile protected var is_stopped = false - def stop: Unit = { is_stopped = true } + def stop(): Unit = { is_stopped = true } def stopped: Boolean = { - if (Thread.interrupted) is_stopped = true + if (Thread.interrupted()) is_stopped = true is_stopped } - def interrupt_handler[A](e: => A): A = POSIX_Interrupt.handler { stop } { e } + def interrupt_handler[A](e: => A): A = POSIX_Interrupt.handler { stop() } { e } def expose_interrupt(): Unit = if (stopped) throw Exn.Interrupt() override def toString: String = if (stopped) "Progress(stopped)" else "Progress" def bash(script: String, cwd: JFile = null, env: Map[String, String] = Isabelle_System.settings(), redirect: Boolean = false, echo: Boolean = false, watchdog: Time = Time.zero, strict: Boolean = true): Process_Result = { val result = Isabelle_System.bash(script, cwd = cwd, env = env, redirect = redirect, progress_stdout = echo_if(echo, _), progress_stderr = echo_if(echo, _), watchdog = if (watchdog.is_zero) None else Some((watchdog, _ => stopped)), strict = strict) if (strict && stopped) throw Exn.Interrupt() else result } } class Console_Progress(verbose: Boolean = false, stderr: Boolean = false) extends Progress { override def echo(msg: String): Unit = Output.writeln(msg, stdout = !stderr, include_empty = true) override def theory(theory: Progress.Theory): Unit = if (verbose) echo(theory.message) } class File_Progress(path: Path, verbose: Boolean = false) extends Progress { override def echo(msg: String): Unit = File.append(path, msg + "\n") override def theory(theory: Progress.Theory): Unit = if (verbose) echo(theory.message) override def toString: String = path.toString } diff --git a/src/Pure/System/scala.scala b/src/Pure/System/scala.scala --- a/src/Pure/System/scala.scala +++ b/src/Pure/System/scala.scala @@ -1,255 +1,255 @@ /* Title: Pure/System/scala.scala Author: Makarius Support for Scala at runtime. */ package isabelle import java.io.{File => JFile, StringWriter, PrintWriter} import scala.tools.nsc.{GenericRunnerSettings, ConsoleWriter, NewLinePrintWriter} import scala.tools.nsc.interpreter.{IMain, Results} import scala.tools.nsc.interpreter.shell.ReplReporterImpl object Scala { /** registered functions **/ abstract class Fun(val name: String) extends Function[String, String] { override def toString: String = name def position: Properties.T = here.position def here: Scala_Project.Here def apply(arg: String): String } class Functions(val functions: Fun*) extends Isabelle_System.Service lazy val functions: List[Fun] = Isabelle_System.make_services(classOf[Functions]).flatMap(_.functions) /** demo functions **/ object Echo extends Fun("echo") { val here = Scala_Project.here def apply(arg: String): String = arg } object Sleep extends Fun("sleep") { val here = Scala_Project.here def apply(seconds: String): String = { val t = seconds match { case Value.Double(s) => Time.seconds(s) case _ => error("Malformed argument: " + quote(seconds)) } val t0 = Time.now() t.sleep val t1 = Time.now() (t1 - t0).toString } } /** compiler **/ object Compiler { def context( error: String => Unit = Exn.error, jar_dirs: List[JFile] = Nil): Context = { def find_jars(dir: JFile): List[String] = File.find_files(dir, file => file.getName.endsWith(".jar")). map(File.absolute_name) val class_path = for { prop <- List("isabelle.scala.classpath", "java.class.path") path = System.getProperty(prop, "") if path != "\"\"" elem <- space_explode(JFile.pathSeparatorChar, path) } yield elem val settings = new GenericRunnerSettings(error) settings.classpath.value = (class_path ::: jar_dirs.flatMap(find_jars)).mkString(JFile.pathSeparator) new Context(settings) } def default_print_writer: PrintWriter = new NewLinePrintWriter(new ConsoleWriter, true) class Context private [Compiler](val settings: GenericRunnerSettings) { override def toString: String = settings.toString def interpreter( print_writer: PrintWriter = default_print_writer, class_loader: ClassLoader = null): IMain = { new IMain(settings, new ReplReporterImpl(settings, print_writer)) { override def parentClassLoader: ClassLoader = if (class_loader == null) super.parentClassLoader else class_loader } } def toplevel(interpret: Boolean, source: String): List[String] = { val out = new StringWriter val interp = interpreter(new PrintWriter(out)) val marker = '\u000b' val ok = interp.withLabel(marker.toString) { if (interpret) interp.interpret(source) == Results.Success else (new interp.ReadEvalPrint).compile(source) } - out.close + out.close() val Error = """(?s)^\S* error: (.*)$""".r val errors = space_explode(marker, Library.strip_ansi_color(out.toString)). collect({ case Error(msg) => "Scala error: " + Library.trim_line(msg) }) if (!ok && errors.isEmpty) List("Error") else errors } } } object Toplevel extends Fun("scala_toplevel") { val here = Scala_Project.here def apply(arg: String): String = { val (interpret, source) = YXML.parse_body(arg) match { case Nil => (false, "") case List(XML.Text(source)) => (false, source) case body => import XML.Decode._; pair(bool, string)(body) } val errors = try { Compiler.context().toplevel(interpret, source) } catch { case ERROR(msg) => List(msg) } locally { import XML.Encode._; YXML.string_of_body(list(string)(errors)) } } } /** invoke Scala functions from ML **/ /* invoke function */ object Tag extends Enumeration { val NULL, OK, ERROR, FAIL, INTERRUPT = Value } def function(name: String, arg: String): (Tag.Value, String) = functions.find(fun => fun.name == name) match { case Some(fun) => Exn.capture { fun(arg) } match { case Exn.Res(null) => (Tag.NULL, "") case Exn.Res(res) => (Tag.OK, res) case Exn.Exn(Exn.Interrupt()) => (Tag.INTERRUPT, "") case Exn.Exn(e) => (Tag.ERROR, Exn.message(e)) } case None => (Tag.FAIL, "Unknown Isabelle/Scala function: " + quote(name)) } /* protocol handler */ class Handler extends Session.Protocol_Handler { private var session: Session = null private var futures = Map.empty[String, Future[Unit]] override def init(session: Session): Unit = synchronized { this.session = session } override def exit(): Unit = synchronized { for ((id, future) <- futures) cancel(id, future) futures = Map.empty } private def result(id: String, tag: Scala.Tag.Value, res: String): Unit = synchronized { if (futures.isDefinedAt(id)) { session.protocol_command("Scala.result", id, tag.id.toString, res) futures -= id } } private def cancel(id: String, future: Future[Unit]): Unit = { - future.cancel + future.cancel() result(id, Scala.Tag.INTERRUPT, "") } private def invoke_scala(msg: Prover.Protocol_Output): Boolean = synchronized { msg.properties match { case Markup.Invoke_Scala(name, id, thread) => def body: Unit = { val (tag, res) = Scala.function(name, msg.text) result(id, tag, res) } val future = if (thread) { Future.thread(name = Isabelle_Thread.make_name(base = "invoke_scala"))(body) } else Future.fork(body) futures += (id -> future) true case _ => false } } private def cancel_scala(msg: Prover.Protocol_Output): Boolean = synchronized { msg.properties match { case Markup.Cancel_Scala(id) => futures.get(id) match { case Some(future) => cancel(id, future) case None => } true case _ => false } } override val functions = List( Markup.Invoke_Scala.name -> invoke_scala, Markup.Cancel_Scala.name -> cancel_scala) } } class Scala_Functions extends Scala.Functions( Scala.Echo, Scala.Sleep, Scala.Toplevel, Doc.Doc_Names, Bash.Process, Bibtex.Check_Database, Isabelle_System.Make_Directory, Isabelle_System.Copy_Dir, Isabelle_System.Copy_File, Isabelle_System.Copy_File_Base, Isabelle_System.Rm_Tree, Isabelle_System.Download, Isabelle_Tool.Isabelle_Tools) diff --git a/src/Pure/System/system_channel.scala b/src/Pure/System/system_channel.scala --- a/src/Pure/System/system_channel.scala +++ b/src/Pure/System/system_channel.scala @@ -1,46 +1,46 @@ /* Title: Pure/System/system_channel.scala Author: Makarius Socket-based system channel for inter-process communication. */ package isabelle import java.io.{InputStream, OutputStream} import java.net.{ServerSocket, InetAddress} object System_Channel { def apply(): System_Channel = new System_Channel } class System_Channel private { private val server = new ServerSocket(0, 50, Server.localhost) val address: String = Server.print_address(server.getLocalPort) val password: String = UUID.random().toString override def toString: String = address - def shutdown(): Unit = server.close + def shutdown(): Unit = server.close() def rendezvous(): (OutputStream, InputStream) = { val socket = server.accept try { val out_stream = socket.getOutputStream val in_stream = socket.getInputStream if (Byte_Message.read_line(in_stream).map(_.text) == Some(password)) (out_stream, in_stream) else { - out_stream.close - in_stream.close + out_stream.close() + in_stream.close() error("Failed to connect system channel: bad password") } } finally { shutdown() } } } diff --git a/src/Pure/System/tty_loop.scala b/src/Pure/System/tty_loop.scala --- a/src/Pure/System/tty_loop.scala +++ b/src/Pure/System/tty_loop.scala @@ -1,68 +1,68 @@ /* Title: Pure/System/tty_loop.scala Author: Makarius Line-oriented TTY loop. */ package isabelle import java.io.{IOException, Writer, Reader, InputStreamReader, BufferedReader} class TTY_Loop( writer: Writer, reader: Reader, writer_lock: AnyRef = new Object) { private val console_output = Future.thread[Unit]("console_output", uninterruptible = true) { try { val result = new StringBuilder(100) var finished = false while (!finished) { var c = -1 var done = false while (!done && (result.isEmpty || reader.ready)) { c = reader.read if (c >= 0) result.append(c.asInstanceOf[Char]) else done = true } if (result.nonEmpty) { System.out.print(result.toString) System.out.flush() - result.clear + result.clear() } else { reader.close() finished = true } } } catch { case e: IOException => case Exn.Interrupt() => } } private val console_input = Future.thread[Unit]("console_input", uninterruptible = true) { val console_reader = new BufferedReader(new InputStreamReader(System.in)) try { var finished = false while (!finished) { console_reader.readLine() match { case null => writer.close() finished = true case line => writer_lock.synchronized { writer.write(line) writer.write("\n") writer.flush() } } } } catch { case e: IOException => case Exn.Interrupt() => } } def join: Unit = { console_output.join; console_input.join } - def cancel: Unit = console_input.cancel + def cancel(): Unit = console_input.cancel() } diff --git a/src/Pure/Thy/file_format.scala b/src/Pure/Thy/file_format.scala --- a/src/Pure/Thy/file_format.scala +++ b/src/Pure/Thy/file_format.scala @@ -1,98 +1,98 @@ /* Title: Pure/Thy/file_format.scala Author: Makarius Support for user-defined file formats, associated with active session. */ package isabelle object File_Format { sealed case class Theory_Context(name: Document.Node.Name, content: String) /* registry */ lazy val registry: Registry = new Registry(Isabelle_System.make_services(classOf[File_Format])) final class Registry private [File_Format](file_formats: List[File_Format]) { override def toString: String = file_formats.mkString("File_Format.Environment(", ",", ")") def get(name: String): Option[File_Format] = file_formats.find(_.detect(name)) def get(name: Document.Node.Name): Option[File_Format] = get(name.node) def get_theory(name: Document.Node.Name): Option[File_Format] = get(name.theory) def is_theory(name: Document.Node.Name): Boolean = get_theory(name).isDefined def start_session(session: isabelle.Session): Session = new Session(file_formats.map(_.start(session)).filterNot(_ == No_Agent)) } /* session */ final class Session private[File_Format](agents: List[Agent]) { override def toString: String = agents.mkString("File_Format.Session(", ", ", ")") def prover_options(options: Options): Options = agents.foldLeft(options) { case (opts, agent) => agent.prover_options(opts) } - def stop_session: Unit = agents.foreach(_.stop) + def stop_session: Unit = agents.foreach(_.stop()) } trait Agent { def prover_options(options: Options): Options = options - def stop: Unit = {} + def stop(): Unit = {} } object No_Agent extends Agent } abstract class File_Format extends Isabelle_System.Service { def format_name: String override def toString: String = "File_Format(" + format_name + ")" def file_ext: String def detect(name: String): Boolean = name.endsWith("." + file_ext) /* implicit theory context: name and content */ def theory_suffix: String = "" def theory_content(name: String): String = "" def make_theory_name(resources: Resources, name: Document.Node.Name): Option[Document.Node.Name] = { for { (_, thy) <- Thy_Header.split_file_name(name.node) if detect(name.node) && theory_suffix.nonEmpty } yield { val thy_node = resources.append(name.node, Path.explode(theory_suffix)) Document.Node.Name(thy_node, name.master_dir, thy) } } def make_theory_content(resources: Resources, thy_name: Document.Node.Name): Option[String] = { for { (prefix, suffix) <- Thy_Header.split_file_name(thy_name.node) if detect(prefix) && suffix == theory_suffix (_, thy) <- Thy_Header.split_file_name(prefix) s <- proper_string(theory_content(thy)) } yield s } def html_document(snapshot: Document.Snapshot): Option[Presentation.HTML_Document] = None /* PIDE session agent */ def start(session: isabelle.Session): File_Format.Agent = File_Format.No_Agent } diff --git a/src/Pure/Thy/sessions.scala b/src/Pure/Thy/sessions.scala --- a/src/Pure/Thy/sessions.scala +++ b/src/Pure/Thy/sessions.scala @@ -1,1520 +1,1520 @@ /* Title: Pure/Thy/sessions.scala Author: Makarius Cumulative session information. */ package isabelle import java.io.{File => JFile} import java.nio.ByteBuffer import java.nio.channels.FileChannel import java.nio.file.StandardOpenOption import scala.collection.immutable.{SortedSet, SortedMap} import scala.collection.mutable object Sessions { /* session and theory names */ val ROOTS: Path = Path.explode("ROOTS") val ROOT: Path = Path.explode("ROOT") val roots_name: String = "ROOTS" val root_name: String = "ROOT" val theory_name: String = "Pure.Sessions" val UNSORTED = "Unsorted" val DRAFT = "Draft" def is_pure(name: String): Boolean = name == Thy_Header.PURE def exclude_session(name: String): Boolean = name == "" || name == DRAFT def exclude_theory(name: String): Boolean = name == root_name || name == "README" || name == "index" || name == "bib" /* ROOTS file format */ class File_Format extends isabelle.File_Format { val format_name: String = roots_name val file_ext = "" override def detect(name: String): Boolean = Thy_Header.split_file_name(name) match { case Some((_, file_name)) => file_name == roots_name case None => false } override def theory_suffix: String = "ROOTS_file" override def theory_content(name: String): String = """theory "ROOTS" imports Pure begin ROOTS_file """ + Outer_Syntax.quote_string(name) + """ end""" } /* base info and source dependencies */ sealed case class Base( pos: Position.T = Position.none, session_directories: Map[JFile, String] = Map.empty, global_theories: Map[String, String] = Map.empty, session_theories: List[Document.Node.Name] = Nil, document_theories: List[Document.Node.Name] = Nil, loaded_theories: Graph[String, Outer_Syntax] = Graph.string, used_theories: List[(Document.Node.Name, Options)] = Nil, load_commands: Map[Document.Node.Name, List[Command_Span.Span]] = Map.empty, known_theories: Map[String, Document.Node.Entry] = Map.empty, known_loaded_files: Map[String, List[Path]] = Map.empty, overall_syntax: Outer_Syntax = Outer_Syntax.empty, imported_sources: List[(Path, SHA1.Digest)] = Nil, sources: List[(Path, SHA1.Digest)] = Nil, session_graph_display: Graph_Display.Graph = Graph_Display.empty_graph, errors: List[String] = Nil) { override def toString: String = "Sessions.Base(loaded_theories = " + loaded_theories.size + ", used_theories = " + used_theories.length + ")" def theory_qualifier(name: String): String = global_theories.getOrElse(name, Long_Name.qualifier(name)) def theory_qualifier(name: Document.Node.Name): String = theory_qualifier(name.theory) def loaded_theory(name: String): Boolean = loaded_theories.defined(name) def loaded_theory(name: Document.Node.Name): Boolean = loaded_theory(name.theory) def loaded_theory_syntax(name: String): Option[Outer_Syntax] = if (loaded_theory(name)) Some(loaded_theories.get_node(name)) else None def loaded_theory_syntax(name: Document.Node.Name): Option[Outer_Syntax] = loaded_theory_syntax(name.theory) def theory_syntax(name: Document.Node.Name): Outer_Syntax = loaded_theory_syntax(name) getOrElse overall_syntax def node_syntax(nodes: Document.Nodes, name: Document.Node.Name): Outer_Syntax = nodes(name).syntax orElse loaded_theory_syntax(name) getOrElse overall_syntax } sealed case class Deps(sessions_structure: Structure, session_bases: Map[String, Base]) { override def toString: String = "Sessions.Deps(" + sessions_structure + ")" def is_empty: Boolean = session_bases.isEmpty def apply(name: String): Base = session_bases(name) def get(name: String): Option[Base] = session_bases.get(name) def imported_sources(name: String): List[SHA1.Digest] = session_bases(name).imported_sources.map(_._2) def sources(name: String): List[SHA1.Digest] = session_bases(name).sources.map(_._2) def errors: List[String] = (for { (name, base) <- session_bases.iterator if base.errors.nonEmpty } yield cat_lines(base.errors) + "\nThe error(s) above occurred in session " + quote(name) + Position.here(base.pos) ).toList def check_errors: Deps = errors match { case Nil => this case errs => error(cat_lines(errs)) } } def deps(sessions_structure: Structure, progress: Progress = new Progress, inlined_files: Boolean = false, verbose: Boolean = false, list_files: Boolean = false, check_keywords: Set[String] = Set.empty): Deps = { var cache_sources = Map.empty[JFile, SHA1.Digest] def check_sources(paths: List[Path]): List[(Path, SHA1.Digest)] = { for { path <- paths file = path.file if cache_sources.isDefinedAt(file) || file.isFile } yield { cache_sources.get(file) match { case Some(digest) => (path, digest) case None => val digest = SHA1.digest(file) cache_sources = cache_sources + (file -> digest) (path, digest) } } } val session_bases = sessions_structure.imports_topological_order.foldLeft(Map("" -> sessions_structure.bootstrap)) { case (session_bases, session_name) => progress.expose_interrupt() val info = sessions_structure(session_name) try { val deps_base = info.deps_base(session_bases) val resources = new Resources(sessions_structure, deps_base) if (verbose || list_files) { val groups = if (info.groups.isEmpty) "" else info.groups.mkString(" (", " ", ")") progress.echo("Session " + info.chapter_session + groups) } val dependencies = resources.session_dependencies(info) val overall_syntax = dependencies.overall_syntax val session_theories = dependencies.theories.filter(name => deps_base.theory_qualifier(name) == session_name) val theory_files = dependencies.theories.map(_.path) dependencies.load_commands val (load_commands, load_commands_errors) = try { if (inlined_files) (dependencies.load_commands, Nil) else (Nil, Nil) } catch { case ERROR(msg) => (Nil, List(msg)) } val loaded_files = load_commands.map({ case (name, spans) => dependencies.loaded_files(name, spans) }) val session_files = (theory_files ::: loaded_files.flatMap(_._2) ::: info.document_files.map(file => info.dir + file._1 + file._2)).map(_.expand) val imported_files = if (inlined_files) dependencies.imported_files else Nil if (list_files) progress.echo(cat_lines(session_files.map(_.implode).sorted.map(" " + _))) if (check_keywords.nonEmpty) { Check_Keywords.check_keywords( progress, overall_syntax.keywords, check_keywords, theory_files) } val session_graph_display: Graph_Display.Graph = { def session_node(name: String): Graph_Display.Node = Graph_Display.Node("[" + name + "]", "session." + name) def node(name: Document.Node.Name): Graph_Display.Node = { val qualifier = deps_base.theory_qualifier(name) if (qualifier == info.name) Graph_Display.Node(name.theory_base_name, "theory." + name.theory) else session_node(qualifier) } val required_sessions = dependencies.loaded_theories.all_preds(dependencies.theories.map(_.theory)) .map(theory => deps_base.theory_qualifier(theory)) .filter(name => name != info.name && sessions_structure.defined(name)) val required_subgraph = sessions_structure.imports_graph .restrict(sessions_structure.imports_graph.all_preds(required_sessions).toSet) .transitive_closure .restrict(required_sessions.toSet) .transitive_reduction_acyclic val graph0 = required_subgraph.topological_order.foldLeft(Graph_Display.empty_graph) { case (g, session) => val a = session_node(session) val bs = required_subgraph.imm_preds(session).toList.map(session_node) bs.foldLeft((a :: bs).foldLeft(g)(_.default_node(_, Nil)))(_.add_edge(_, a)) } dependencies.entries.foldLeft(graph0) { case (g, entry) => val a = node(entry.name) val bs = entry.header.imports.map(node).filterNot(_ == a) bs.foldLeft((a :: bs).foldLeft(g)(_.default_node(_, Nil)))(_.add_edge(_, a)) } } val known_theories = dependencies.entries.iterator.map(entry => entry.name.theory -> entry). foldLeft(deps_base.known_theories)(_ + _) val known_loaded_files = deps_base.known_loaded_files ++ loaded_files val import_errors = { val known_sessions = sessions_structure.imports_requirements(List(session_name)).toSet for { name <- dependencies.theories qualifier = deps_base.theory_qualifier(name) if !known_sessions(qualifier) } yield "Bad import of theory " + quote(name.toString) + ": need to include sessions " + quote(qualifier) + " in ROOT" } val document_errors = info.document_theories.flatMap( { case (thy, pos) => val parent_sessions = if (sessions_structure.build_graph.defined(session_name)) { sessions_structure.build_requirements(List(session_name)) } else Nil def err(msg: String): Option[String] = Some(msg + " " + quote(thy) + Position.here(pos)) known_theories.get(thy).map(_.name) match { case None => err("Unknown document theory") case Some(name) => val qualifier = deps_base.theory_qualifier(name) if (session_theories.contains(name)) { err("Redundant document theory from this session:") } else if (parent_sessions.contains(qualifier)) None else if (dependencies.theories.contains(name)) None else err("Document theory from other session not imported properly:") } }) val document_theories = info.document_theories.map({ case (thy, _) => known_theories(thy).name }) val dir_errors = { val ok = info.dirs.map(_.canonical_file).toSet val bad = (for { name <- session_theories.iterator path = name.master_dir_path if !ok(path.canonical_file) path1 = File.relative_path(info.dir.canonical, path).getOrElse(path) } yield (path1, name)).toList val bad_dirs = (for { (path1, _) <- bad } yield path1.toString).distinct.sorted val errs1 = for { (path1, name) <- bad } yield "Implicit use of directory " + path1 + " for theory " + quote(name.toString) val errs2 = if (bad_dirs.isEmpty) Nil else List("Implicit use of session directories: " + commas(bad_dirs)) val errs3 = for (p <- info.dirs if !p.is_dir) yield "No such directory: " + p val errs4 = (for { name <- session_theories.iterator name1 <- resources.find_theory_node(name.theory) if name.node != name1.node } yield "Incoherent theory file import:\n " + name.path + " vs. \n " + name1.path) .toList errs1 ::: errs2 ::: errs3 ::: errs4 } val sources_errors = for (p <- session_files if !p.is_file) yield "No such file: " + p val path_errors = try { Path.check_case_insensitive(session_files ::: imported_files); Nil } catch { case ERROR(msg) => List(msg) } val bibtex_errors = try { info.bibtex_entries; Nil } catch { case ERROR(msg) => List(msg) } val base = Base( pos = info.pos, session_directories = sessions_structure.session_directories, global_theories = sessions_structure.global_theories, session_theories = session_theories, document_theories = document_theories, loaded_theories = dependencies.loaded_theories, used_theories = dependencies.theories_adjunct, load_commands = load_commands.toMap, known_theories = known_theories, known_loaded_files = known_loaded_files, overall_syntax = overall_syntax, imported_sources = check_sources(imported_files), sources = check_sources(session_files), session_graph_display = session_graph_display, errors = dependencies.errors ::: load_commands_errors ::: import_errors ::: document_errors ::: dir_errors ::: sources_errors ::: path_errors ::: bibtex_errors) session_bases + (info.name -> base) } catch { case ERROR(msg) => cat_error(msg, "The error(s) above occurred in session " + quote(info.name) + Position.here(info.pos)) } } Deps(sessions_structure, session_bases) } /* base info */ sealed case class Base_Info( session: String, sessions_structure: Structure, errors: List[String], base: Base, infos: List[Info]) { def check: Base_Info = if (errors.isEmpty) this else error(cat_lines(errors)) } def base_info(options: Options, session: String, progress: Progress = new Progress, dirs: List[Path] = Nil, include_sessions: List[String] = Nil, session_ancestor: Option[String] = None, session_requirements: Boolean = false): Base_Info = { val full_sessions = load_structure(options, dirs = dirs) val selected_sessions = full_sessions.selection(Selection(sessions = session :: session_ancestor.toList)) val info = selected_sessions(session) val ancestor = session_ancestor orElse info.parent val (session1, infos1) = if (session_requirements && ancestor.isDefined) { val deps = Sessions.deps(selected_sessions, progress = progress) val base = deps(session) val ancestor_loaded = deps.get(ancestor.get) match { case Some(ancestor_base) if !selected_sessions.imports_requirements(List(ancestor.get)).contains(session) => ancestor_base.loaded_theories.defined _ case _ => error("Bad ancestor " + quote(ancestor.get) + " for session " + quote(session)) } val required_theories = for { thy <- base.loaded_theories.keys if !ancestor_loaded(thy) && base.theory_qualifier(thy) != session } yield thy if (required_theories.isEmpty) (ancestor.get, Nil) else { val other_name = info.name + "_requirements(" + ancestor.get + ")" Isabelle_System.isabelle_tmp_prefix() (other_name, List( make_info(info.options, dir_selected = false, dir = Path.explode("$ISABELLE_TMP_PREFIX"), chapter = info.chapter, Session_Entry( pos = info.pos, name = other_name, groups = info.groups, path = ".", parent = ancestor, description = "Required theory imports from other sessions", options = Nil, imports = info.deps, directories = Nil, theories = List((Nil, required_theories.map(thy => ((thy, Position.none), false)))), document_theories = Nil, document_files = Nil, export_files = Nil)))) } } else (session, Nil) val full_sessions1 = if (infos1.isEmpty) full_sessions else load_structure(options, dirs = dirs, infos = infos1) val selected_sessions1 = full_sessions1.selection(Selection(sessions = session1 :: session :: include_sessions)) val deps1 = Sessions.deps(selected_sessions1, progress = progress) Base_Info(session1, full_sessions1, deps1.errors, deps1(session1), infos1) } /* cumulative session info */ sealed case class Info( name: String, chapter: String, dir_selected: Boolean, pos: Position.T, groups: List[String], dir: Path, parent: Option[String], description: String, directories: List[Path], options: Options, imports: List[String], theories: List[(Options, List[(String, Position.T)])], global_theories: List[String], document_theories: List[(String, Position.T)], document_files: List[(Path, Path)], export_files: List[(Path, Int, List[String])], meta_digest: SHA1.Digest) { def chapter_session: String = chapter + "/" + name def relative_path(info1: Info): String = if (name == info1.name) "" else if (chapter == info1.chapter) "../" + info1.name + "/" else "../../" + info1.chapter_session + "/" def deps: List[String] = parent.toList ::: imports def deps_base(session_bases: String => Base): Base = { val parent_base = session_bases(parent.getOrElse("")) val imports_bases = imports.map(session_bases) parent_base.copy( known_theories = (for { base <- imports_bases.iterator (_, entry) <- base.known_theories.iterator } yield (entry.name.theory -> entry)).foldLeft(parent_base.known_theories)(_ + _), known_loaded_files = imports_bases.iterator.map(_.known_loaded_files). foldLeft(parent_base.known_loaded_files)(_ ++ _)) } def dirs: List[Path] = dir :: directories def timeout: Time = Time.seconds(options.real("timeout") * options.real("timeout_scale")) def document_enabled: Boolean = options.string("document") match { case "" | "false" => false case "pdf" => true case doc => error("Bad document specification " + quote(doc)) } def document_variants: List[Presentation.Document_Variant] = { val variants = Library.space_explode(':', options.string("document_variants")). map(Presentation.Document_Variant.parse) val dups = Library.duplicates(variants.map(_.name)) if (dups.nonEmpty) error("Duplicate document variants: " + commas_quote(dups)) variants } def documents: List[Presentation.Document_Variant] = { val variants = document_variants if (!document_enabled || document_files.isEmpty) Nil else variants } def document_output: Option[Path] = options.string("document_output") match { case "" => None case s => Some(dir + Path.explode(s)) } def browser_info: Boolean = options.bool("browser_info") lazy val bibtex_entries: List[Text.Info[String]] = (for { (document_dir, file) <- document_files.iterator if Bibtex.is_bibtex(file.file_name) info <- Bibtex.entries(File.read(dir + document_dir + file)).iterator } yield info).toList def record_proofs: Boolean = options.int("record_proofs") >= 2 def is_afp: Boolean = chapter == AFP.chapter def is_afp_bulky: Boolean = is_afp && groups.exists(AFP.groups_bulky.contains) } def make_info(options: Options, dir_selected: Boolean, dir: Path, chapter: String, entry: Session_Entry): Info = { try { val name = entry.name if (exclude_session(name)) error("Bad session name") if (is_pure(name) && entry.parent.isDefined) error("Illegal parent session") if (!is_pure(name) && !entry.parent.isDefined) error("Missing parent session") val session_path = dir + Path.explode(entry.path) val directories = entry.directories.map(dir => session_path + Path.explode(dir)) val session_options = options ++ entry.options val theories = entry.theories.map({ case (opts, thys) => (session_options ++ opts, thys.map({ case ((thy, pos), _) => if (exclude_theory(thy)) error("Bad theory name " + quote(thy) + Position.here(pos)) else (thy, pos) })) }) val global_theories = for { (_, thys) <- entry.theories; ((thy, pos), global) <- thys if global } yield { val thy_name = Path.explode(thy).file_name if (Long_Name.is_qualified(thy_name)) error("Bad qualified name for global theory " + quote(thy_name) + Position.here(pos)) else thy_name } val conditions = theories.flatMap(thys => space_explode(',', thys._1.string("condition"))).distinct.sorted. map(x => (x, Isabelle_System.getenv(x) != "")) val document_files = entry.document_files.map({ case (s1, s2) => (Path.explode(s1), Path.explode(s2)) }) val export_files = entry.export_files.map({ case (dir, prune, pats) => (Path.explode(dir), prune, pats) }) val meta_digest = SHA1.digest( (name, chapter, entry.parent, entry.directories, entry.options, entry.imports, entry.theories_no_position, conditions, entry.document_theories_no_position, entry.document_files) .toString) Info(name, chapter, dir_selected, entry.pos, entry.groups, session_path, entry.parent, entry.description, directories, session_options, entry.imports, theories, global_theories, entry.document_theories, document_files, export_files, meta_digest) } catch { case ERROR(msg) => error(msg + "\nThe error(s) above occurred in session entry " + quote(entry.name) + Position.here(entry.pos)) } } object Selection { val empty: Selection = Selection() val all: Selection = Selection(all_sessions = true) def session(session: String): Selection = Selection(sessions = List(session)) } sealed case class Selection( requirements: Boolean = false, all_sessions: Boolean = false, base_sessions: List[String] = Nil, exclude_session_groups: List[String] = Nil, exclude_sessions: List[String] = Nil, session_groups: List[String] = Nil, sessions: List[String] = Nil) { def ++ (other: Selection): Selection = Selection( requirements = requirements || other.requirements, all_sessions = all_sessions || other.all_sessions, base_sessions = Library.merge(base_sessions, other.base_sessions), exclude_session_groups = Library.merge(exclude_session_groups, other.exclude_session_groups), exclude_sessions = Library.merge(exclude_sessions, other.exclude_sessions), session_groups = Library.merge(session_groups, other.session_groups), sessions = Library.merge(sessions, other.sessions)) } object Structure { val empty: Structure = make(Nil) def make(infos: List[Info]): Structure = { def add_edges(graph: Graph[String, Info], kind: String, edges: Info => Iterable[String]) : Graph[String, Info] = { def add_edge(pos: Position.T, name: String, g: Graph[String, Info], parent: String) = { if (!g.defined(parent)) error("Bad " + kind + " session " + quote(parent) + " for " + quote(name) + Position.here(pos)) try { g.add_edge_acyclic(parent, name) } catch { case exn: Graph.Cycles[_] => error(cat_lines(exn.cycles.map(cycle => "Cyclic session dependency of " + cycle.map(c => quote(c.toString)).mkString(" via "))) + Position.here(pos)) } } graph.iterator.foldLeft(graph) { case (g, (name, (info, _))) => edges(info).foldLeft(g)(add_edge(info.pos, name, _, _)) } } val info_graph = infos.foldLeft(Graph.string[Info]) { case (graph, info) => if (graph.defined(info.name)) error("Duplicate session " + quote(info.name) + Position.here(info.pos) + Position.here(graph.get_node(info.name).pos)) else graph.new_node(info.name, info) } val build_graph = add_edges(info_graph, "parent", _.parent) val imports_graph = add_edges(build_graph, "imports", _.imports) val session_positions: List[(String, Position.T)] = (for ((name, (info, _)) <- info_graph.iterator) yield (name, info.pos)).toList val session_directories: Map[JFile, String] = (for { session <- imports_graph.topological_order.iterator info = info_graph.get_node(session) dir <- info.dirs.iterator } yield (info, dir)).foldLeft(Map.empty[JFile, String]) { case (dirs, (info, dir)) => val session = info.name val canonical_dir = dir.canonical_file dirs.get(canonical_dir) match { case Some(session1) => val info1 = info_graph.get_node(session1) error("Duplicate use of directory " + dir + "\n for session " + quote(session1) + Position.here(info1.pos) + "\n vs. session " + quote(session) + Position.here(info.pos)) case None => dirs + (canonical_dir -> session) } } val global_theories: Map[String, String] = (for { session <- imports_graph.topological_order.iterator info = info_graph.get_node(session) thy <- info.global_theories.iterator } yield (info, thy)).foldLeft(Thy_Header.bootstrap_global_theories.toMap) { case (global, (info, thy)) => val qualifier = info.name global.get(thy) match { case Some(qualifier1) if qualifier != qualifier1 => error("Duplicate global theory " + quote(thy) + Position.here(info.pos)) case _ => global + (thy -> qualifier) } } new Structure( session_positions, session_directories, global_theories, build_graph, imports_graph) } } final class Structure private[Sessions]( val session_positions: List[(String, Position.T)], val session_directories: Map[JFile, String], val global_theories: Map[String, String], val build_graph: Graph[String, Info], val imports_graph: Graph[String, Info]) { sessions_structure => def bootstrap: Base = Base( session_directories = session_directories, global_theories = global_theories, overall_syntax = Thy_Header.bootstrap_syntax) def dest_session_directories: List[(String, String)] = for ((file, session) <- session_directories.toList) yield (File.standard_path(file), session) lazy val chapters: SortedMap[String, List[Info]] = build_graph.iterator.foldLeft(SortedMap.empty[String, List[Info]]) { case (chs, (_, (info, _))) => chs + (info.chapter -> (info :: chs.getOrElse(info.chapter, Nil))) } def build_graph_display: Graph_Display.Graph = Graph_Display.make_graph(build_graph) def imports_graph_display: Graph_Display.Graph = Graph_Display.make_graph(imports_graph) def defined(name: String): Boolean = imports_graph.defined(name) def apply(name: String): Info = imports_graph.get_node(name) def get(name: String): Option[Info] = if (defined(name)) Some(apply(name)) else None def theory_qualifier(name: String): String = global_theories.getOrElse(name, Long_Name.qualifier(name)) def check_sessions(names: List[String]): Unit = { val bad_sessions = SortedSet(names.filterNot(defined): _*).toList if (bad_sessions.nonEmpty) error("Undefined session(s): " + commas_quote(bad_sessions)) } def check_sessions(sel: Selection): Unit = check_sessions(sel.base_sessions ::: sel.exclude_sessions ::: sel.sessions) private def selected(graph: Graph[String, Info], sel: Selection): List[String] = { check_sessions(sel) val select_group = sel.session_groups.toSet val select_session = sel.sessions.toSet ++ imports_graph.all_succs(sel.base_sessions) val selected0 = if (sel.all_sessions) graph.keys else { (for { (name, (info, _)) <- graph.iterator if info.dir_selected || select_session(name) || graph.get_node(name).groups.exists(select_group) } yield name).toList } if (sel.requirements) (graph.all_preds(selected0).toSet -- selected0).toList else selected0 } def selection(sel: Selection): Structure = { check_sessions(sel) val excluded = { val exclude_group = sel.exclude_session_groups.toSet val exclude_group_sessions = (for { (name, (info, _)) <- imports_graph.iterator if imports_graph.get_node(name).groups.exists(exclude_group) } yield name).toList imports_graph.all_succs(exclude_group_sessions ::: sel.exclude_sessions).toSet } def restrict(graph: Graph[String, Info]): Graph[String, Info] = { val sessions = graph.all_preds(selected(graph, sel)).filterNot(excluded) graph.restrict(graph.all_preds(sessions).toSet) } new Structure( session_positions, session_directories, global_theories, restrict(build_graph), restrict(imports_graph)) } def selection(session: String): Structure = selection(Selection.session(session)) def selection_deps( selection: Selection, progress: Progress = new Progress, loading_sessions: Boolean = false, inlined_files: Boolean = false, verbose: Boolean = false): Deps = { val deps = Sessions.deps(sessions_structure.selection(selection), progress = progress, inlined_files = inlined_files, verbose = verbose) if (loading_sessions) { val selection_size = deps.sessions_structure.build_graph.size if (selection_size > 1) progress.echo("Loading " + selection_size + " sessions ...") } deps } def hierarchy(session: String): List[String] = build_graph.all_preds(List(session)) def build_selection(sel: Selection): List[String] = selected(build_graph, sel) def build_descendants(ss: List[String]): List[String] = build_graph.all_succs(ss) def build_requirements(ss: List[String]): List[String] = build_graph.all_preds_rev(ss) def build_topological_order: List[String] = build_graph.topological_order def imports_selection(sel: Selection): List[String] = selected(imports_graph, sel) def imports_descendants(ss: List[String]): List[String] = imports_graph.all_succs(ss) def imports_requirements(ss: List[String]): List[String] = imports_graph.all_preds_rev(ss) def imports_topological_order: List[String] = imports_graph.topological_order def bibtex_entries: List[(String, List[String])] = build_topological_order.flatMap(name => apply(name).bibtex_entries match { case Nil => None case entries => Some(name -> entries.map(_.info)) }) def session_chapters: List[(String, String)] = imports_topological_order.map(name => name -> apply(name).chapter) override def toString: String = imports_graph.keys_iterator.mkString("Sessions.Structure(", ", ", ")") } /* parser */ private val CHAPTER = "chapter" private val SESSION = "session" private val IN = "in" private val DESCRIPTION = "description" private val DIRECTORIES = "directories" private val OPTIONS = "options" private val SESSIONS = "sessions" private val THEORIES = "theories" private val GLOBAL = "global" private val DOCUMENT_THEORIES = "document_theories" private val DOCUMENT_FILES = "document_files" private val EXPORT_FILES = "export_files" val root_syntax: Outer_Syntax = Outer_Syntax.empty + "(" + ")" + "+" + "," + "=" + "[" + "]" + GLOBAL + IN + (CHAPTER, Keyword.THY_DECL) + (SESSION, Keyword.THY_DECL) + (DESCRIPTION, Keyword.QUASI_COMMAND) + (DIRECTORIES, Keyword.QUASI_COMMAND) + (OPTIONS, Keyword.QUASI_COMMAND) + (SESSIONS, Keyword.QUASI_COMMAND) + (THEORIES, Keyword.QUASI_COMMAND) + (DOCUMENT_THEORIES, Keyword.QUASI_COMMAND) + (DOCUMENT_FILES, Keyword.QUASI_COMMAND) + (EXPORT_FILES, Keyword.QUASI_COMMAND) abstract class Entry sealed case class Chapter(name: String) extends Entry sealed case class Session_Entry( pos: Position.T, name: String, groups: List[String], path: String, parent: Option[String], description: String, options: List[Options.Spec], imports: List[String], directories: List[String], theories: List[(List[Options.Spec], List[((String, Position.T), Boolean)])], document_theories: List[(String, Position.T)], document_files: List[(String, String)], export_files: List[(String, Int, List[String])]) extends Entry { def theories_no_position: List[(List[Options.Spec], List[(String, Boolean)])] = theories.map({ case (a, b) => (a, b.map({ case ((c, _), d) => (c, d) })) }) def document_theories_no_position: List[String] = document_theories.map(_._1) } private object Parser extends Options.Parser { private val chapter: Parser[Chapter] = { val chapter_name = atom("chapter name", _.is_name) command(CHAPTER) ~! chapter_name ^^ { case _ ~ a => Chapter(a) } } private val session_entry: Parser[Session_Entry] = { val option = option_name ~ opt($$$("=") ~! option_value ^^ { case _ ~ x => x }) ^^ { case x ~ y => (x, y) } val options = $$$("[") ~> rep1sep(option, $$$(",")) <~ $$$("]") val theory_entry = position(theory_name) ~ opt_keyword(GLOBAL) ^^ { case x ~ y => (x, y) } val theories = $$$(THEORIES) ~! ((options | success(Nil)) ~ rep1(theory_entry)) ^^ { case _ ~ (x ~ y) => (x, y) } val in_path = $$$("(") ~! ($$$(IN) ~ path ~ $$$(")")) ^^ { case _ ~ (_ ~ x ~ _) => x } val document_theories = $$$(DOCUMENT_THEORIES) ~! rep1(position(name)) ^^ { case _ ~ x => x } val document_files = $$$(DOCUMENT_FILES) ~! ((in_path | success("document")) ~ rep1(path)) ^^ { case _ ~ (x ~ y) => y.map((x, _)) } val prune = $$$("[") ~! (nat ~ $$$("]")) ^^ { case _ ~ (x ~ _) => x } | success(0) val export_files = $$$(EXPORT_FILES) ~! ((in_path | success("export")) ~ prune ~ rep1(embedded)) ^^ { case _ ~ (x ~ y ~ z) => (x, y, z) } command(SESSION) ~! (position(session_name) ~ (($$$("(") ~! (rep1(name) <~ $$$(")")) ^^ { case _ ~ x => x }) | success(Nil)) ~ (($$$(IN) ~! path ^^ { case _ ~ x => x }) | success(".")) ~ ($$$("=") ~! (opt(session_name ~! $$$("+") ^^ { case x ~ _ => x }) ~ (($$$(DESCRIPTION) ~! text ^^ { case _ ~ x => x }) | success("")) ~ (($$$(OPTIONS) ~! options ^^ { case _ ~ x => x }) | success(Nil)) ~ (($$$(SESSIONS) ~! rep1(session_name) ^^ { case _ ~ x => x }) | success(Nil)) ~ (($$$(DIRECTORIES) ~! rep1(path) ^^ { case _ ~ x => x }) | success(Nil)) ~ rep(theories) ~ (opt(document_theories) ^^ (x => x.getOrElse(Nil))) ~ (rep(document_files) ^^ (x => x.flatten)) ~ rep(export_files)))) ^^ { case _ ~ ((a, pos) ~ b ~ c ~ (_ ~ (d ~ e ~ f ~ g ~ h ~ i ~ j ~ k ~ l))) => Session_Entry(pos, a, b, c, d, e, f, g, h, i, j, k, l) } } def parse_root(path: Path): List[Entry] = { val toks = Token.explode(root_syntax.keywords, File.read(path)) val start = Token.Pos.file(path.implode) parse_all(rep(chapter | session_entry), Token.reader(toks, start)) match { case Success(result, _) => result case bad => error(bad.toString) } } } def parse_root(path: Path): List[Entry] = Parser.parse_root(path) def parse_root_entries(path: Path): List[Session_Entry] = for (entry <- Parser.parse_root(path) if entry.isInstanceOf[Session_Entry]) yield entry.asInstanceOf[Session_Entry] def read_root(options: Options, select: Boolean, path: Path): List[Info] = { var entry_chapter = UNSORTED val infos = new mutable.ListBuffer[Info] parse_root(path).foreach { case Chapter(name) => entry_chapter = name case entry: Session_Entry => infos += make_info(options, select, path.dir, entry_chapter, entry) } infos.toList } def parse_roots(roots: Path): List[String] = { for { line <- split_lines(File.read(roots)) if !(line == "" || line.startsWith("#")) } yield line } /* load sessions from certain directories */ private def is_session_dir(dir: Path): Boolean = (dir + ROOT).is_file || (dir + ROOTS).is_file private def check_session_dir(dir: Path): Path = if (is_session_dir(dir)) File.pwd() + dir.expand else error("Bad session root directory (missing ROOT or ROOTS): " + dir.expand.toString) def directories(dirs: List[Path], select_dirs: List[Path]): List[(Boolean, Path)] = { val default_dirs = Isabelle_System.components().filter(is_session_dir) for { (select, dir) <- (default_dirs ::: dirs).map((false, _)) ::: select_dirs.map((true, _)) } yield (select, dir.canonical) } def load_structure(options: Options, dirs: List[Path] = Nil, select_dirs: List[Path] = Nil, infos: List[Info] = Nil): Structure = { def load_dir(select: Boolean, dir: Path): List[(Boolean, Path)] = load_root(select, dir) ::: load_roots(select, dir) def load_root(select: Boolean, dir: Path): List[(Boolean, Path)] = { val root = dir + ROOT if (root.is_file) List((select, root)) else Nil } def load_roots(select: Boolean, dir: Path): List[(Boolean, Path)] = { val roots = dir + ROOTS if (roots.is_file) { for { entry <- parse_roots(roots) dir1 = try { check_session_dir(dir + Path.explode(entry)) } catch { case ERROR(msg) => error(msg + "\nThe error(s) above occurred in session catalog " + roots.toString) } res <- load_dir(select, dir1) } yield res } else Nil } val roots = for { (select, dir) <- directories(dirs, select_dirs) res <- load_dir(select, check_session_dir(dir)) } yield res val unique_roots = roots.foldLeft(Map.empty[JFile, (Boolean, Path)]) { case (m, (select, path)) => val file = path.canonical_file m.get(file) match { case None => m + (file -> (select, path)) case Some((select1, path1)) => m + (file -> (select1 || select, path1)) } }.toList.map(_._2) Structure.make(unique_roots.flatMap(p => read_root(options, p._1, p._2)) ::: infos) } /* Isabelle tool wrapper */ val isabelle_tool = Isabelle_Tool("sessions", "explore structure of Isabelle sessions", Scala_Project.here, args => { var base_sessions: List[String] = Nil var select_dirs: List[Path] = Nil var requirements = false var exclude_session_groups: List[String] = Nil var all_sessions = false var dirs: List[Path] = Nil var session_groups: List[String] = Nil var exclude_sessions: List[String] = Nil val getopts = Getopts(""" Usage: isabelle sessions [OPTIONS] [SESSIONS ...] Options are: -B NAME include session NAME and all descendants -D DIR include session directory and select its sessions -R refer to requirements of selected sessions -X NAME exclude sessions from group NAME and all descendants -a select all sessions -d DIR include session directory -g NAME select session group NAME -x NAME exclude session NAME and all descendants Explore the structure of Isabelle sessions and print result names in topological order (on stdout). """, "B:" -> (arg => base_sessions = base_sessions ::: List(arg)), "D:" -> (arg => select_dirs = select_dirs ::: List(Path.explode(arg))), "R" -> (_ => requirements = true), "X:" -> (arg => exclude_session_groups = exclude_session_groups ::: List(arg)), "a" -> (_ => all_sessions = true), "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))), "g:" -> (arg => session_groups = session_groups ::: List(arg)), "x:" -> (arg => exclude_sessions = exclude_sessions ::: List(arg))) val sessions = getopts(args) val options = Options.init() val selection = Selection(requirements = requirements, all_sessions = all_sessions, base_sessions = base_sessions, exclude_session_groups = exclude_session_groups, exclude_sessions = exclude_sessions, session_groups = session_groups, sessions = sessions) val sessions_structure = load_structure(options, dirs = dirs, select_dirs = select_dirs).selection(selection) for (name <- sessions_structure.imports_topological_order) { Output.writeln(name, stdout = true) } }) /** heap file with SHA1 digest **/ private val sha1_prefix = "SHA1:" def read_heap_digest(heap: Path): Option[String] = { if (heap.is_file) { using(FileChannel.open(heap.file.toPath, StandardOpenOption.READ))(file => { val len = file.size val n = sha1_prefix.length + SHA1.digest_length if (len >= n) { file.position(len - n) val buf = ByteBuffer.allocate(n) var i = 0 var m = 0 do { m = file.read(buf) if (m != -1) i += m } while (m != -1 && n > i) if (i == n) { val prefix = new String(buf.array(), 0, sha1_prefix.length, UTF8.charset) val s = new String(buf.array(), sha1_prefix.length, SHA1.digest_length, UTF8.charset) if (prefix == sha1_prefix) Some(s) else None } else None } else None }) } else None } def write_heap_digest(heap: Path): String = read_heap_digest(heap) match { case None => val s = SHA1.digest(heap).rep File.append(heap, sha1_prefix + s) s case Some(s) => s } /** persistent store **/ object Session_Info { val session_name = SQL.Column.string("session_name").make_primary_key // Build_Log.Session_Info val session_timing = SQL.Column.bytes("session_timing") val command_timings = SQL.Column.bytes("command_timings") val theory_timings = SQL.Column.bytes("theory_timings") val ml_statistics = SQL.Column.bytes("ml_statistics") val task_statistics = SQL.Column.bytes("task_statistics") val errors = SQL.Column.bytes("errors") val build_log_columns = List(session_name, session_timing, command_timings, theory_timings, ml_statistics, task_statistics, errors) // Build.Session_Info val sources = SQL.Column.string("sources") val input_heaps = SQL.Column.string("input_heaps") val output_heap = SQL.Column.string("output_heap") val return_code = SQL.Column.int("return_code") val build_columns = List(sources, input_heaps, output_heap, return_code) val table = SQL.Table("isabelle_session_info", build_log_columns ::: build_columns) } class Database_Context private[Sessions]( val store: Sessions.Store, database_server: Option[SQL.Database]) extends AutoCloseable { def cache: XML.Cache = store.cache - def close: Unit = database_server.foreach(_.close) + def close(): Unit = database_server.foreach(_.close()) def output_database[A](session: String)(f: SQL.Database => A): A = database_server match { case Some(db) => f(db) case None => using(store.open_database(session, output = true))(f) } def input_database[A](session: String)(f: (SQL.Database, String) => Option[A]): Option[A] = database_server match { case Some(db) => f(db, session) case None => store.try_open_database(session) match { case Some(db) => using(db)(f(_, session)) case None => None } } def read_export( sessions: List[String], theory_name: String, name: String): Option[Export.Entry] = { val attempts = database_server match { case Some(db) => sessions.view.map(session_name => Export.read_entry(db, store.cache, session_name, theory_name, name)) case None => sessions.view.map(session_name => store.try_open_database(session_name) match { case Some(db) => using(db)(Export.read_entry(_, store.cache, session_name, theory_name, name)) case None => None }) } attempts.collectFirst({ case Some(entry) => entry }) } def get_export( session_hierarchy: List[String], theory_name: String, name: String): Export.Entry = read_export(session_hierarchy, theory_name, name) getOrElse Export.empty_entry(theory_name, name) override def toString: String = { val s = database_server match { case Some(db) => db.toString case None => "input_dirs = " + store.input_dirs.map(_.absolute).mkString(", ") } "Database_Context(" + s + ")" } } def store(options: Options, cache: XML.Cache = XML.Cache.make()): Store = new Store(options, cache) class Store private[Sessions](val options: Options, val cache: XML.Cache) { store => override def toString: String = "Store(output_dir = " + output_dir.absolute + ")" /* directories */ val system_output_dir: Path = Path.explode("$ISABELLE_HEAPS_SYSTEM/$ML_IDENTIFIER") val user_output_dir: Path = Path.explode("$ISABELLE_HEAPS/$ML_IDENTIFIER") def system_heaps: Boolean = options.bool("system_heaps") val output_dir: Path = if (system_heaps) system_output_dir else user_output_dir val input_dirs: List[Path] = if (system_heaps) List(system_output_dir) else List(user_output_dir, system_output_dir) def presentation_dir: Path = if (system_heaps) Path.explode("$ISABELLE_BROWSER_INFO_SYSTEM") else Path.explode("$ISABELLE_BROWSER_INFO") /* file names */ def heap(name: String): Path = Path.basic(name) def database(name: String): Path = Path.basic("log") + Path.basic(name).ext("db") def log(name: String): Path = Path.basic("log") + Path.basic(name) def log_gz(name: String): Path = log(name).ext("gz") def output_heap(name: String): Path = output_dir + heap(name) def output_database(name: String): Path = output_dir + database(name) def output_log(name: String): Path = output_dir + log(name) def output_log_gz(name: String): Path = output_dir + log_gz(name) def prepare_output_dir(): Unit = Isabelle_System.make_directory(output_dir + Path.basic("log")) /* heap */ def find_heap(name: String): Option[Path] = input_dirs.map(_ + heap(name)).find(_.is_file) def find_heap_digest(name: String): Option[String] = find_heap(name).flatMap(read_heap_digest) def the_heap(name: String): Path = find_heap(name) getOrElse error("Missing heap image for session " + quote(name) + " -- expected in:\n" + cat_lines(input_dirs.map(dir => " " + dir.expand.implode))) /* database */ def database_server: Boolean = options.bool("build_database_server") def open_database_server(): SQL.Database = PostgreSQL.open_database( user = options.string("build_database_user"), password = options.string("build_database_password"), database = options.string("build_database_name"), host = options.string("build_database_host"), port = options.int("build_database_port"), ssh = options.proper_string("build_database_ssh_host").map(ssh_host => SSH.open_session(options, host = ssh_host, user = options.string("build_database_ssh_user"), port = options.int("build_database_ssh_port"))), ssh_close = true) def open_database_context(): Database_Context = new Database_Context(store, if (database_server) Some(open_database_server()) else None) def try_open_database(name: String, output: Boolean = false): Option[SQL.Database] = { def check(db: SQL.Database): Option[SQL.Database] = - if (output || session_info_exists(db)) Some(db) else { db.close; None } + if (output || session_info_exists(db)) Some(db) else { db.close(); None } if (database_server) check(open_database_server()) else if (output) Some(SQLite.open_database(output_database(name))) else { (for { dir <- input_dirs.view path = dir + database(name) if path.is_file db <- check(SQLite.open_database(path)) } yield db).headOption } } def open_database(name: String, output: Boolean = false): SQL.Database = try_open_database(name, output = output) getOrElse error("Missing build database for session " + quote(name)) def clean_output(name: String): (Boolean, Boolean) = { val relevant_db = database_server && { try_open_database(name) match { case Some(db) => try { db.transaction { val relevant_db = session_info_defined(db, name) init_session_info(db, name) relevant_db } - } finally { db.close } + } finally { db.close() } case None => false } } val del = for { dir <- (if (system_heaps) List(user_output_dir, system_output_dir) else List(user_output_dir)) file <- List(heap(name), database(name), log(name), log_gz(name)) path = dir + file if path.is_file } yield path.file.delete val relevant = relevant_db || del.nonEmpty val ok = del.forall(b => b) (relevant, ok) } /* SQL database content */ def read_bytes(db: SQL.Database, name: String, column: SQL.Column): Bytes = db.using_statement(Session_Info.table.select(List(column), Session_Info.session_name.where_equal(name)))(stmt => { val res = stmt.execute_query() if (!res.next()) Bytes.empty else res.bytes(column) }) def read_properties(db: SQL.Database, name: String, column: SQL.Column): List[Properties.T] = Properties.uncompress(read_bytes(db, name, column), cache = cache) /* session info */ def init_session_info(db: SQL.Database, name: String): Unit = { db.transaction { db.create_table(Session_Info.table) db.using_statement( - Session_Info.table.delete(Session_Info.session_name.where_equal(name)))(_.execute) + Session_Info.table.delete(Session_Info.session_name.where_equal(name)))(_.execute()) db.create_table(Export.Data.table) db.using_statement( - Export.Data.table.delete(Export.Data.session_name.where_equal(name)))(_.execute) + Export.Data.table.delete(Export.Data.session_name.where_equal(name)))(_.execute()) db.create_table(Presentation.Data.table) db.using_statement( Presentation.Data.table.delete( - Presentation.Data.session_name.where_equal(name)))(_.execute) + Presentation.Data.session_name.where_equal(name)))(_.execute()) } } def session_info_exists(db: SQL.Database): Boolean = { val tables = db.tables tables.contains(Session_Info.table.name) && tables.contains(Export.Data.table.name) } def session_info_defined(db: SQL.Database, name: String): Boolean = db.transaction { session_info_exists(db) && { db.using_statement( Session_Info.table.select(List(Session_Info.session_name), Session_Info.session_name.where_equal(name)))(stmt => stmt.execute_query().next()) } } def write_session_info( db: SQL.Database, name: String, build_log: Build_Log.Session_Info, build: Build.Session_Info): Unit = { db.transaction { db.using_statement(Session_Info.table.insert())(stmt => { stmt.string(1) = name stmt.bytes(2) = Properties.encode(build_log.session_timing) stmt.bytes(3) = Properties.compress(build_log.command_timings, cache = cache.xz) stmt.bytes(4) = Properties.compress(build_log.theory_timings, cache = cache.xz) stmt.bytes(5) = Properties.compress(build_log.ml_statistics, cache = cache.xz) stmt.bytes(6) = Properties.compress(build_log.task_statistics, cache = cache.xz) stmt.bytes(7) = Build_Log.compress_errors(build_log.errors, cache = cache.xz) stmt.string(8) = build.sources stmt.string(9) = cat_lines(build.input_heaps) stmt.string(10) = build.output_heap getOrElse "" stmt.int(11) = build.return_code stmt.execute() }) } } def read_session_timing(db: SQL.Database, name: String): Properties.T = Properties.decode(read_bytes(db, name, Session_Info.session_timing), cache = cache) def read_command_timings(db: SQL.Database, name: String): List[Properties.T] = read_properties(db, name, Session_Info.command_timings) def read_theory_timings(db: SQL.Database, name: String): List[Properties.T] = read_properties(db, name, Session_Info.theory_timings) def read_ml_statistics(db: SQL.Database, name: String): List[Properties.T] = read_properties(db, name, Session_Info.ml_statistics) def read_task_statistics(db: SQL.Database, name: String): List[Properties.T] = read_properties(db, name, Session_Info.task_statistics) def read_theories(db: SQL.Database, name: String): List[String] = read_theory_timings(db, name).flatMap(Markup.Name.unapply) def read_errors(db: SQL.Database, name: String): List[String] = Build_Log.uncompress_errors(read_bytes(db, name, Session_Info.errors), cache = cache) def read_build(db: SQL.Database, name: String): Option[Build.Session_Info] = { if (db.tables.contains(Session_Info.table.name)) { db.using_statement(Session_Info.table.select(Session_Info.build_columns, Session_Info.session_name.where_equal(name)))(stmt => { val res = stmt.execute_query() if (!res.next()) None else { Some( Build.Session_Info( res.string(Session_Info.sources), split_lines(res.string(Session_Info.input_heaps)), res.string(Session_Info.output_heap) match { case "" => None case s => Some(s) }, res.int(Session_Info.return_code))) } }) } else None } } } diff --git a/src/Pure/Tools/build.scala b/src/Pure/Tools/build.scala --- a/src/Pure/Tools/build.scala +++ b/src/Pure/Tools/build.scala @@ -1,679 +1,679 @@ /* Title: Pure/Tools/build.scala Author: Makarius Options: :folding=explicit: Build and manage Isabelle sessions. */ package isabelle import scala.collection.immutable.SortedSet import scala.annotation.tailrec object Build { /** auxiliary **/ /* persistent build info */ sealed case class Session_Info( sources: String, input_heaps: List[String], output_heap: Option[String], return_code: Int) { def ok: Boolean = return_code == 0 } /* queue with scheduling information */ private object Queue { type Timings = (List[Properties.T], Double) def load_timings(progress: Progress, store: Sessions.Store, session_name: String): Timings = { val no_timings: Timings = (Nil, 0.0) store.try_open_database(session_name) match { case None => no_timings case Some(db) => def ignore_error(msg: String) = { progress.echo_warning("Ignoring bad database " + db + (if (msg == "") "" else "\n" + msg)) no_timings } try { val command_timings = store.read_command_timings(db, session_name) val session_timing = store.read_session_timing(db, session_name) match { case Markup.Elapsed(t) => t case _ => 0.0 } (command_timings, session_timing) } catch { case ERROR(msg) => ignore_error(msg) case exn: java.lang.Error => ignore_error(Exn.message(exn)) case _: XML.Error => ignore_error("") } - finally { db.close } + finally { db.close() } } } def make_session_timing(sessions_structure: Sessions.Structure, timing: Map[String, Double]) : Map[String, Double] = { val maximals = sessions_structure.build_graph.maximals.toSet def desc_timing(session_name: String): Double = { if (maximals.contains(session_name)) timing(session_name) else { val descendants = sessions_structure.build_descendants(List(session_name)).toSet val g = sessions_structure.build_graph.restrict(descendants) (0.0 :: g.maximals.flatMap(desc => { val ps = g.all_preds(List(desc)) if (ps.exists(p => !timing.isDefinedAt(p))) None else Some(ps.map(timing(_)).sum) })).max } } timing.keySet.iterator.map(name => (name -> desc_timing(name))).toMap.withDefaultValue(0.0) } def apply(progress: Progress, sessions_structure: Sessions.Structure, store: Sessions.Store) : Queue = { val graph = sessions_structure.build_graph val names = graph.keys val timings = names.map(name => (name, load_timings(progress, store, name))) val command_timings = timings.map({ case (name, (ts, _)) => (name, ts) }).toMap.withDefaultValue(Nil) val session_timing = make_session_timing(sessions_structure, timings.map({ case (name, (_, t)) => (name, t) }).toMap) object Ordering extends scala.math.Ordering[String] { def compare(name1: String, name2: String): Int = session_timing(name2) compare session_timing(name1) match { case 0 => sessions_structure(name2).timeout compare sessions_structure(name1).timeout match { case 0 => name1 compare name2 case ord => ord } case ord => ord } } new Queue(graph, SortedSet(names: _*)(Ordering), command_timings) } } private class Queue( graph: Graph[String, Sessions.Info], order: SortedSet[String], val command_timings: String => List[Properties.T]) { def is_inner(name: String): Boolean = !graph.is_maximal(name) def is_empty: Boolean = graph.is_empty def - (name: String): Queue = new Queue(graph.del_node(name), order - name, command_timings) def dequeue(skip: String => Boolean): Option[(String, Sessions.Info)] = { val it = order.iterator.dropWhile(name => skip(name) || !graph.is_minimal(name)) if (it.hasNext) { val name = it.next(); Some((name, graph.get_node(name))) } else None } } /** build with results **/ class Results private[Build](results: Map[String, (Option[Process_Result], Sessions.Info)]) { def sessions: Set[String] = results.keySet def infos: List[Sessions.Info] = results.values.map(_._2).toList def cancelled(name: String): Boolean = results(name)._1.isEmpty def apply(name: String): Process_Result = results(name)._1.getOrElse(Process_Result(1)) def info(name: String): Sessions.Info = results(name)._2 val rc: Int = results.iterator.map({ case (_, (Some(r), _)) => r.rc case (_, (None, _)) => 1 }). foldLeft(0)(_ max _) def ok: Boolean = rc == 0 override def toString: String = rc.toString } def session_finished(session_name: String, process_result: Process_Result): String = "Finished " + session_name + " (" + process_result.timing.message_resources + ")" def session_timing(session_name: String, build_log: Build_Log.Session_Info): String = { val props = build_log.session_timing val threads = Markup.Session_Timing.Threads.unapply(props) getOrElse 1 val timing = Markup.Timing_Properties.parse(props) "Timing " + session_name + " (" + threads + " threads, " + timing.message_factor + ")" } def build( options: Options, selection: Sessions.Selection = Sessions.Selection.empty, presentation: Presentation.Context = Presentation.Context.none, progress: Progress = new Progress, check_unknown_files: Boolean = false, build_heap: Boolean = false, clean_build: Boolean = false, dirs: List[Path] = Nil, select_dirs: List[Path] = Nil, infos: List[Sessions.Info] = Nil, numa_shuffling: Boolean = false, max_jobs: Int = 1, list_files: Boolean = false, check_keywords: Set[String] = Set.empty, fresh_build: Boolean = false, no_build: Boolean = false, soft_build: Boolean = false, verbose: Boolean = false, export_files: Boolean = false): Results = { val build_options = options + "completion_limit=0" + "editor_tracing_messages=0" + "kodkod_scala=false" + ("pide_reports=" + options.bool("build_pide_reports")) val store = Sessions.store(build_options) Isabelle_Fonts.init() /* session selection and dependencies */ val full_sessions = Sessions.load_structure(build_options, dirs = dirs, select_dirs = select_dirs, infos = infos) val full_sessions_selection = full_sessions.imports_selection(selection) def sources_stamp(deps: Sessions.Deps, session_name: String): String = { val digests = full_sessions(session_name).meta_digest :: deps.sources(session_name) ::: deps.imported_sources(session_name) SHA1.digest_set(digests).toString } val deps = { val deps0 = Sessions.deps(full_sessions.selection(selection), progress = progress, inlined_files = true, verbose = verbose, list_files = list_files, check_keywords = check_keywords).check_errors if (soft_build && !fresh_build) { val outdated = deps0.sessions_structure.build_topological_order.flatMap(name => store.try_open_database(name) match { case Some(db) => using(db)(store.read_build(_, name)) match { case Some(build) if build.ok && build.sources == sources_stamp(deps0, name) => None case _ => Some(name) } case None => Some(name) }) Sessions.deps(full_sessions.selection(Sessions.Selection(sessions = outdated)), progress = progress, inlined_files = true).check_errors } else deps0 } /* check unknown files */ if (check_unknown_files) { val source_files = (for { (_, base) <- deps.session_bases.iterator (path, _) <- base.sources.iterator } yield path).toList val exclude_files = List(Path.explode("$POLYML_EXE")).map(_.canonical_file) val unknown_files = Mercurial.check_files(source_files)._2. filterNot(path => exclude_files.contains(path.canonical_file)) if (unknown_files.nonEmpty) { progress.echo_warning("Unknown files (not part of the underlying Mercurial repository):" + unknown_files.map(path => path.expand.implode).sorted.mkString("\n ", "\n ", "")) } } /* main build process */ val queue = Queue(progress, deps.sessions_structure, store) store.prepare_output_dir() if (clean_build) { for (name <- full_sessions.imports_descendants(full_sessions_selection)) { val (relevant, ok) = store.clean_output(name) if (relevant) { if (ok) progress.echo("Cleaned " + name) else progress.echo(name + " FAILED to clean") } } } // scheduler loop case class Result( current: Boolean, heap_digest: Option[String], process: Option[Process_Result], info: Sessions.Info) { def ok: Boolean = process match { case None => false case Some(res) => res.rc == 0 } } def sleep(): Unit = - Isabelle_Thread.interrupt_handler(_ => progress.stop) { Time.seconds(0.5).sleep } + Isabelle_Thread.interrupt_handler(_ => progress.stop()) { Time.seconds(0.5).sleep } val numa_nodes = new NUMA.Nodes(numa_shuffling) @tailrec def loop( pending: Queue, running: Map[String, (List[String], Build_Job)], results: Map[String, Result]): Map[String, Result] = { def used_node(i: Int): Boolean = running.iterator.exists( { case (_, (_, job)) => job.numa_node.isDefined && job.numa_node.get == i }) if (pending.is_empty) results else { if (progress.stopped) { - for ((_, (_, job)) <- running) job.terminate + for ((_, (_, job)) <- running) job.terminate() } running.find({ case (_, (_, job)) => job.is_finished }) match { case Some((session_name, (input_heaps, job))) => //{{{ finish job val (process_result, heap_digest) = job.join val log_lines = process_result.out_lines.filterNot(Protocol_Message.Marker.test) val process_result_tail = { val tail = job.info.options.int("process_output_tail") process_result.copy( out_lines = "(see also " + store.output_log(session_name).file.toString + ")" :: (if (tail == 0) log_lines else log_lines.drop(log_lines.length - tail max 0))) } val build_log = Build_Log.Log_File(session_name, process_result.out_lines). parse_session_info( command_timings = true, theory_timings = true, ml_statistics = true, task_statistics = true) // write log file if (process_result.ok) { File.write_gzip(store.output_log_gz(session_name), terminate_lines(log_lines)) } else File.write(store.output_log(session_name), terminate_lines(log_lines)) // write database using(store.open_database(session_name, output = true))(db => store.write_session_info(db, session_name, build_log = if (process_result.timeout) build_log.error("Timeout") else build_log, build = Session_Info(sources_stamp(deps, session_name), input_heaps, heap_digest, process_result.rc))) // messages process_result.err_lines.foreach(progress.echo) if (process_result.ok) { if (verbose) progress.echo(session_timing(session_name, build_log)) progress.echo(session_finished(session_name, process_result)) } else { progress.echo(session_name + " FAILED") if (!process_result.interrupted) progress.echo(process_result_tail.out) } loop(pending - session_name, running - session_name, results + (session_name -> Result(false, heap_digest, Some(process_result_tail), job.info))) //}}} case None if running.size < (max_jobs max 1) => //{{{ check/start next job pending.dequeue(running.isDefinedAt) match { case Some((session_name, info)) => val ancestor_results = deps.sessions_structure.build_requirements(List(session_name)). filterNot(_ == session_name).map(results(_)) val ancestor_heaps = ancestor_results.flatMap(_.heap_digest) val do_store = build_heap || Sessions.is_pure(session_name) || queue.is_inner(session_name) val (current, heap_digest) = { store.try_open_database(session_name) match { case Some(db) => using(db)(store.read_build(_, session_name)) match { case Some(build) => val heap_digest = store.find_heap_digest(session_name) val current = !fresh_build && build.ok && build.sources == sources_stamp(deps, session_name) && build.input_heaps == ancestor_heaps && build.output_heap == heap_digest && !(do_store && heap_digest.isEmpty) (current, heap_digest) case None => (false, None) } case None => (false, None) } } val all_current = current && ancestor_results.forall(_.current) if (all_current) loop(pending - session_name, running, results + (session_name -> Result(true, heap_digest, Some(Process_Result(0)), info))) else if (no_build) { progress.echo_if(verbose, "Skipping " + session_name + " ...") loop(pending - session_name, running, results + (session_name -> Result(false, heap_digest, Some(Process_Result(1)), info))) } else if (ancestor_results.forall(_.ok) && !progress.stopped) { progress.echo((if (do_store) "Building " else "Running ") + session_name + " ...") store.clean_output(session_name) using(store.open_database(session_name, output = true))( store.init_session_info(_, session_name)) val numa_node = numa_nodes.next(used_node) val job = new Build_Job(progress, session_name, info, deps, store, do_store, verbose, numa_node, queue.command_timings(session_name)) loop(pending, running + (session_name -> (ancestor_heaps, job)), results) } else { progress.echo(session_name + " CANCELLED") loop(pending - session_name, running, results + (session_name -> Result(false, heap_digest, None, info))) } case None => sleep(); loop(pending, running, results) } ///}}} case None => sleep(); loop(pending, running, results) } } } /* build results */ val results = { val results0 = if (deps.is_empty) { progress.echo_warning("Nothing to build") Map.empty[String, Result] } else Isabelle_Thread.uninterruptible { loop(queue, Map.empty, Map.empty) } new Results( (for ((name, result) <- results0.iterator) yield (name, (result.process, result.info))).toMap) } if (export_files) { for (name <- full_sessions.imports_selection(selection).iterator if results(name).ok) { val info = results.info(name) if (info.export_files.nonEmpty) { progress.echo("Exporting " + info.name + " ...") for ((dir, prune, pats) <- info.export_files) { Export.export_files(store, name, info.dir + dir, progress = if (verbose) progress else new Progress, export_prune = prune, export_patterns = pats) } } } } if (results.rc != 0 && (verbose || !no_build)) { val unfinished = (for { name <- results.sessions.iterator if !results(name).ok } yield name).toList.sorted progress.echo("Unfinished session(s): " + commas(unfinished)) } /* PDF/HTML presentation */ if (!no_build && !progress.stopped && results.ok) { val selected = full_sessions_selection.toSet val presentation_chapters = (for { session_name <- deps.sessions_structure.build_topological_order.iterator info = results.info(session_name) if selected(session_name) && presentation.enabled(info) && results(session_name).ok } yield (info.chapter, (session_name, info.description))).toList if (presentation_chapters.nonEmpty) { val presentation_dir = presentation.dir(store) progress.echo("Presentation in " + presentation_dir.absolute) val resources = Resources.empty val html_context = Presentation.html_context() using(store.open_database_context())(db_context => for ((_, (session_name, _)) <- presentation_chapters) { progress.expose_interrupt() progress.echo("Presenting " + session_name + " ...") Presentation.session_html( resources, session_name, deps, db_context, progress = progress, verbose = verbose, html_context = html_context, elements = Presentation.elements1, presentation = presentation) }) val browser_chapters = presentation_chapters.groupBy(_._1). map({ case (chapter, es) => (chapter, es.map(_._2)) }).filterNot(_._2.isEmpty) for ((chapter, entries) <- browser_chapters) Presentation.update_chapter_index(presentation_dir, chapter, entries) if (browser_chapters.nonEmpty) Presentation.make_global_index(presentation_dir) } } results } /* Isabelle tool wrapper */ val isabelle_tool = Isabelle_Tool("build", "build and manage Isabelle sessions", Scala_Project.here, args => { val build_options = Word.explode(Isabelle_System.getenv("ISABELLE_BUILD_OPTIONS")) var base_sessions: List[String] = Nil var select_dirs: List[Path] = Nil var numa_shuffling = false var presentation = Presentation.Context.none var requirements = false var soft_build = false var exclude_session_groups: List[String] = Nil var all_sessions = false var build_heap = false var clean_build = false var dirs: List[Path] = Nil var export_files = false var fresh_build = false var session_groups: List[String] = Nil var max_jobs = 1 var check_keywords: Set[String] = Set.empty var list_files = false var no_build = false var options = Options.init(opts = build_options) var verbose = false var exclude_sessions: List[String] = Nil val getopts = Getopts(""" Usage: isabelle build [OPTIONS] [SESSIONS ...] Options are: -B NAME include session NAME and all descendants -D DIR include session directory and select its sessions -N cyclic shuffling of NUMA CPU nodes (performance tuning) -P DIR enable HTML/PDF presentation in directory (":" for default) -R refer to requirements of selected sessions -S soft build: only observe changes of sources, not heap images -X NAME exclude sessions from group NAME and all descendants -a select all sessions -b build heap images -c clean build -d DIR include session directory -e export files from session specification into file-system -f fresh build -g NAME select session group NAME -j INT maximum number of parallel jobs (default 1) -k KEYWORD check theory sources for conflicts with proposed keywords -l list session source files -n no build -- test dependencies only -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -v verbose -x NAME exclude session NAME and all descendants Build and manage Isabelle sessions, depending on implicit settings: """ + Library.prefix_lines(" ", Build_Log.Settings.show()) + "\n", "B:" -> (arg => base_sessions = base_sessions ::: List(arg)), "D:" -> (arg => select_dirs = select_dirs ::: List(Path.explode(arg))), "N" -> (_ => numa_shuffling = true), "P:" -> (arg => presentation = Presentation.Context.make(arg)), "R" -> (_ => requirements = true), "S" -> (_ => soft_build = true), "X:" -> (arg => exclude_session_groups = exclude_session_groups ::: List(arg)), "a" -> (_ => all_sessions = true), "b" -> (_ => build_heap = true), "c" -> (_ => clean_build = true), "d:" -> (arg => dirs = dirs ::: List(Path.explode(arg))), "e" -> (_ => export_files = true), "f" -> (_ => fresh_build = true), "g:" -> (arg => session_groups = session_groups ::: List(arg)), "j:" -> (arg => max_jobs = Value.Int.parse(arg)), "k:" -> (arg => check_keywords = check_keywords + arg), "l" -> (_ => list_files = true), "n" -> (_ => no_build = true), "o:" -> (arg => options = options + arg), "v" -> (_ => verbose = true), "x:" -> (arg => exclude_sessions = exclude_sessions ::: List(arg))) val sessions = getopts(args) val progress = new Console_Progress(verbose = verbose) val start_date = Date.now() if (verbose) { progress.echo( "Started at " + Build_Log.print_date(start_date) + " (" + Isabelle_System.getenv("ML_IDENTIFIER") + " on " + Isabelle_System.hostname() +")") progress.echo(Build_Log.Settings.show() + "\n") } val results = progress.interrupt_handler { build(options, selection = Sessions.Selection( requirements = requirements, all_sessions = all_sessions, base_sessions = base_sessions, exclude_session_groups = exclude_session_groups, exclude_sessions = exclude_sessions, session_groups = session_groups, sessions = sessions), presentation = presentation, progress = progress, check_unknown_files = Mercurial.is_repository(Path.explode("~~")), build_heap = build_heap, clean_build = clean_build, dirs = dirs, select_dirs = select_dirs, numa_shuffling = NUMA.enabled_warning(progress, numa_shuffling), max_jobs = max_jobs, list_files = list_files, check_keywords = check_keywords, fresh_build = fresh_build, no_build = no_build, soft_build = soft_build, verbose = verbose, export_files = export_files) } val end_date = Date.now() val elapsed_time = end_date.time - start_date.time if (verbose) { progress.echo("\nFinished at " + Build_Log.print_date(end_date)) } val total_timing = results.sessions.iterator.map(a => results(a).timing).foldLeft(Timing.zero)(_ + _). copy(elapsed = elapsed_time) progress.echo(total_timing.message_resources) sys.exit(results.rc) }) /* build logic image */ def build_logic(options: Options, logic: String, progress: Progress = new Progress, build_heap: Boolean = false, dirs: List[Path] = Nil, fresh: Boolean = false, strict: Boolean = false): Int = { val selection = Sessions.Selection.session(logic) val rc = if (!fresh && build(options, selection = selection, build_heap = build_heap, no_build = true, dirs = dirs).ok) 0 else { progress.echo("Build started for Isabelle/" + logic + " ...") Build.build(options, selection = selection, progress = progress, build_heap = build_heap, fresh_build = fresh, dirs = dirs).rc } if (strict && rc != 0) error("Failed to build Isabelle/" + logic) else rc } } diff --git a/src/Pure/Tools/build_job.scala b/src/Pure/Tools/build_job.scala --- a/src/Pure/Tools/build_job.scala +++ b/src/Pure/Tools/build_job.scala @@ -1,538 +1,538 @@ /* Title: Pure/Tools/build_job.scala Author: Makarius Build job running prover process, with rudimentary PIDE session. */ package isabelle import scala.collection.mutable object Build_Job { /* theory markup/messages from database */ def read_theory( db_context: Sessions.Database_Context, resources: Resources, session: String, theory: String, unicode_symbols: Boolean = false): Option[Command] = { def read(name: String): Export.Entry = db_context.get_export(List(session), theory, name) def read_xml(name: String): XML.Body = YXML.parse_body( Symbol.output(unicode_symbols, UTF8.decode_permissive(read(name).uncompressed)), cache = db_context.cache) (read(Export.DOCUMENT_ID).text, split_lines(read(Export.FILES).text)) match { case (Value.Long(id), thy_file :: blobs_files) => val node_name = resources.file_node(Path.explode(thy_file), theory = theory) val results = Command.Results.make( for (elem @ XML.Elem(Markup(_, Markup.Serial(i)), _) <- read_xml(Export.MESSAGES)) yield i -> elem) val blobs = blobs_files.map(file => { val path = Path.explode(file) val name = resources.file_node(path) val src_path = File.relative_path(node_name.master_dir_path, path).getOrElse(path) Command.Blob(name, src_path, None) }) val blobs_xml = for (i <- (1 to blobs.length).toList) yield read_xml(Export.MARKUP + i) val blobs_info = Command.Blobs_Info( for { (Command.Blob(name, src_path, _), xml) <- blobs zip blobs_xml } yield { val text = XML.content(xml) val chunk = Symbol.Text_Chunk(text) val digest = SHA1.digest(Symbol.encode(text)) Exn.Res(Command.Blob(name, src_path, Some((digest, chunk)))) }) val thy_xml = read_xml(Export.MARKUP) val thy_source = XML.content(thy_xml) val markups_index = Command.Markup_Index.markup :: blobs.map(Command.Markup_Index.blob) val markups = Command.Markups.make( for ((index, xml) <- markups_index.zip(thy_xml :: blobs_xml)) yield index -> Markup_Tree.from_XML(xml)) val command = Command.unparsed(thy_source, theory = true, id = id, node_name = node_name, blobs_info = blobs_info, results = results, markups = markups) Some(command) case _ => None } } /* print messages */ def print_log( options: Options, session_name: String, theories: List[String] = Nil, verbose: Boolean = false, progress: Progress = new Progress, margin: Double = Pretty.default_margin, breakgain: Double = Pretty.default_breakgain, metric: Pretty.Metric = Symbol.Metric, unicode_symbols: Boolean = false): Unit = { val store = Sessions.store(options) val resources = Resources.empty val session = new Session(options, resources) using(store.open_database_context())(db_context => { val result = db_context.input_database(session_name)((db, _) => { val theories = store.read_theories(db, session_name) val errors = store.read_errors(db, session_name) store.read_build(db, session_name).map(info => (theories, errors, info.return_code)) }) result match { case None => error("Missing build database for session " + quote(session_name)) case Some((used_theories, errors, rc)) => val bad_theories = theories.filterNot(used_theories.toSet) if (bad_theories.nonEmpty) error("Unknown theories " + commas_quote(bad_theories)) val print_theories = if (theories.isEmpty) used_theories else used_theories.filter(theories.toSet) for (thy <- print_theories) { val thy_heading = "\nTheory " + quote(thy) + ":" read_theory(db_context, resources, session_name, thy, unicode_symbols = unicode_symbols) match { case None => progress.echo(thy_heading + " MISSING") case Some(command) => val snapshot = Document.State.init.snippet(command) val rendering = new Rendering(snapshot, options, session) val messages = rendering.text_messages(Text.Range.full) .filter(message => verbose || Protocol.is_exported(message.info)) if (messages.nonEmpty) { val line_document = Line.Document(command.source) progress.echo(thy_heading) for (Text.Info(range, elem) <- messages) { val line = line_document.position(range.start).line1 val pos = Position.Line_File(line, command.node_name.node) progress.echo( Protocol.message_text(elem, heading = true, pos = pos, margin = margin, breakgain = breakgain, metric = metric)) } } } } if (errors.nonEmpty) { val msg = Symbol.output(unicode_symbols, cat_lines(errors)) progress.echo("\nBuild errors:\n" + Output.error_message_text(msg)) } if (rc != 0) progress.echo("\n" + Process_Result.print_return_code(rc)) } }) } /* Isabelle tool wrapper */ val isabelle_tool = Isabelle_Tool("log", "print messages from build database", Scala_Project.here, args => { /* arguments */ var unicode_symbols = false var theories: List[String] = Nil var margin = Pretty.default_margin var options = Options.init() var verbose = false val getopts = Getopts(""" Usage: isabelle log [OPTIONS] SESSION Options are: -T NAME restrict to given theories (multiple options possible) -U output Unicode symbols -m MARGIN margin for pretty printing (default: """ + margin + """) -o OPTION override Isabelle system OPTION (via NAME=VAL or NAME) -v print all messages, including information, tracing etc. Print messages from the build database of the given session, without any checks against current sources: results from a failed build can be printed as well. """, "T:" -> (arg => theories = theories ::: List(arg)), "U" -> (_ => unicode_symbols = true), "m:" -> (arg => margin = Value.Double.parse(arg)), "o:" -> (arg => options = options + arg), "v" -> (_ => verbose = true)) val more_args = getopts(args) val session_name = more_args match { case List(session_name) => session_name case _ => getopts.usage() } val progress = new Console_Progress() print_log(options, session_name, theories = theories, verbose = verbose, margin = margin, progress = progress, unicode_symbols = unicode_symbols) }) } class Build_Job(progress: Progress, session_name: String, val info: Sessions.Info, deps: Sessions.Deps, store: Sessions.Store, do_store: Boolean, verbose: Boolean, val numa_node: Option[Int], command_timings0: List[Properties.T]) { val options: Options = NUMA.policy_options(info.options, numa_node) private val sessions_structure = deps.sessions_structure private val future_result: Future[Process_Result] = Future.thread("build", uninterruptible = true) { val parent = info.parent.getOrElse("") val base = deps(parent) val result_base = deps(session_name) val env = Isabelle_System.settings() + ("ISABELLE_ML_DEBUGGER" -> options.bool("ML_debugger").toString) val is_pure = Sessions.is_pure(session_name) val use_prelude = if (is_pure) Thy_Header.ml_roots.map(_._1) else Nil val eval_store = if (do_store) { (if (info.theories.nonEmpty) List("ML_Heap.share_common_data ()") else Nil) ::: List("ML_Heap.save_child " + ML_Syntax.print_string_bytes(File.platform_path(store.output_heap(session_name)))) } else Nil val resources = new Resources(sessions_structure, base, command_timings = command_timings0) val session = new Session(options, resources) { override val cache: XML.Cache = store.cache override def build_blobs_info(name: Document.Node.Name): Command.Blobs_Info = { result_base.load_commands.get(name.expand) match { case Some(spans) => val syntax = result_base.theory_syntax(name) Command.build_blobs_info(syntax, name, spans) case None => Command.Blobs_Info.none } } } object Build_Session_Errors { private val promise: Promise[List[String]] = Future.promise def result: Exn.Result[List[String]] = promise.join_result - def cancel: Unit = promise.cancel + def cancel(): Unit = promise.cancel() def apply(errs: List[String]): Unit = { try { promise.fulfill(errs) } catch { case _: IllegalStateException => } } } val export_consumer = Export.consumer(store.open_database(session_name, output = true), store.cache) val stdout = new StringBuilder(1000) val stderr = new StringBuilder(1000) val command_timings = new mutable.ListBuffer[Properties.T] val theory_timings = new mutable.ListBuffer[Properties.T] val session_timings = new mutable.ListBuffer[Properties.T] val runtime_statistics = new mutable.ListBuffer[Properties.T] val task_statistics = new mutable.ListBuffer[Properties.T] def fun( name: String, acc: mutable.ListBuffer[Properties.T], unapply: Properties.T => Option[Properties.T]): (String, Session.Protocol_Function) = { name -> ((msg: Prover.Protocol_Output) => unapply(msg.properties) match { case Some(props) => acc += props; true case _ => false }) } session.init_protocol_handler(new Session.Protocol_Handler { - override def exit(): Unit = Build_Session_Errors.cancel + override def exit(): Unit = Build_Session_Errors.cancel() private def build_session_finished(msg: Prover.Protocol_Output): Boolean = { val (rc, errors) = try { val (rc, errs) = { import XML.Decode._ pair(int, list(x => x))(Symbol.decode_yxml(msg.text)) } val errors = for (err <- errs) yield { val prt = Protocol_Message.expose_no_reports(err) Pretty.string_of(prt, metric = Symbol.Metric) } (rc, errors) } catch { case ERROR(err) => (2, List(err)) } session.protocol_command("Prover.stop", rc.toString) Build_Session_Errors(errors) true } private def loading_theory(msg: Prover.Protocol_Output): Boolean = msg.properties match { case Markup.Loading_Theory(Markup.Name(name)) => progress.theory(Progress.Theory(name, session = session_name)) false case _ => false } private def export(msg: Prover.Protocol_Output): Boolean = msg.properties match { case Protocol.Export(args) => export_consumer(session_name, args, msg.bytes) true case _ => false } override val functions = List( Markup.Build_Session_Finished.name -> build_session_finished, Markup.Loading_Theory.name -> loading_theory, Markup.EXPORT -> export, fun(Markup.Theory_Timing.name, theory_timings, Markup.Theory_Timing.unapply), fun(Markup.Session_Timing.name, session_timings, Markup.Session_Timing.unapply), fun(Markup.Task_Statistics.name, task_statistics, Markup.Task_Statistics.unapply)) }) session.command_timings += Session.Consumer("command_timings") { case Session.Command_Timing(props) => for { elapsed <- Markup.Elapsed.unapply(props) elapsed_time = Time.seconds(elapsed) if elapsed_time.is_relevant && elapsed_time >= options.seconds("command_timing_threshold") } command_timings += props.filter(Markup.command_timing_property) } session.runtime_statistics += Session.Consumer("ML_statistics") { case Session.Runtime_Statistics(props) => runtime_statistics += props } session.finished_theories += Session.Consumer[Document.Snapshot]("finished_theories") { case snapshot => val rendering = new Rendering(snapshot, options, session) def export(name: String, xml: XML.Body, compress: Boolean = true): Unit = { val theory_name = snapshot.node_name.theory val args = Protocol.Export.Args(theory_name = theory_name, name = name, compress = compress) val bytes = Bytes(Symbol.encode(YXML.string_of_body(xml))) if (!bytes.is_empty) export_consumer(session_name, args, bytes) } def export_text(name: String, text: String, compress: Boolean = true): Unit = export(name, List(XML.Text(text)), compress = compress) for (command <- snapshot.snippet_command) { export_text(Export.DOCUMENT_ID, command.id.toString, compress = false) } export_text(Export.FILES, cat_lines(snapshot.node_files.map(_.symbolic.node)), compress = false) for (((_, xml), i) <- snapshot.xml_markup_blobs().zipWithIndex) { export(Export.MARKUP + (i + 1), xml) } export(Export.MARKUP, snapshot.xml_markup()) export(Export.MESSAGES, snapshot.messages.map(_._1)) val citations = Library.distinct(rendering.citations(Text.Range.full).map(_.info)) export_text(Export.CITATIONS, cat_lines(citations)) } session.all_messages += Session.Consumer[Any]("build_session_output") { case msg: Prover.Output => val message = msg.message if (msg.is_stdout) { stdout ++= Symbol.encode(XML.content(message)) } else if (msg.is_stderr) { stderr ++= Symbol.encode(XML.content(message)) } else if (msg.is_exit) { val err = "Prover terminated" + (msg.properties match { case Markup.Process_Result(result) => ": " + result.print_rc case _ => "" }) Build_Session_Errors(List(err)) } case _ => } val eval_main = Command_Line.ML_tool("Isabelle_Process.init_build ()" :: eval_store) val process = Isabelle_Process(session, options, sessions_structure, store, logic = parent, raw_ml_system = is_pure, use_prelude = use_prelude, eval_main = eval_main, cwd = info.dir.file, env = env) val build_errors = - Isabelle_Thread.interrupt_handler(_ => process.terminate) { - Exn.capture { process.await_startup } match { + Isabelle_Thread.interrupt_handler(_ => process.terminate()) { + Exn.capture { process.await_startup() } match { case Exn.Res(_) => val resources_yxml = resources.init_session_yxml val args_yxml = YXML.string_of_body( { import XML.Encode._ pair(string, list(pair(Options.encode, list(pair(string, properties)))))( (session_name, info.theories)) }) session.protocol_command("build_session", resources_yxml, args_yxml) Build_Session_Errors.result case Exn.Exn(exn) => Exn.Res(List(Exn.message(exn))) } } val process_result = - Isabelle_Thread.interrupt_handler(_ => process.terminate) { process.await_shutdown } + Isabelle_Thread.interrupt_handler(_ => process.terminate()) { process.await_shutdown() } session.stop() val export_errors = export_consumer.shutdown(close = true).map(Output.error_message_text) val (document_output, document_errors) = try { if (build_errors.isInstanceOf[Exn.Res[_]] && process_result.ok && info.documents.nonEmpty) { using(store.open_database_context())(db_context => { val documents = Presentation.build_documents(session_name, deps, db_context, output_sources = info.document_output, output_pdf = info.document_output, progress = progress, verbose = verbose) db_context.output_database(session_name)(db => documents.foreach(_.write(db, session_name))) (documents.flatMap(_.log_lines), Nil) }) } else (Nil, Nil) } catch { case exn: Presentation.Build_Error => (exn.log_lines, List(exn.message)) case Exn.Interrupt.ERROR(msg) => (Nil, List(msg)) } val result = { val theory_timing = theory_timings.iterator.map( { case props @ Markup.Name(name) => name -> props }).toMap val used_theory_timings = for { (name, _) <- deps(session_name).used_theories } yield theory_timing.getOrElse(name.theory, Markup.Name(name.theory)) val more_output = Library.trim_line(stdout.toString) :: command_timings.toList.map(Protocol.Command_Timing_Marker.apply) ::: used_theory_timings.map(Protocol.Theory_Timing_Marker.apply) ::: session_timings.toList.map(Protocol.Session_Timing_Marker.apply) ::: runtime_statistics.toList.map(Protocol.ML_Statistics_Marker.apply) ::: task_statistics.toList.map(Protocol.Task_Statistics_Marker.apply) ::: document_output process_result.output(more_output) .error(Library.trim_line(stderr.toString)) .errors_rc(export_errors ::: document_errors) } build_errors match { case Exn.Res(build_errs) => val errs = build_errs ::: document_errors if (errs.isEmpty) result else { result.error_rc.output( errs.flatMap(s => split_lines(Output.error_message_text(s))) ::: errs.map(Protocol.Error_Message_Marker.apply)) } case Exn.Exn(Exn.Interrupt()) => if (result.ok) result.copy(rc = Exn.Interrupt.return_code) else result case Exn.Exn(exn) => throw exn } } - def terminate: Unit = future_result.cancel + def terminate(): Unit = future_result.cancel() def is_finished: Boolean = future_result.is_finished private val timeout_request: Option[Event_Timer.Request] = { if (info.timeout > Time.zero) - Some(Event_Timer.request(Time.now() + info.timeout) { terminate }) + Some(Event_Timer.request(Time.now() + info.timeout) { terminate() }) else None } def join: (Process_Result, Option[String]) = { val result1 = future_result.join val was_timeout = timeout_request match { case None => false - case Some(request) => !request.cancel + case Some(request) => !request.cancel() } val result2 = if (result1.ok) result1 else if (was_timeout) result1.error(Output.error_message_text("Timeout")).timeout_rc else if (result1.interrupted) result1.error(Output.error_message_text("Interrupt")) else result1 val heap_digest = if (result2.ok && do_store && store.output_heap(session_name).is_file) Some(Sessions.write_heap_digest(store.output_heap(session_name))) else None (result2, heap_digest) } } diff --git a/src/Pure/Tools/server.scala b/src/Pure/Tools/server.scala --- a/src/Pure/Tools/server.scala +++ b/src/Pure/Tools/server.scala @@ -1,598 +1,598 @@ /* Title: Pure/Tools/server.scala Author: Makarius Resident Isabelle servers. Message formats: - short message (single line): NAME ARGUMENT - long message (multiple lines): BYTE_LENGTH NAME ARGUMENT Argument formats: - Unit as empty string - XML.Elem in YXML notation - JSON.T in standard notation */ package isabelle import java.io.{BufferedInputStream, BufferedOutputStream, InputStreamReader, OutputStreamWriter, IOException} import java.net.{Socket, SocketException, SocketTimeoutException, ServerSocket, InetAddress} object Server { /* message argument */ object Argument { def is_name_char(c: Char): Boolean = Symbol.is_ascii_letter(c) || Symbol.is_ascii_digit(c) || c == '_' || c == '.' def split(msg: String): (String, String) = { val name = msg.takeWhile(is_name_char) val argument = msg.substring(name.length).dropWhile(Symbol.is_ascii_blank) (name, argument) } def print(arg: Any): String = arg match { case () => "" case t: XML.Elem => YXML.string_of_tree(t) case t: JSON.T => JSON.Format(t) } def parse(argument: String): Any = if (argument == "") () else if (YXML.detect_elem(argument)) YXML.parse_elem(argument) else JSON.parse(argument, strict = false) def unapply(argument: String): Option[Any] = try { Some(parse(argument)) } catch { case ERROR(_) => None } } /* input command */ type Command_Body = PartialFunction[(Context, Any), Any] abstract class Command(val command_name: String) { def command_body: Command_Body override def toString: String = command_name } class Commands(commands: Command*) extends Isabelle_System.Service { def entries: List[Command] = commands.toList } private lazy val command_table: Map[String, Command] = Isabelle_System.make_services(classOf[Commands]).flatMap(_.entries). foldLeft(Map.empty[String, Command]) { case (cmds, cmd) => val name = cmd.command_name if (cmds.isDefinedAt(name)) error("Duplicate Isabelle server command: " + quote(name)) else cmds + (name -> cmd) } /* output reply */ class Error(val message: String, val json: JSON.Object.T = JSON.Object.empty) extends RuntimeException(message) def json_error(exn: Throwable): JSON.Object.T = exn match { case e: Error => Reply.error_message(e.message) ++ e.json case ERROR(msg) => Reply.error_message(msg) case _ if Exn.is_interrupt(exn) => Reply.error_message(Exn.message(exn)) case _ => JSON.Object.empty } object Reply extends Enumeration { val OK, ERROR, FINISHED, FAILED, NOTE = Value def message(msg: String, kind: String = ""): JSON.Object.T = JSON.Object(Markup.KIND -> proper_string(kind).getOrElse(Markup.WRITELN), "message" -> msg) def error_message(msg: String): JSON.Object.T = message(msg, kind = Markup.ERROR) def unapply(msg: String): Option[(Reply.Value, Any)] = { if (msg == "") None else { val (name, argument) = Argument.split(msg) for { reply <- try { Some(withName(name)) } catch { case _: NoSuchElementException => None } arg <- Argument.unapply(argument) } yield (reply, arg) } } } /* handler: port, password, thread */ abstract class Handler(port0: Int) { val socket: ServerSocket = new ServerSocket(port0, 50, Server.localhost) def port: Int = socket.getLocalPort val password: String = UUID.random_string() override def toString: String = print(port, password) def handle(connection: Server.Connection): Unit private lazy val thread: Thread = Isabelle_Thread.fork(name = "server_handler") { var finished = false while (!finished) { Exn.capture(socket.accept) match { case Exn.Res(client) => Isabelle_Thread.fork(name = "client") { using(Connection(client))(connection => if (connection.read_password(password)) handle(connection)) } case Exn.Exn(_) => finished = true } } } - def start: Unit = thread + def start(): Unit = thread def join: Unit = thread.join - def stop: Unit = { socket.close; join } + def stop(): Unit = { socket.close(); join } } /* socket connection */ object Connection { def apply(socket: Socket): Connection = new Connection(socket) } class Connection private(socket: Socket) extends AutoCloseable { override def toString: String = socket.toString - def close(): Unit = socket.close + def close(): Unit = socket.close() def set_timeout(t: Time): Unit = socket.setSoTimeout(t.ms.toInt) private val in = new BufferedInputStream(socket.getInputStream) private val out = new BufferedOutputStream(socket.getOutputStream) private val out_lock: AnyRef = new Object def tty_loop(): TTY_Loop = new TTY_Loop( new OutputStreamWriter(out), new InputStreamReader(in), writer_lock = out_lock) def read_password(password: String): Boolean = try { Byte_Message.read_line(in).map(_.text) == Some(password) } catch { case _: IOException => false } def read_message(): Option[String] = try { Byte_Message.read_line_message(in).map(_.text) } catch { case _: IOException => None } def await_close(): Unit = try { Byte_Message.read(in, 1); socket.close() } catch { case _: IOException => } def write_message(msg: String): Unit = out_lock.synchronized { Byte_Message.write_line_message(out, Bytes(UTF8.bytes(msg))) } def reply(r: Reply.Value, arg: Any): Unit = { val argument = Argument.print(arg) write_message(if (argument == "") r.toString else r.toString + " " + argument) } def reply_ok(arg: Any): Unit = reply(Reply.OK, arg) def reply_error(arg: Any): Unit = reply(Reply.ERROR, arg) def reply_error_message(message: String, more: JSON.Object.Entry*): Unit = reply_error(Reply.error_message(message) ++ more) def notify(arg: Any): Unit = reply(Reply.NOTE, arg) } /* context with output channels */ class Context private[Server](val server: Server, connection: Connection) extends AutoCloseable { context => def command_list: List[String] = command_table.keys.toList.sorted def reply(r: Reply.Value, arg: Any): Unit = connection.reply(r, arg) def notify(arg: Any): Unit = connection.notify(arg) def message(kind: String, msg: String, more: JSON.Object.Entry*): Unit = notify(Reply.message(msg, kind = kind) ++ more) def writeln(msg: String, more: JSON.Object.Entry*): Unit = message(Markup.WRITELN, msg, more:_*) def warning(msg: String, more: JSON.Object.Entry*): Unit = message(Markup.WARNING, msg, more:_*) def error_message(msg: String, more: JSON.Object.Entry*): Unit = message(Markup.ERROR, msg, more:_*) def progress(more: JSON.Object.Entry*): Connection_Progress = new Connection_Progress(context, more:_*) override def toString: String = connection.toString /* asynchronous tasks */ private val _tasks = Synchronized(Set.empty[Task]) def make_task(body: Task => JSON.Object.T): Task = { val task = new Task(context, body) _tasks.change(_ + task) task } def remove_task(task: Task): Unit = _tasks.change(_ - task) def cancel_task(id: UUID.T): Unit = - _tasks.change(tasks => { tasks.find(task => task.id == id).foreach(_.cancel); tasks }) + _tasks.change(tasks => { tasks.find(task => task.id == id).foreach(_.cancel()); tasks }) def close(): Unit = { - while(_tasks.change_result(tasks => { tasks.foreach(_.cancel); (tasks.nonEmpty, tasks) })) + while(_tasks.change_result(tasks => { tasks.foreach(_.cancel()); (tasks.nonEmpty, tasks) })) { _tasks.value.foreach(_.join) } } } class Connection_Progress private[Server](context: Context, more: JSON.Object.Entry*) extends Progress { override def echo(msg: String): Unit = context.writeln(msg, more:_*) override def echo_warning(msg: String): Unit = context.warning(msg, more:_*) override def echo_error_message(msg: String): Unit = context.error_message(msg, more:_*) override def theory(theory: Progress.Theory): Unit = { val entries: List[JSON.Object.Entry] = List("theory" -> theory.theory, "session" -> theory.session) ::: (theory.percentage match { case None => Nil case Some(p) => List("percentage" -> p) }) context.writeln(theory.message, entries ::: more.toList:_*) } override def nodes_status(nodes_status: Document_Status.Nodes_Status): Unit = { val json = for ((name, node_status) <- nodes_status.present) yield name.json + ("status" -> node_status.json) context.notify(JSON.Object(Markup.KIND -> Markup.NODES_STATUS, Markup.NODES_STATUS -> json)) } override def toString: String = context.toString } class Task private[Server](val context: Context, body: Task => JSON.Object.T) { task => val id: UUID.T = UUID.random() val ident: JSON.Object.Entry = ("task" -> id.toString) val progress: Connection_Progress = context.progress(ident) - def cancel: Unit = progress.stop + def cancel(): Unit = progress.stop() private lazy val thread = Isabelle_Thread.fork(name = "server_task") { Exn.capture { body(task) } match { case Exn.Res(res) => context.reply(Reply.FINISHED, res + ident) case Exn.Exn(exn) => val err = json_error(exn) if (err.isEmpty) throw exn else context.reply(Reply.FAILED, err + ident) } - progress.stop + progress.stop() context.remove_task(task) } def start: Unit = thread def join: Unit = thread.join } /* server info */ val localhost_name: String = "127.0.0.1" def localhost: InetAddress = InetAddress.getByName(localhost_name) def print_address(port: Int): String = localhost_name + ":" + port def print(port: Int, password: String): String = print_address(port) + " (password " + quote(password) + ")" object Info { private val Pattern = ("""server "([^"]*)" = \Q""" + localhost_name + """\E:(\d+) \(password "([^"]*)"\)""").r def parse(s: String): Option[Info] = s match { case Pattern(name, Value.Int(port), password) => Some(Info(name, port, password)) case _ => None } def apply(name: String, port: Int, password: String): Info = new Info(name, port, password) } class Info private(val name: String, val port: Int, val password: String) { def address: String = print_address(port) override def toString: String = "server " + quote(name) + " = " + print(port, password) def connection(): Connection = { val connection = Connection(new Socket(localhost, port)) connection.write_message(password) connection } - def active(): Boolean = + def active: Boolean = try { using(connection())(connection => { connection.set_timeout(Time.seconds(2.0)) connection.read_message() match { case Some(Reply(Reply.OK, _)) => true case _ => false } }) } catch { case _: IOException => false case _: SocketException => false case _: SocketTimeoutException => false } } /* per-user servers */ val default_name = "isabelle" object Data { val database = Path.explode("$ISABELLE_HOME_USER/servers.db") val name = SQL.Column.string("name").make_primary_key val port = SQL.Column.int("port") val password = SQL.Column.string("password") val table = SQL.Table("isabelle_servers", List(name, port, password)) } def list(db: SQLite.Database): List[Info] = if (db.tables.contains(Data.table.name)) { db.using_statement(Data.table.select())(stmt => stmt.execute_query().iterator(res => Info( res.string(Data.name), res.int(Data.port), res.string(Data.password))).toList.sortBy(_.name)) } else Nil def find(db: SQLite.Database, name: String): Option[Info] = list(db).find(server_info => server_info.name == name && server_info.active) def init( name: String = default_name, port: Int = 0, existing_server: Boolean = false, log: Logger = No_Logger): (Info, Option[Server]) = { using(SQLite.open_database(Data.database))(db => { db.transaction { Isabelle_System.chmod("600", Data.database) db.create_table(Data.table) list(db).filterNot(_.active).foreach(server_info => db.using_statement(Data.table.delete(Data.name.where_equal(server_info.name)))( - _.execute)) + _.execute())) } db.transaction { find(db, name) match { case Some(server_info) => (server_info, None) case None => if (existing_server) error("Isabelle server " + quote(name) + " not running") val server = new Server(port, log) val server_info = Info(name, server.port, server.password) - db.using_statement(Data.table.delete(Data.name.where_equal(name)))(_.execute) + db.using_statement(Data.table.delete(Data.name.where_equal(name)))(_.execute()) db.using_statement(Data.table.insert())(stmt => { stmt.string(1) = server_info.name stmt.int(2) = server_info.port stmt.string(3) = server_info.password stmt.execute() }) - server.start + server.start() (server_info, Some(server)) } } }) } def exit(name: String = default_name): Boolean = { using(SQLite.open_database(Data.database))(db => db.transaction { find(db, name) match { case Some(server_info) => using(server_info.connection())(_.write_message("shutdown")) while(server_info.active) { Time.seconds(0.05).sleep } true case None => false } }) } /* Isabelle tool wrapper */ val isabelle_tool = Isabelle_Tool("server", "manage resident Isabelle servers", Scala_Project.here, args => { var console = false var log_file: Option[Path] = None var operation_list = false var operation_exit = false var name = default_name var port = 0 var existing_server = false val getopts = Getopts(""" Usage: isabelle server [OPTIONS] Options are: -L FILE logging on FILE -c console interaction with specified server -l list servers (alternative operation) -n NAME explicit server name (default: """ + default_name + """) -p PORT explicit server port -s assume existing server, no implicit startup -x exit specified server (alternative operation) Manage resident Isabelle servers. """, "L:" -> (arg => log_file = Some(Path.explode(File.standard_path(arg)))), "c" -> (_ => console = true), "l" -> (_ => operation_list = true), "n:" -> (arg => name = arg), "p:" -> (arg => port = Value.Int.parse(arg)), "s" -> (_ => existing_server = true), "x" -> (_ => operation_exit = true)) val more_args = getopts(args) if (more_args.nonEmpty) getopts.usage() if (operation_list) { for { server_info <- using(SQLite.open_database(Data.database))(list) if server_info.active } Output.writeln(server_info.toString, stdout = true) } else if (operation_exit) { val ok = Server.exit(name) sys.exit(if (ok) 0 else 2) } else { val log = Logger.make(log_file) val (server_info, server) = init(name, port = port, existing_server = existing_server, log = log) Output.writeln(server_info.toString, stdout = true) if (console) { using(server_info.connection())(connection => connection.tty_loop().join) } server.foreach(_.join) } }) } class Server private(port0: Int, val log: Logger) extends Server.Handler(port0) { server => private val _sessions = Synchronized(Map.empty[UUID.T, Headless.Session]) def err_session(id: UUID.T): Nothing = error("No session " + Library.single_quote(id.toString)) def the_session(id: UUID.T): Headless.Session = _sessions.value.getOrElse(id, err_session(id)) def add_session(entry: (UUID.T, Headless.Session)): Unit = _sessions.change(_ + entry) def remove_session(id: UUID.T): Headless.Session = { _sessions.change_result(sessions => sessions.get(id) match { case Some(session) => (session, sessions - id) case None => err_session(id) }) } def shutdown(): Unit = { - server.socket.close + server.socket.close() val sessions = _sessions.change_result(sessions => (sessions, Map.empty)) for ((_, session) <- sessions) { try { val result = session.stop() if (!result.ok) log("Session shutdown failed: " + result.print_rc) } catch { case ERROR(msg) => log("Session shutdown failed: " + msg) } } } override def join: Unit = { super.join; shutdown() } override def handle(connection: Server.Connection): Unit = { using(new Server.Context(server, connection))(context => { connection.reply_ok( JSON.Object( "isabelle_id" -> Isabelle_System.isabelle_id(), "isabelle_version" -> Distribution.version)) var finished = false while (!finished) { connection.read_message() match { case None => finished = true case Some("") => context.notify("Command 'help' provides list of commands") case Some(msg) => val (name, argument) = Server.Argument.split(msg) Server.command_table.get(name) match { case Some(cmd) => argument match { case Server.Argument(arg) => if (cmd.command_body.isDefinedAt((context, arg))) { Exn.capture { cmd.command_body((context, arg)) } match { case Exn.Res(task: Server.Task) => connection.reply_ok(JSON.Object(task.ident)) task.start case Exn.Res(res) => connection.reply_ok(res) case Exn.Exn(exn) => val err = Server.json_error(exn) if (err.isEmpty) throw exn else connection.reply_error(err) } } else { connection.reply_error_message( "Bad argument for command " + Library.single_quote(name), "argument" -> argument) } case _ => connection.reply_error_message( "Malformed argument for command " + Library.single_quote(name), "argument" -> argument) } case None => connection.reply_error("Bad command " + Library.single_quote(name)) } } } }) } } diff --git a/src/Pure/Tools/spell_checker.scala b/src/Pure/Tools/spell_checker.scala --- a/src/Pure/Tools/spell_checker.scala +++ b/src/Pure/Tools/spell_checker.scala @@ -1,289 +1,289 @@ /* Title: Pure/Tools/spell_checker.scala Author: Makarius Spell checker with completion, based on JOrtho (see https://sourceforge.net/projects/jortho). */ package isabelle import java.lang.Class import scala.collection.mutable import scala.annotation.tailrec import scala.collection.immutable.SortedMap object Spell_Checker { /* words within text */ def marked_words(base: Text.Offset, text: String, mark: Text.Info[String] => Boolean) : List[Text.Info[String]] = { val result = new mutable.ListBuffer[Text.Info[String]] var offset = 0 def apostrophe(c: Int): Boolean = c == '\'' && (offset + 1 == text.length || text(offset + 1) != '\'') @tailrec def scan(pred: Int => Boolean): Unit = { if (offset < text.length) { val c = text.codePointAt(offset) if (pred(c)) { offset += Character.charCount(c) scan(pred) } } } while (offset < text.length) { scan(c => !Character.isLetter(c)) val start = offset scan(c => Character.isLetterOrDigit(c) || apostrophe(c)) val stop = offset if (stop - start >= 2) { val info = Text.Info(Text.Range(base + start, base + stop), text.substring(start, stop)) if (mark(info)) result += info } } result.toList } def current_word(rendering: Rendering, range: Text.Range): Option[Text.Info[String]] = { for { spell_range <- rendering.spell_checker_point(range) text <- rendering.get_text(spell_range) info <- marked_words(spell_range.start, text, info => info.range.overlaps(range)).headOption } yield info } /* dictionaries */ class Dictionary private[Spell_Checker](val path: Path) { val lang: String = path.drop_ext.file_name val user_path: Path = Path.explode("$ISABELLE_HOME_USER/dictionaries") + Path.basic(lang) override def toString: String = lang } private object Decl { def apply(name: String, include: Boolean): String = if (include) name else "-" + name def unapply(decl: String): Option[(String, Boolean)] = { val decl1 = decl.trim if (decl1 == "" || decl1.startsWith("#")) None else Library.try_unprefix("-", decl1.trim) match { case None => Some((decl1, true)) case Some(decl2) => Some((decl2, false)) } } } - def dictionaries(): List[Dictionary] = + def dictionaries: List[Dictionary] = for { path <- Path.split(Isabelle_System.getenv("JORTHO_DICTIONARIES")) if path.is_file } yield new Dictionary(path) /* create spell checker */ def apply(dictionary: Dictionary): Spell_Checker = new Spell_Checker(dictionary) private sealed case class Update(include: Boolean, permanent: Boolean) } class Spell_Checker private(dictionary: Spell_Checker.Dictionary) { override def toString: String = dictionary.toString /* main dictionary content */ private var dict = new Object private var updates = SortedMap.empty[String, Spell_Checker.Update] private def included_iterator(): Iterator[String] = for { (word, upd) <- updates.iterator if upd.include } yield word private def excluded(word: String): Boolean = updates.get(word) match { case Some(upd) => !upd.include case None => false } private def load(): Unit = { val main_dictionary = split_lines(File.read_gzip(dictionary.path)) val permanent_updates = if (dictionary.user_path.is_file) for { Spell_Checker.Decl(word, include) <- split_lines(File.read(dictionary.user_path)) } yield (word, Spell_Checker.Update(include, true)) else Nil updates = updates -- (for ((name, upd) <- updates.iterator; if upd.permanent) yield name) ++ permanent_updates val factory_class = Class.forName("com.inet.jortho.DictionaryFactory") val factory = Untyped.constructor(factory_class).newInstance() val add = Untyped.method(factory_class, "add", classOf[String]) for { word <- main_dictionary.iterator ++ included_iterator() if !excluded(word) } add.invoke(factory, word) dict = Untyped.method(factory_class, "create").invoke(factory) } load() private def save(): Unit = { val permanent_decls = (for { (word, upd) <- updates.iterator if upd.permanent } yield Spell_Checker.Decl(word, upd.include)).toList if (permanent_decls.nonEmpty || dictionary.user_path.is_file) { val header = """# User updates for spell-checker dictionary # # * each line contains at most one word # * extra blanks are ignored # * lines starting with "#" are stripped # * lines starting with "-" indicate excluded words # #:mode=text:encoding=UTF-8: """ Isabelle_System.make_directory(dictionary.user_path.expand.dir) File.write(dictionary.user_path, header + cat_lines(permanent_decls)) } } def update(word: String, include: Boolean, permanent: Boolean): Unit = { updates += (word -> Spell_Checker.Update(include, permanent)) if (include) { if (permanent) save() Untyped.method(dict.getClass, "add", classOf[String]).invoke(dict, word) } else { save(); load() } } def reset(): Unit = { updates = SortedMap.empty load() } def reset_enabled(): Int = updates.valuesIterator.count(upd => !upd.permanent) /* check known words */ def contains(word: String): Boolean = Untyped.method(dict.getClass.getSuperclass, "exist", classOf[String]). invoke(dict, word).asInstanceOf[java.lang.Boolean].booleanValue def check(word: String): Boolean = word match { case Word.Case(c) if c != Word.Lowercase => contains(word) || contains(Word.lowercase(word)) case _ => contains(word) } def marked_words(base: Text.Offset, text: String): List[Text.Info[String]] = Spell_Checker.marked_words(base, text, info => !check(info.info)) /* completion: suggestions for unknown words */ private def suggestions(word: String): Option[List[String]] = { val res = Untyped.method(dict.getClass.getSuperclass, "searchSuggestions", classOf[String]). invoke(dict, word).asInstanceOf[java.util.List[AnyRef]].toArray.toList.map(_.toString) if (res.isEmpty) None else Some(res) } def complete(word: String): List[String] = if (check(word)) Nil else { val word_case = Word.Case.unapply(word) def recover_case(s: String) = word_case match { case Some(c) => Word.Case(c, s) case None => s } val result = word_case match { case Some(c) if c != Word.Lowercase => suggestions(word) orElse suggestions(Word.lowercase(word)) case _ => suggestions(word) } result.getOrElse(Nil).map(recover_case) } def completion(rendering: Rendering, caret: Text.Offset): Option[Completion.Result] = { val caret_range = rendering.before_caret_range(caret) for { word <- Spell_Checker.current_word(rendering, caret_range) words = complete(word.info) if words.nonEmpty descr = "(from dictionary " + quote(dictionary.toString) + ")" items = words.map(w => Completion.Item(word.range, word.info, "", List(w, descr), w, 0, false)) } yield Completion.Result(word.range, word.info, false, items) } } class Spell_Checker_Variable { private val no_spell_checker: (String, Option[Spell_Checker]) = ("", None) private var current_spell_checker = no_spell_checker def get: Option[Spell_Checker] = synchronized { current_spell_checker._2 } def update(options: Options): Unit = synchronized { if (options.bool("spell_checker")) { val lang = options.string("spell_checker_dictionary") if (current_spell_checker._1 != lang) { Spell_Checker.dictionaries.find(_.lang == lang) match { case Some(dictionary) => val spell_checker = Exn.capture { Spell_Checker(dictionary) } match { case Exn.Res(spell_checker) => Some(spell_checker) case Exn.Exn(_) => None } current_spell_checker = (lang, spell_checker) case None => current_spell_checker = no_spell_checker } } } else current_spell_checker = no_spell_checker } } diff --git a/src/Tools/Graphview/graph_panel.scala b/src/Tools/Graphview/graph_panel.scala --- a/src/Tools/Graphview/graph_panel.scala +++ b/src/Tools/Graphview/graph_panel.scala @@ -1,360 +1,360 @@ /* Title: Tools/Graphview/graph_panel.scala Author: Markus Kaiser, TU Muenchen Author: Makarius GUI panel for graph layout. */ package isabelle.graphview import isabelle._ import java.awt.{Dimension, Graphics2D, Point, Rectangle} import java.awt.geom.{AffineTransform, Point2D} import javax.imageio.ImageIO import javax.swing.{JScrollPane, JComponent, SwingUtilities} import javax.swing.border.EmptyBorder import scala.swing.{BorderPanel, Button, CheckBox, Action, FileChooser, Panel, ScrollPane} import scala.swing.event.{Event, Key, MousePressed, MouseDragged, MouseClicked, MouseEvent} class Graph_Panel(val graphview: Graphview) extends BorderPanel { graph_panel => /** graph **/ /* painter */ private val painter = new Panel { override def paint(gfx: Graphics2D): Unit = { super.paintComponent(gfx) gfx.setColor(graphview.background_color) gfx.fillRect(0, 0, peer.getWidth, peer.getHeight) gfx.transform(Transform()) graphview.paint(gfx) } } def set_preferred_size(): Unit = { val box = graphview.bounding_box() val s = Transform.scale_discrete painter.preferredSize = new Dimension((box.width * s).ceil.toInt, (box.height * s).ceil.toInt) painter.revalidate() } def refresh(): Unit = { if (painter != null) { set_preferred_size() painter.repaint() } } def scroll_to_node(node: Graph_Display.Node): Unit = { val gap = graphview.metrics.gap val info = graphview.layout.get_node(node) val t = Transform() val p = t.transform(new Point2D.Double(info.x - info.width2 - gap, info.y - info.height2 - gap), null) val q = t.transform(new Point2D.Double(info.x + info.width2 + gap, info.y + info.height2 + gap), null) painter.peer.scrollRectToVisible( new Rectangle(p.getX.toInt, p.getY.toInt, (q.getX - p.getX).ceil.toInt, (q.getY - p.getY).ceil.toInt)) } /* scrollable graph pane */ private val graph_pane: ScrollPane = new ScrollPane(painter) { tooltip = "" override lazy val peer: JScrollPane = new JScrollPane with SuperMixin { override def getToolTipText(event: java.awt.event.MouseEvent): String = graphview.find_node(Transform.pane_to_graph_coordinates(event.getPoint)) match { case Some(node) => graphview.model.full_graph.get_node(node) match { case Nil => null case content => graphview.make_tooltip(graph_pane.peer, event.getX, event.getY, content) } case None => null } } horizontalScrollBarPolicy = ScrollPane.BarPolicy.Always verticalScrollBarPolicy = ScrollPane.BarPolicy.Always listenTo(mouse.moves) listenTo(mouse.clicks) reactions += { case MousePressed(_, p, _, _, _) => Mouse_Interaction.pressed(p) painter.repaint() case MouseDragged(_, to, _) => Mouse_Interaction.dragged(to) painter.repaint() case e @ MouseClicked(_, p, m, n, _) => Mouse_Interaction.clicked(p, m, n, SwingUtilities.isRightMouseButton(e.peer)) painter.repaint() } contents = painter } graph_pane.peer.getVerticalScrollBar.setUnitIncrement(10) /* transform */ def rescale(s: Double): Unit = { Transform.scale = s if (zoom != null) zoom.set_item((Transform.scale_discrete * 100).floor.toInt) refresh() } def fit_to_window(): Unit = { Transform.fit_to_window() refresh() } private object Transform { private var _scale: Double = 1.0 def scale: Double = _scale def scale_=(s: Double): Unit = { _scale = (s min 10.0) max 0.1 } def scale_discrete: Double = { val font_height = GUI.line_metrics(graphview.metrics.font).getHeight.toInt (scale * font_height).floor / font_height } def apply(): AffineTransform = { val box = graphview.bounding_box() val t = AffineTransform.getScaleInstance(scale_discrete, scale_discrete) t.translate(- box.x, - box.y) t } def fit_to_window(): Unit = { if (graphview.visible_graph.is_empty) rescale(1.0) else { val box = graphview.bounding_box() rescale((size.width / box.width) min (size.height / box.height)) } } def pane_to_graph_coordinates(at: Point2D): Point2D = { val s = Transform.scale_discrete val p = Transform().inverseTransform(graph_pane.peer.getViewport.getViewPosition, null) p.setLocation(p.getX + at.getX / s, p.getY + at.getY / s) p } } /* interaction */ graphview.model.Colors.events += { case _ => painter.repaint() } graphview.model.Mutators.events += { case _ => painter.repaint() } private object Mouse_Interaction { private var draginfo: (Point, List[Graph_Display.Node], List[Layout.Dummy]) = (new Point(0, 0), Nil, Nil) def pressed(at: Point): Unit = { val c = Transform.pane_to_graph_coordinates(at) val l = graphview.find_node(c) match { case Some(node) => if (graphview.Selection.contains(node)) graphview.Selection.get() else List(node) case None => Nil } val d = l match { case Nil => graphview.find_dummy(c).toList case _ => Nil } draginfo = (at, l, d) } def dragged(to: Point): Unit = { val (from, p, d) = draginfo val s = Transform.scale_discrete val dx = to.x - from.x val dy = to.y - from.y (p, d) match { case (Nil, Nil) => val r = graph_pane.peer.getViewport.getViewRect r.translate(- dx, - dy) painter.peer.scrollRectToVisible(r) case (Nil, ds) => ds.foreach(d => graphview.translate_vertex(d, dx / s, dy / s)) case (ls, _) => ls.foreach(l => graphview.translate_vertex(Layout.Node(l), dx / s, dy / s)) } draginfo = (to, p, d) } def clicked(at: Point, m: Key.Modifiers, clicks: Int, right_click: Boolean): Unit = { val c = Transform.pane_to_graph_coordinates(at) if (clicks < 2) { if (right_click) { // FIXME if (false) { val menu = Popups(graph_panel, graphview.find_node(c), graphview.Selection.get()) menu.show(graph_pane.peer, at.x, at.y) } } else { (graphview.find_node(c), m) match { case (Some(node), Key.Modifier.Control) => graphview.Selection.add(node) case (None, Key.Modifier.Control) => case (Some(node), Key.Modifier.Shift) => graphview.Selection.add(node) case (None, Key.Modifier.Shift) => case (Some(node), _) => graphview.Selection.clear() graphview.Selection.add(node) case (None, _) => graphview.Selection.clear() } } } } } /** controls **/ private val mutator_dialog = new Mutator_Dialog(graphview, graphview.model.Mutators, "Filters", "Hide", false) private val color_dialog = new Mutator_Dialog(graphview, graphview.model.Colors, "Colorations") private val chooser = new FileChooser { fileSelectionMode = FileChooser.SelectionMode.FilesOnly title = "Save image (.png or .pdf)" } private val show_content = new CheckBox() { selected = graphview.show_content action = Action("Show content") { graphview.show_content = selected graphview.update_layout() refresh() } tooltip = "Show full node content within graph layout" } private val show_arrow_heads = new CheckBox() { selected = graphview.show_arrow_heads action = Action("Show arrow heads") { graphview.show_arrow_heads = selected painter.repaint() } tooltip = "Draw edges with explicit arrow heads" } private val show_dummies = new CheckBox() { selected = graphview.show_dummies action = Action("Show dummies") { graphview.show_dummies = selected painter.repaint() } tooltip = "Draw dummy nodes within graph layout, for easier mouse dragging" } private val save_image = new Button { action = Action("Save image") { chooser.showSaveDialog(this) match { case FileChooser.Result.Approve => try { Graph_File.write(chooser.selectedFile, graphview) } catch { case ERROR(msg) => GUI.error_dialog(this.peer, "Error", GUI.scrollable_text(msg)) } case _ => } } tooltip = "Save current graph layout as PNG or PDF" } private val zoom = new GUI.Zoom_Box { def changed = rescale(0.01 * factor) } private val fit_window = new Button { action = Action("Fit to window") { fit_to_window() } tooltip = "Zoom to fit window width and height" } private val update_layout = new Button { action = Action("Update layout") { graphview.update_layout() refresh() } tooltip = "Regenerate graph layout according to built-in algorithm" } private val editor_style = new CheckBox() { selected = graphview.editor_style action = Action("Editor style") { graphview.editor_style = selected graphview.update_layout() refresh() } tooltip = "Use editor font and colors for painting" } - private val colorations = new Button { action = Action("Colorations") { color_dialog.open } } - private val filters = new Button { action = Action("Filters") { mutator_dialog.open } } + private val colorations = new Button { action = Action("Colorations") { color_dialog.open() } } + private val filters = new Button { action = Action("Filters") { mutator_dialog.open() } } private val controls = Wrap_Panel( List(show_content, show_arrow_heads, show_dummies, save_image, zoom, fit_window, update_layout, editor_style)) // FIXME colorations, filters /** main layout **/ layout(graph_pane) = BorderPanel.Position.Center layout(controls) = BorderPanel.Position.North rescale(1.0) } diff --git a/src/Tools/Graphview/mutator_dialog.scala b/src/Tools/Graphview/mutator_dialog.scala --- a/src/Tools/Graphview/mutator_dialog.scala +++ b/src/Tools/Graphview/mutator_dialog.scala @@ -1,381 +1,381 @@ /* Title: Tools/Graphview/mutator_dialog.scala Author: Markus Kaiser, TU Muenchen Author: Makarius Mutator selection and configuration dialog. */ package isabelle.graphview import isabelle._ import java.awt.Color import java.awt.FocusTraversalPolicy import javax.swing.JColorChooser import javax.swing.border.EmptyBorder import scala.swing.{Dialog, Button, BoxPanel, Swing, Orientation, ComboBox, Action, Dimension, BorderPanel, ScrollPane, Label, CheckBox, Alignment, Component, TextField} import scala.swing.event.ValueChanged class Mutator_Dialog( graphview: Graphview, container: Mutator_Container, caption: String, reverse_caption: String = "Inverse", show_color_chooser: Boolean = true) extends Dialog { title = caption private var _panels: List[Mutator_Panel] = Nil private def panels = _panels private def panels_=(panels: List[Mutator_Panel]): Unit = { _panels = panels paint_panels() } container.events += { case Mutator_Event.Add(m) => add_panel(new Mutator_Panel(m)) case Mutator_Event.New_List(ms) => panels = get_panels(ms) } override def open(): Unit = { if (!visible) panels = get_panels(container()) - super.open + super.open() } minimumSize = new Dimension(700, 200) preferredSize = new Dimension(1000, 300) peer.setFocusTraversalPolicy(Focus_Traveral_Policy) private def get_panels(m: List[Mutator.Info]): List[Mutator_Panel] = m.filter({ case Mutator.Info(_, _, Mutator.Identity()) => false case _ => true }) .map(m => new Mutator_Panel(m)) private def get_mutators(panels: List[Mutator_Panel]): List[Mutator.Info] = panels.map(panel => panel.get_mutator) private def movePanelUp(m: Mutator_Panel) = { def moveUp(l: List[Mutator_Panel]): List[Mutator_Panel] = l match { case x :: y :: xs => if (y == m) y :: x :: xs else x :: moveUp(y :: xs) case _ => l } panels = moveUp(panels) } private def movePanelDown(m: Mutator_Panel) = { def moveDown(l: List[Mutator_Panel]): List[Mutator_Panel] = l match { case x :: y :: xs => if (x == m) y :: x :: xs else x :: moveDown(y :: xs) case _ => l } panels = moveDown(panels) } private def removePanel(m: Mutator_Panel): Unit = { panels = panels.filter(_ != m).toList } private def add_panel(m: Mutator_Panel): Unit = { panels = panels ::: List(m) } def paint_panels(): Unit = { - Focus_Traveral_Policy.clear - filter_panel.contents.clear + Focus_Traveral_Policy.clear() + filter_panel.contents.clear() panels.map(x => { filter_panel.contents += x Focus_Traveral_Policy.addAll(x.focusList) }) filter_panel.contents += Swing.VGlue Focus_Traveral_Policy.add(mutator_box.peer.asInstanceOf[java.awt.Component]) Focus_Traveral_Policy.add(add_button.peer) Focus_Traveral_Policy.add(apply_button.peer) Focus_Traveral_Policy.add(cancel_button.peer) - filter_panel.revalidate - filter_panel.repaint + filter_panel.revalidate() + filter_panel.repaint() } val filter_panel: BoxPanel = new BoxPanel(Orientation.Vertical) {} private val mutator_box = new ComboBox[Mutator](container.available) private val add_button = new Button { action = Action("Add") { add_panel( new Mutator_Panel(Mutator.Info(true, graphview.Colors.next(), mutator_box.selection.item))) } } private val apply_button = new Button { action = Action("Apply") { container(get_mutators(panels)) } } private val reset_button = new Button { action = Action("Reset") { panels = get_panels(container()) } } private val cancel_button = new Button { - action = Action("Close") { close } + action = Action("Close") { close() } } defaultButton = cancel_button private val botPanel = new BoxPanel(Orientation.Horizontal) { border = new EmptyBorder(10, 0, 0, 0) contents += mutator_box contents += Swing.RigidBox(new Dimension(10, 0)) contents += add_button contents += Swing.HGlue contents += Swing.RigidBox(new Dimension(30, 0)) contents += apply_button contents += Swing.RigidBox(new Dimension(5, 0)) contents += reset_button contents += Swing.RigidBox(new Dimension(5, 0)) contents += cancel_button } contents = new BorderPanel { border = new EmptyBorder(5, 5, 5, 5) layout(new ScrollPane(filter_panel)) = BorderPanel.Position.Center layout(botPanel) = BorderPanel.Position.South } private class Mutator_Panel(initials: Mutator.Info) extends BoxPanel(Orientation.Horizontal) { private val inputs: List[(String, Input)] = get_inputs() var focusList = List.empty[java.awt.Component] private val enabledBox = new Check_Box_Input("Enabled", initials.enabled) border = new EmptyBorder(5, 5, 5, 5) maximumSize = new Dimension(Integer.MAX_VALUE, 30) background = initials.color contents += new Label(initials.mutator.name) { preferredSize = new Dimension(175, 20) horizontalAlignment = Alignment.Left if (initials.mutator.description != "") tooltip = initials.mutator.description } contents += Swing.RigidBox(new Dimension(10, 0)) contents += enabledBox contents += Swing.RigidBox(new Dimension(5, 0)) focusList = enabledBox.peer :: focusList inputs.map({ case (n, c) => contents += Swing.RigidBox(new Dimension(10, 0)) if (n != "") { contents += new Label(n) contents += Swing.RigidBox(new Dimension(5, 0)) } contents += c.asInstanceOf[Component] focusList = c.asInstanceOf[Component].peer :: focusList }) { val t = this contents += Swing.HGlue contents += Swing.RigidBox(new Dimension(10, 0)) if (show_color_chooser) { contents += new Button { maximumSize = new Dimension(20, 20) action = Action("Color") { t.background = JColorChooser.showDialog(t.peer, "Choose new Color", t.background) } focusList = this.peer :: focusList } contents += Swing.RigidBox(new Dimension(2, 0)) } contents += new Button { maximumSize = new Dimension(20, 20) action = Action("Up") { movePanelUp(t) } focusList = this.peer :: focusList } contents += Swing.RigidBox(new Dimension(2, 0)) contents += new Button { maximumSize = new Dimension(20, 20) action = Action("Down") { movePanelDown(t) } focusList = this.peer :: focusList } contents += Swing.RigidBox(new Dimension(2, 0)) contents += new Button { maximumSize = new Dimension(20, 20) action = Action("Del") { removePanel(t) } focusList = this.peer :: focusList } } focusList = focusList.reverse def get_mutator: Mutator.Info = { val model = graphview.model val m = initials.mutator match { case Mutator.Identity() => Mutator.Identity() case Mutator.Node_Expression(r, _, _, _) => val r1 = inputs(2)._2.string Mutator.Node_Expression( if (Library.make_regex(r1).isDefined) r1 else r, inputs(3)._2.bool, // "Parents" means "Show parents" or "Matching Children" inputs(1)._2.bool, inputs(0)._2.bool) case Mutator.Node_List(_, _, _, _) => Mutator.Node_List( for { ident <- space_explode(',', inputs(2)._2.string) node <- model.find_node(ident) } yield node, inputs(3)._2.bool, // "Parents" means "Show parents" or "Matching Children" inputs(1)._2.bool, inputs(0)._2.bool) case Mutator.Edge_Endpoints(_) => (model.find_node(inputs(0)._2.string), model.find_node(inputs(1)._2.string)) match { case (Some(node1), Some(node2)) => Mutator.Edge_Endpoints((node1, node2)) case _ => Mutator.Identity() } case Mutator.Add_Node_Expression(r) => val r1 = inputs(0)._2.string Mutator.Add_Node_Expression(if (Library.make_regex(r1).isDefined) r1 else r) case Mutator.Add_Transitive_Closure(_, _) => Mutator.Add_Transitive_Closure( inputs(0)._2.bool, inputs(1)._2.bool) case _ => Mutator.Identity() } Mutator.Info(enabledBox.selected, background, m) } private def get_inputs(): List[(String, Input)] = initials.mutator match { case Mutator.Node_Expression(regex, reverse, check_parents, check_children) => List( ("", new Check_Box_Input("Parents", check_children)), ("", new Check_Box_Input("Children", check_parents)), ("Regex", new Text_Field_Input(regex, x => Library.make_regex(x).isDefined)), ("", new Check_Box_Input(reverse_caption, reverse))) case Mutator.Node_List(list, reverse, check_parents, check_children) => List( ("", new Check_Box_Input("Parents", check_children)), ("", new Check_Box_Input("Children", check_parents)), ("Names", new Text_Field_Input(list.map(_.ident).mkString(","))), ("", new Check_Box_Input(reverse_caption, reverse))) case Mutator.Edge_Endpoints((source, dest)) => List( ("Source", new Text_Field_Input(source.ident)), ("Destination", new Text_Field_Input(dest.ident))) case Mutator.Add_Node_Expression(regex) => List(("Regex", new Text_Field_Input(regex, x => Library.make_regex(x).isDefined))) case Mutator.Add_Transitive_Closure(parents, children) => List( ("", new Check_Box_Input("Parents", parents)), ("", new Check_Box_Input("Children", children))) case _ => Nil } } private trait Input { def string: String def bool: Boolean } private class Text_Field_Input(txt: String, check: String => Boolean = (_: String) => true) extends TextField(txt) with Input { preferredSize = new Dimension(125, 18) private val default_foreground = foreground reactions += { case ValueChanged(_) => foreground = if (check(text)) default_foreground else graphview.error_color } def string = text def bool = true } private class Check_Box_Input(txt: String, s: Boolean) extends CheckBox(txt) with Input { selected = s def string = "" def bool = selected } private object Focus_Traveral_Policy extends FocusTraversalPolicy { private var items = Vector.empty[java.awt.Component] def add(c: java.awt.Component): Unit = { items = items :+ c } def addAll(cs: IterableOnce[java.awt.Component]): Unit = { items = items ++ cs } def clear(): Unit = { items = Vector.empty } def getComponentAfter(root: java.awt.Container, c: java.awt.Component): java.awt.Component = { val i = items.indexOf(c) if (i < 0) getDefaultComponent(root) else items((i + 1) % items.length) } def getComponentBefore(root: java.awt.Container, c: java.awt.Component): java.awt.Component = { val i = items.indexOf(c) if (i < 0) getDefaultComponent(root) else items((i - 1) % items.length) } def getFirstComponent(root: java.awt.Container): java.awt.Component = if (items.nonEmpty) items(0) else null def getDefaultComponent(root: java.awt.Container): java.awt.Component = getFirstComponent(root) def getLastComponent(root: java.awt.Container): java.awt.Component = if (items.nonEmpty) items.last else null } } \ No newline at end of file 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,562 +1,562 @@ /* 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 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 -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), "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, modes = modes, log = log) // prevent spurious garbage on the main protocol channel val orig_out = System.out try { System.setOut(new PrintStream(new OutputStream { def write(n: Int): Unit = {} })) 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, session_dirs: List[Path] = Nil, include_sessions: List[String] = Nil, session_ancestor: Option[String] = None, session_requirements: 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 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 + if (invoke_load) delay_load.invoke() } private val file_watcher = File_Watcher(sync_documents, options.seconds("vscode_load_delay")) 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 def build(no_build: Boolean = false): Build.Results = Build.build(options, selection = Sessions.Selection.session(base_info.session), build_heap = true, no_build = no_build, dirs = session_dirs, infos = base_info.infos) if (!build(no_build = true).ok) { val start_msg = "Build started for Isabelle/" + base_info.session + " ..." 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) 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(session, options, base_info.sessions_structure, Sessions.store(options), - modes = modes, logic = base_info.session).await_startup + modes = modes, logic = base_info.session).await_startup() reply_ok("Welcome to Isabelle/" + base_info.session + " (" + Distribution.version + ")") } 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.isDefined) 2 else 0) } /* 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.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 LSP.Symbols_Request(()) => channel.write(LSP.Symbols()) 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/jEdit/src-base/jedit_lib.scala b/src/Tools/jEdit/src-base/jedit_lib.scala --- a/src/Tools/jEdit/src-base/jedit_lib.scala +++ b/src/Tools/jEdit/src-base/jedit_lib.scala @@ -1,25 +1,25 @@ /* Title: Tools/jEdit/src-base/jedit_lib.scala Author: Makarius Misc library functions for jEdit. */ package isabelle.jedit_base import isabelle._ import org.gjt.sp.jedit.{jEdit, View} object JEdit_Lib { def request_focus_view(alt_view: View = null): Unit = { val view = if (alt_view != null) alt_view else jEdit.getActiveView() if (view != null) { val text_area = view.getTextArea - if (text_area != null) text_area.requestFocus + if (text_area != null) text_area.requestFocus() } } } \ No newline at end of file diff --git a/src/Tools/jEdit/src/active.scala b/src/Tools/jEdit/src/active.scala --- a/src/Tools/jEdit/src/active.scala +++ b/src/Tools/jEdit/src/active.scala @@ -1,110 +1,110 @@ /* Title: Tools/jEdit/src/active.scala Author: Makarius Active areas within the document. */ package isabelle.jedit import isabelle._ import org.gjt.sp.jedit.{ServiceManager, View} object Active { abstract class Handler { def handle( view: View, text: String, elem: XML.Elem, doc_view: Document_View, snapshot: Document.Snapshot): Boolean } def handlers: List[Handler] = ServiceManager.getServiceNames(classOf[Handler]).toList .map(ServiceManager.getService(classOf[Handler], _)) def action(view: View, text: String, elem: XML.Elem): Unit = { GUI_Thread.require {} Document_View.get(view.getTextArea) match { case Some(doc_view) => doc_view.rich_text_area.robust_body(()) { val snapshot = doc_view.model.snapshot() if (!snapshot.is_outdated) { handlers.find(_.handle(view, text, elem, doc_view, snapshot)) } } case None => } } class Misc_Handler extends Active.Handler { override def handle( view: View, text: String, elem: XML.Elem, doc_view: Document_View, snapshot: Document.Snapshot): Boolean = { val text_area = doc_view.text_area val model = doc_view.model val buffer = model.buffer elem match { case XML.Elem(Markup(Markup.BROWSER, _), body) => Isabelle_Thread.fork(name = "browser") { val graph_file = Isabelle_System.tmp_file("graph") File.write(graph_file, XML.content(body)) Isabelle_System.bash("isabelle browser -c " + File.bash_path(graph_file) + " &") } true case XML.Elem(Markup(Markup.THEORY_EXPORTS, props), _) => GUI_Thread.later { val name = Markup.Name.unapply(props) getOrElse "" PIDE.editor.hyperlink_file(true, Isabelle_Export.vfs_prefix + name).follow(view) } true case XML.Elem(Markup(Markup.JEDIT_ACTION, _), body) => GUI_Thread.later { view.getInputHandler.invokeAction(XML.content(body)) } true case XML.Elem(Markup(Markup.SIMP_TRACE_PANEL, props), _) => val link = props match { case Position.Id(id) => PIDE.editor.hyperlink_command(true, snapshot, id) case _ => None } GUI_Thread.later { link.foreach(_.follow(view)) view.getDockableWindowManager.showDockableWindow("isabelle-simplifier-trace") } true case XML.Elem(Markup(Markup.SENDBACK, props), _) => if (buffer.isEditable) { props match { case Position.Id(id) => Isabelle.edit_command(snapshot, text_area, props.contains(Markup.PADDING_COMMAND), id, text) case _ => if (props.contains(Markup.PADDING_LINE)) Isabelle.insert_line_padding(text_area, text) else text_area.setSelectedText(text) } - text_area.requestFocus + text_area.requestFocus() } true case Protocol.Dialog(id, serial, result) => model.session.dialog_result(id, serial, result) true case _ => false } } } } diff --git a/src/Tools/jEdit/src/completion_popup.scala b/src/Tools/jEdit/src/completion_popup.scala --- a/src/Tools/jEdit/src/completion_popup.scala +++ b/src/Tools/jEdit/src/completion_popup.scala @@ -1,714 +1,714 @@ /* Title: Tools/jEdit/src/completion_popup.scala Author: Makarius Completion popup. */ package isabelle.jedit import isabelle._ import java.awt.{Color, Font, Point, BorderLayout, Dimension} import java.awt.event.{KeyEvent, KeyListener, MouseEvent, MouseAdapter, FocusAdapter, FocusEvent} import java.io.{File => JFile} import javax.swing.{JPanel, JComponent, JLayeredPane, SwingUtilities} import javax.swing.border.LineBorder import javax.swing.text.DefaultCaret import scala.swing.{ListView, ScrollPane} import scala.swing.event.MouseClicked import org.gjt.sp.jedit.View import org.gjt.sp.jedit.textarea.{JEditTextArea, TextArea, Selection} import org.gjt.sp.jedit.gui.{HistoryTextField, KeyEventWorkaround} import org.gjt.sp.util.StandardUtilities object Completion_Popup { /** items with HTML rendering **/ private class Item(val item: Completion.Item) { private val html = item.description match { case a :: bs => "" + HTML.output(Symbol.print_newlines(a)) + "" + HTML.output(bs.map(b => " " + Symbol.print_newlines(b)).mkString) + "" case Nil => "" } override def toString: String = html } /** jEdit text area **/ object Text_Area { private val key = new Object def apply(text_area: TextArea): Option[Completion_Popup.Text_Area] = { GUI_Thread.require {} text_area.getClientProperty(key) match { case text_area_completion: Completion_Popup.Text_Area => Some(text_area_completion) case _ => None } } def active_range(text_area: TextArea): Option[Text.Range] = apply(text_area) match { case Some(text_area_completion) => text_area_completion.active_range case None => None } def action(text_area: TextArea, word_only: Boolean): Boolean = apply(text_area) match { case Some(text_area_completion) => if (text_area_completion.active_range.isDefined) text_area_completion.action(word_only = word_only) else text_area_completion.action(immediate = true, explicit = true, word_only = word_only) true case None => false } def exit(text_area: JEditTextArea): Unit = { GUI_Thread.require {} apply(text_area) match { case None => case Some(text_area_completion) => text_area_completion.deactivate() text_area.putClientProperty(key, null) } } def init(text_area: JEditTextArea): Completion_Popup.Text_Area = { exit(text_area) val text_area_completion = new Text_Area(text_area) text_area.putClientProperty(key, text_area_completion) text_area_completion.activate() text_area_completion } def dismissed(text_area: TextArea): Boolean = { GUI_Thread.require {} apply(text_area) match { case Some(text_area_completion) => text_area_completion.dismissed() case None => false } } } class Text_Area private(text_area: JEditTextArea) { // owned by GUI thread private var completion_popup: Option[Completion_Popup] = None def active_range: Option[Text.Range] = completion_popup match { case Some(completion) => completion.active_range case None => None } /* rendering */ def rendering(rendering: JEdit_Rendering, line_range: Text.Range): Option[Text.Info[Color]] = { active_range match { case Some(range) => range.try_restrict(line_range) case None => val caret = text_area.getCaretPosition if (line_range.contains(caret)) { rendering.before_caret_range(caret).try_restrict(line_range) match { case Some(range) if !range.is_singularity => val range0 = Completion.Result.merge(Completion.History.empty, syntax_completion(Completion.History.empty, true, Some(rendering)), rendering.path_completion(caret), Document_Model.bibtex_completion(Completion.History.empty, rendering, caret)) .map(_.range) rendering.semantic_completion(range0, range) match { case None => range0 case Some(Text.Info(_, Completion.No_Completion)) => None case Some(Text.Info(range1, _: Completion.Names)) => Some(range1) } case _ => None } } else None } }.map(range => Text.Info(range, rendering.completion_color)) /* syntax completion */ def syntax_completion( history: Completion.History, explicit: Boolean, opt_rendering: Option[JEdit_Rendering]): Option[Completion.Result] = { val buffer = text_area.getBuffer val unicode = Isabelle_Encoding.is_active(buffer) Isabelle.buffer_syntax(buffer) match { case Some(syntax) => val context = (for { rendering <- opt_rendering if PIDE.options.bool("jedit_completion_context") caret_range = rendering.before_caret_range(text_area.getCaretPosition) context <- rendering.language_context(caret_range) } yield context) getOrElse syntax.language_context val caret = text_area.getCaretPosition val line_range = JEdit_Lib.line_range(buffer, text_area.getCaretLine) val line_start = line_range.start for { line_text <- JEdit_Lib.get_text(buffer, line_range) result <- syntax.complete( history, unicode, explicit, line_start, line_text, caret - line_start, context) } yield result case None => None } } /* completion action: text area */ private def insert(item: Completion.Item): Unit = { GUI_Thread.require {} val buffer = text_area.getBuffer val range = item.range if (buffer.isEditable) { JEdit_Lib.buffer_edit(buffer) { JEdit_Lib.get_text(buffer, range) match { case Some(text) if text == item.original => text_area.getSelectionAtOffset(text_area.getCaretPosition) match { /*rectangular selection as "tall caret"*/ case selection : Selection.Rect if selection.getStart(buffer, text_area.getCaretLine) == range.stop => text_area.moveCaretPosition(range.stop) (0 until Character.codePointCount(item.original, 0, item.original.length)) .foreach(_ => text_area.backspace()) text_area.setSelectedText(selection, item.replacement) text_area.moveCaretPosition(text_area.getCaretPosition + item.move) /*other selections: rectangular, multiple range, ...*/ case selection if selection != null && selection.getStart(buffer, text_area.getCaretLine) == range.start && selection.getEnd(buffer, text_area.getCaretLine) == range.stop => text_area.moveCaretPosition(range.stop + item.move) text_area.getSelection.foreach(text_area.setSelectedText(_, item.replacement)) /*no selection*/ case _ => buffer.remove(range.start, range.length) buffer.insert(range.start, item.replacement) text_area.moveCaretPosition(range.start + item.replacement.length + item.move) Isabelle.indent_input(text_area) } case _ => } } } } def action( immediate: Boolean = false, explicit: Boolean = false, delayed: Boolean = false, word_only: Boolean = false): Boolean = { val view = text_area.getView val layered = view.getLayeredPane val buffer = text_area.getBuffer val painter = text_area.getPainter val history = PIDE.plugin.completion_history.value val unicode = Isabelle_Encoding.is_active(buffer) def open_popup(result: Completion.Result): Unit = { val font = painter.getFont.deriveFont( Font_Info.main_size(PIDE.options.real("jedit_popup_font_scale"))) val range = result.range val loc1 = text_area.offsetToXY(range.start) if (loc1 != null) { val loc2 = SwingUtilities.convertPoint(painter, loc1.x, loc1.y + painter.getLineHeight, layered) val items = result.items.map(new Item(_)) val completion = new Completion_Popup(Some(range), layered, loc2, font, items) { override def complete(item: Completion.Item): Unit = { PIDE.plugin.completion_history.update(item) insert(item) } override def propagate(evt: KeyEvent): Unit = { if (view.getKeyEventInterceptor == null) JEdit_Lib.propagate_key(view, evt) else if (view.getKeyEventInterceptor == inner_key_listener) { try { view.setKeyEventInterceptor(null) JEdit_Lib.propagate_key(view, evt) } finally { if (isDisplayable) view.setKeyEventInterceptor(inner_key_listener) } } if (evt.getID == KeyEvent.KEY_TYPED) input(evt) } override def shutdown(focus: Boolean): Unit = { if (view.getKeyEventInterceptor == inner_key_listener) view.setKeyEventInterceptor(null) - if (focus) text_area.requestFocus + if (focus) text_area.requestFocus() JEdit_Lib.invalidate_range(text_area, range) } } dismissed() completion_popup = Some(completion) view.setKeyEventInterceptor(completion.inner_key_listener) JEdit_Lib.invalidate_range(text_area, range) Pretty_Tooltip.dismissed_all() completion.show_popup(false) } } if (buffer.isEditable) { val caret = text_area.getCaretPosition val opt_rendering = Document_View.get(text_area).map(_.get_rendering()) val result0 = syntax_completion(history, explicit, opt_rendering) val (no_completion, semantic_completion) = { opt_rendering match { case Some(rendering) => rendering.semantic_completion_result(history, unicode, result0.map(_.range), rendering.before_caret_range(caret)) case None => (false, None) } } if (no_completion) false else { val result = { val result1 = if (word_only) None else Completion.Result.merge(history, semantic_completion, result0) opt_rendering match { case None => result1 case Some(rendering) => Completion.Result.merge(history, result1, JEdit_Spell_Checker.completion(rendering, explicit, caret), rendering.path_completion(caret), Document_Model.bibtex_completion(history, rendering, caret)) } } result match { case Some(result) => result.items match { case List(item) if result.unique && item.immediate && immediate => insert(item) true case _ :: _ if !delayed => open_popup(result) false case _ => false } case None => false } } } else false } /* input key events */ def input(evt: KeyEvent): Unit = { GUI_Thread.require {} if (!evt.isConsumed) { val special = JEdit_Lib.special_key(evt) if (PIDE.options.bool("jedit_completion")) { dismissed() if (evt.getKeyChar != '\b') { val immediate = PIDE.options.bool("jedit_completion_immediate") if (PIDE.options.seconds("jedit_completion_delay").is_zero && !special) { input_delay.revoke() action(immediate = immediate) } else { if (!special && action(immediate = immediate, delayed = true)) input_delay.revoke() else input_delay.invoke() } } } val selection = text_area.getSelection() if (!special && (selection == null || selection.isEmpty)) Isabelle.indent_input(text_area) } } private val input_delay = Delay.last(PIDE.options.seconds("jedit_completion_delay"), gui = true) { action() } /* dismiss popup */ def dismissed(): Boolean = { GUI_Thread.require {} completion_popup match { case Some(completion) => completion.hide_popup() completion_popup = None true case None => false } } /* activation */ private val outer_key_listener = JEdit_Lib.key_listener(key_typed = input) private def activate(): Unit = text_area.addKeyListener(outer_key_listener) private def deactivate(): Unit = { dismissed() text_area.removeKeyListener(outer_key_listener) } } /** history text field **/ class History_Text_Field( name: String = null, instant_popups: Boolean = false, enter_adds_to_history: Boolean = true, syntax: Outer_Syntax = Outer_Syntax.empty) extends HistoryTextField(name, instant_popups, enter_adds_to_history) { text_field => // see https://forums.oracle.com/thread/1361677 if (GUI.is_macos_laf) text_field.setCaret(new DefaultCaret) // owned by GUI thread private var completion_popup: Option[Completion_Popup] = None /* dismiss */ private def dismissed(): Boolean = { completion_popup match { case Some(completion) => completion.hide_popup() completion_popup = None true case None => false } } /* insert */ private def insert(item: Completion.Item): Unit = { GUI_Thread.require {} val range = item.range if (text_field.isEditable) { val content = text_field.getText range.try_substring(content) match { case Some(text) if text == item.original => text_field.setText( content.substring(0, range.start) + item.replacement + content.substring(range.stop)) text_field.getCaret.setDot(range.start + item.replacement.length + item.move) case _ => } } } /* completion action: text field */ def action(): Unit = { GUI.layered_pane(text_field) match { case Some(layered) if text_field.isEditable => val history = PIDE.plugin.completion_history.value val caret = text_field.getCaret.getDot val text = text_field.getText val context = syntax.language_context syntax.complete(history, true, false, 0, text, caret, context) match { case Some(result) => val fm = text_field.getFontMetrics(text_field.getFont) val loc = SwingUtilities.convertPoint(text_field, fm.stringWidth(text), fm.getHeight, layered) val items = result.items.map(new Item(_)) val completion = new Completion_Popup(None, layered, loc, text_field.getFont, items) { override def complete(item: Completion.Item): Unit = { PIDE.plugin.completion_history.update(item) insert(item) } override def propagate(evt: KeyEvent): Unit = if (!evt.isConsumed) text_field.processKeyEvent(evt) override def shutdown(focus: Boolean): Unit = - if (focus) text_field.requestFocus + if (focus) text_field.requestFocus() } dismissed() completion_popup = Some(completion) completion.show_popup(true) case None => } case _ => } } /* process key event */ private def process(evt: KeyEvent): Unit = { if (PIDE.options.bool("jedit_completion")) { dismissed() if (evt.getKeyChar != '\b') { val special = JEdit_Lib.special_key(evt) if (PIDE.options.seconds("jedit_completion_delay").is_zero && !special) { process_delay.revoke() action() } else process_delay.invoke() } } } private val process_delay = Delay.last(PIDE.options.seconds("jedit_completion_delay"), gui = true) { action() } override def processKeyEvent(evt0: KeyEvent): Unit = { val evt = KeyEventWorkaround.processKeyEvent(evt0) if (evt != null) { evt.getID match { case KeyEvent.KEY_PRESSED => val key_code = evt.getKeyCode if (key_code == KeyEvent.VK_ESCAPE) { if (dismissed()) evt.consume } case KeyEvent.KEY_TYPED => super.processKeyEvent(evt) process(evt) evt.consume case _ => } if (!evt.isConsumed) super.processKeyEvent(evt) } } } } class Completion_Popup private( opt_range: Option[Text.Range], layered: JLayeredPane, location: Point, font: Font, items: List[Completion_Popup.Item]) extends JPanel(new BorderLayout) { completion => GUI_Thread.require {} require(items.nonEmpty, "no completion items") val multi: Boolean = items.length > 1 /* actions */ def complete(item: Completion.Item): Unit = {} def propagate(evt: KeyEvent): Unit = {} def shutdown(focus: Boolean): Unit = {} /* list view */ private val list_view = new ListView(items) list_view.font = font list_view.selection.intervalMode = ListView.IntervalMode.Single list_view.peer.setFocusTraversalKeysEnabled(false) list_view.peer.setVisibleRowCount(items.length min 8) list_view.peer.setSelectedIndex(0) for (cond <- List(JComponent.WHEN_FOCUSED, JComponent.WHEN_ANCESTOR_OF_FOCUSED_COMPONENT, JComponent.WHEN_IN_FOCUSED_WINDOW)) list_view.peer.setInputMap(cond, null) private def complete_selected(): Boolean = { list_view.selection.items.toList match { case item :: _ => complete(item.item); true case _ => false } } private def move_items(n: Int): Unit = { val i = list_view.peer.getSelectedIndex val j = ((i + n) min (items.length - 1)) max 0 if (i != j) { list_view.peer.setSelectedIndex(j) list_view.peer.ensureIndexIsVisible(j) } } private def move_pages(n: Int): Unit = { val page_size = list_view.peer.getVisibleRowCount - 1 move_items(page_size * n) } /* event handling */ val inner_key_listener: KeyListener = JEdit_Lib.key_listener( key_pressed = (e: KeyEvent) => { if (!e.isConsumed) { e.getKeyCode match { case KeyEvent.VK_ENTER if PIDE.options.bool("jedit_completion_select_enter") => if (complete_selected()) e.consume hide_popup() case KeyEvent.VK_TAB if PIDE.options.bool("jedit_completion_select_tab") => if (complete_selected()) e.consume hide_popup() case KeyEvent.VK_ESCAPE => hide_popup() e.consume case KeyEvent.VK_UP | KeyEvent.VK_KP_UP if multi => move_items(-1) e.consume case KeyEvent.VK_DOWN | KeyEvent.VK_KP_DOWN if multi => move_items(1) e.consume case KeyEvent.VK_PAGE_UP if multi => move_pages(-1) e.consume case KeyEvent.VK_PAGE_DOWN if multi => move_pages(1) e.consume case _ => if (e.isActionKey || e.isAltDown || e.isMetaDown || e.isControlDown) hide_popup() } } propagate(e) }, key_typed = propagate ) list_view.peer.addKeyListener(inner_key_listener) list_view.peer.addMouseListener(new MouseAdapter { override def mouseClicked(e: MouseEvent): Unit = { if (complete_selected()) e.consume hide_popup() } }) list_view.peer.addFocusListener(new FocusAdapter { override def focusLost(e: FocusEvent): Unit = hide_popup() }) /* main content */ override def getFocusTraversalKeysEnabled = false completion.setBorder(new LineBorder(Color.BLACK)) completion.add((new ScrollPane(list_view)).peer.asInstanceOf[JComponent]) /* popup */ def active_range: Option[Text.Range] = if (isDisplayable) opt_range else None private val popup = { val screen = GUI.screen_location(layered, location) val size = { val geometry = JEdit_Lib.window_geometry(completion, completion) val bounds = JEdit_Rendering.popup_bounds val w = geometry.width min (screen.bounds.width * bounds).toInt min layered.getWidth val h = geometry.height min (screen.bounds.height * bounds).toInt min layered.getHeight new Dimension(w, h) } new Popup(layered, completion, screen.relative(layered, size), size) } private def show_popup(focus: Boolean): Unit = { popup.show - if (focus) list_view.requestFocus + if (focus) list_view.requestFocus() } private def hide_popup(): Unit = { shutdown(list_view.peer.isFocusOwner) popup.hide } } diff --git a/src/Tools/jEdit/src/debugger_dockable.scala b/src/Tools/jEdit/src/debugger_dockable.scala --- a/src/Tools/jEdit/src/debugger_dockable.scala +++ b/src/Tools/jEdit/src/debugger_dockable.scala @@ -1,372 +1,372 @@ /* Title: Tools/jEdit/src/debugger_dockable.scala Author: Makarius Dockable window for Isabelle/ML debugger. */ package isabelle.jedit import isabelle._ import isabelle.jedit_base.Dockable import java.awt.{BorderLayout, Dimension} import java.awt.event.{ComponentEvent, ComponentAdapter, KeyEvent, FocusAdapter, FocusEvent, MouseEvent, MouseAdapter} import javax.swing.{JTree, JMenuItem} import javax.swing.tree.{DefaultMutableTreeNode, DefaultTreeModel, TreeSelectionModel} import javax.swing.event.{TreeSelectionEvent, TreeSelectionListener} import scala.collection.immutable.SortedMap import scala.swing.{Button, Label, Component, ScrollPane, SplitPane, Orientation, CheckBox, BorderPanel} import scala.swing.event.ButtonClicked import org.gjt.sp.jedit.{jEdit, View} import org.gjt.sp.jedit.menu.EnhancedMenuItem import org.gjt.sp.jedit.textarea.JEditTextArea object Debugger_Dockable { /* breakpoints */ def get_breakpoint(text_area: JEditTextArea, offset: Text.Offset): Option[(Command, Long)] = { GUI_Thread.require {} Document_View.get(text_area) match { case Some(doc_view) => val rendering = doc_view.get_rendering() val range = JEdit_Lib.point_range(text_area.getBuffer, offset) rendering.breakpoint(range) case None => None } } /* context menu */ def context_menu(text_area: JEditTextArea, offset: Text.Offset): List[JMenuItem] = if (PIDE.session.debugger.is_active() && get_breakpoint(text_area, offset).isDefined) { val context = jEdit.getActionContext() val name = "isabelle.toggle-breakpoint" List(new EnhancedMenuItem(context.getAction(name).getLabel, name, context)) } else Nil } class Debugger_Dockable(view: View, position: String) extends Dockable(view, position) { GUI_Thread.require {} private val debugger = PIDE.session.debugger /* component state -- owned by GUI thread */ private var current_snapshot = Document.Snapshot.init private var current_threads: Debugger.Threads = SortedMap.empty private var current_output: List[XML.Tree] = Nil /* pretty text area */ val pretty_text_area = new Pretty_Text_Area(view) override def detach_operation: Option[() => Unit] = pretty_text_area.detach_operation private def handle_resize(): Unit = { GUI_Thread.require {} pretty_text_area.resize( Font_Info.main(PIDE.options.real("jedit_font_scale") * zoom.factor / 100)) } private def handle_update(): Unit = { GUI_Thread.require {} val new_snapshot = PIDE.editor.current_node_snapshot(view).getOrElse(current_snapshot) val (new_threads, new_output) = debugger.status(tree_selection()) if (new_threads != current_threads) update_tree(new_threads) if (new_output != current_output) pretty_text_area.update(new_snapshot, Command.Results.empty, Pretty.separate(new_output)) current_snapshot = new_snapshot current_threads = new_threads current_output = new_output } /* tree view */ private val root = new DefaultMutableTreeNode("Threads") val tree = new JTree(root) tree.setRowHeight(0) tree.getSelectionModel.setSelectionMode(TreeSelectionModel.SINGLE_TREE_SELECTION) def tree_selection(): Option[Debugger.Context] = tree.getLastSelectedPathComponent match { case node: DefaultMutableTreeNode => node.getUserObject match { case c: Debugger.Context => Some(c) case _ => None } case _ => None } def thread_selection(): Option[String] = tree_selection().map(_.thread_name) private def update_tree(threads: Debugger.Threads): Unit = { val thread_contexts = (for ((a, b) <- threads.iterator) yield Debugger.Context(a, b)).toList val new_tree_selection = tree_selection() match { case Some(c) if thread_contexts.contains(c) => Some(c) case Some(c) if thread_contexts.exists(t => c.thread_name == t.thread_name) => Some(c.reset) case _ => thread_contexts.headOption } tree.clearSelection root.removeAllChildren for (thread <- thread_contexts) { val thread_node = new DefaultMutableTreeNode(thread) for ((debug_state, i) <- thread.debug_states.zipWithIndex) thread_node.add(new DefaultMutableTreeNode(thread.select(i))) root.add(thread_node) } tree.getModel.asInstanceOf[DefaultTreeModel].reload(root) tree.expandRow(0) for (i <- Range.inclusive(tree.getRowCount - 1, 1, -1)) tree.expandRow(i) new_tree_selection match { case Some(c) => val i = (for (t <- thread_contexts.iterator.takeWhile(t => c.thread_name != t.thread_name)) yield t.size).sum tree.addSelectionRow(i + c.index + 1) case None => } tree.revalidate() } def update_vals(): Unit = { tree_selection() match { case Some(c) if c.stack_state.isDefined => debugger.print_vals(c, sml_button.selected, context_field.getText) case Some(c) => debugger.clear_output(c.thread_name) case None => } } tree.addTreeSelectionListener( new TreeSelectionListener { override def valueChanged(e: TreeSelectionEvent): Unit = { update_focus() update_vals() } }) tree.addMouseListener( new MouseAdapter { override def mouseClicked(e: MouseEvent): Unit = { val click = tree.getPathForLocation(e.getX, e.getY) if (click != null && e.getClickCount == 1) update_focus() } }) private val tree_pane = new ScrollPane(Component.wrap(tree)) tree_pane.horizontalScrollBarPolicy = ScrollPane.BarPolicy.Always tree_pane.verticalScrollBarPolicy = ScrollPane.BarPolicy.Always tree_pane.minimumSize = new Dimension(200, 100) /* controls */ private val break_button = new CheckBox("Break") { tooltip = "Break running threads at next possible breakpoint" selected = debugger.is_break() reactions += { case ButtonClicked(_) => debugger.set_break(selected) } } private val continue_button = new Button("Continue") { tooltip = "Continue program on current thread, until next breakpoint" reactions += { case ButtonClicked(_) => thread_selection().map(debugger.continue) } } private val step_button = new Button("Step") { tooltip = "Single-step in depth-first order" reactions += { case ButtonClicked(_) => thread_selection().map(debugger.step) } } private val step_over_button = new Button("Step over") { tooltip = "Single-step within this function" reactions += { case ButtonClicked(_) => thread_selection().map(debugger.step_over) } } private val step_out_button = new Button("Step out") { tooltip = "Single-step outside this function" reactions += { case ButtonClicked(_) => thread_selection().map(debugger.step_out) } } private val context_label = new Label("Context:") { tooltip = "Isabelle/ML context: type theory, Proof.context, Context.generic" } private val context_field = new Completion_Popup.History_Text_Field("isabelle-debugger-context") { override def processKeyEvent(evt: KeyEvent): Unit = { if (evt.getID == KeyEvent.KEY_PRESSED && evt.getKeyCode == KeyEvent.VK_ENTER) eval_expression() super.processKeyEvent(evt) } setColumns(20) setToolTipText(context_label.tooltip) setFont(GUI.imitate_font(getFont, scale = 1.2)) } private val expression_label = new Label("ML:") { tooltip = "Isabelle/ML or Standard ML expression" } private val expression_field = new Completion_Popup.History_Text_Field("isabelle-debugger-expression") { override def processKeyEvent(evt: KeyEvent): Unit = { if (evt.getID == KeyEvent.KEY_PRESSED && evt.getKeyCode == KeyEvent.VK_ENTER) eval_expression() super.processKeyEvent(evt) } { val max = getPreferredSize; max.width = Integer.MAX_VALUE; setMaximumSize(max) } setColumns(40) setToolTipText(expression_label.tooltip) setFont(GUI.imitate_font(getFont, scale = 1.2)) } private val eval_button = new Button("Eval") { tooltip = "Evaluate ML expression within optional context" reactions += { case ButtonClicked(_) => eval_expression() } } private def eval_expression(): Unit = { context_field.addCurrentToHistory() expression_field.addCurrentToHistory() tree_selection() match { case Some(c) if c.debug_index.isDefined => debugger.eval(c, sml_button.selected, context_field.getText, expression_field.getText) case _ => } } private val sml_button = new CheckBox("SML") { tooltip = "Official Standard ML instead of Isabelle/ML" selected = false } private val zoom = new Font_Info.Zoom_Box { def changed = handle_resize() } private val controls = Wrap_Panel( List( break_button, continue_button, step_button, step_over_button, step_out_button, context_label, Component.wrap(context_field), expression_label, Component.wrap(expression_field), eval_button, sml_button, pretty_text_area.search_label, pretty_text_area.search_field, zoom)) add(controls.peer, BorderLayout.NORTH) /* focus */ - override def focusOnDefaultComponent(): Unit = eval_button.requestFocus + override def focusOnDefaultComponent(): Unit = eval_button.requestFocus() addFocusListener(new FocusAdapter { override def focusGained(e: FocusEvent): Unit = update_focus() }) private def update_focus(): Unit = { for (c <- tree_selection()) { debugger.set_focus(c) for { pos <- c.debug_position link <- PIDE.editor.hyperlink_position(false, current_snapshot, pos) } link.follow(view) } JEdit_Lib.jedit_text_areas(view.getBuffer).foreach(_.repaint()) } /* main panel */ val main_panel: SplitPane = new SplitPane(Orientation.Vertical) { oneTouchExpandable = true leftComponent = tree_pane rightComponent = Component.wrap(pretty_text_area) } set_content(main_panel) /* main */ private val main = Session.Consumer[Any](getClass.getName) { case _: Session.Global_Options => GUI_Thread.later { handle_resize() } case Debugger.Update => GUI_Thread.later { break_button.selected = debugger.is_break() handle_update() } } override def init(): Unit = { PIDE.session.global_options += main PIDE.session.debugger_updates += main debugger.init() handle_update() jEdit.propertiesChanged() } override def exit(): Unit = { PIDE.session.global_options -= main PIDE.session.debugger_updates -= main delay_resize.revoke() debugger.exit() jEdit.propertiesChanged() } /* resize */ private val delay_resize = Delay.first(PIDE.options.seconds("editor_update_delay"), gui = true) { handle_resize() } addComponentListener(new ComponentAdapter { override def componentResized(e: ComponentEvent): Unit = delay_resize.invoke() override def componentShown(e: ComponentEvent): Unit = delay_resize.invoke() }) } diff --git a/src/Tools/jEdit/src/document_model.scala b/src/Tools/jEdit/src/document_model.scala --- a/src/Tools/jEdit/src/document_model.scala +++ b/src/Tools/jEdit/src/document_model.scala @@ -1,700 +1,700 @@ /* Title: Tools/jEdit/src/document_model.scala Author: Fabian Immler, TU Munich Author: Makarius Document model connected to jEdit buffer or external file: content of theory node or auxiliary file (blob). */ package isabelle.jedit import isabelle._ import java.io.{File => JFile} import scala.collection.mutable import scala.annotation.tailrec import org.gjt.sp.jedit.View import org.gjt.sp.jedit.Buffer import org.gjt.sp.jedit.buffer.{BufferAdapter, BufferListener, JEditBuffer} object Document_Model { /* document models */ sealed case class State( models: Map[Document.Node.Name, Document_Model] = Map.empty, buffer_models: Map[JEditBuffer, Buffer_Model] = Map.empty, overlays: Document.Overlays = Document.Overlays.empty) { def file_models_iterator: Iterator[(Document.Node.Name, File_Model)] = for { (node_name, model) <- models.iterator if model.isInstanceOf[File_Model] } yield (node_name, model.asInstanceOf[File_Model]) def document_blobs: Document.Blobs = Document.Blobs( (for { (node_name, model) <- models.iterator blob <- model.get_blob } yield (node_name -> blob)).toMap) def open_buffer(session: Session, node_name: Document.Node.Name, buffer: Buffer) : (Buffer_Model, State) = { val old_model = models.get(node_name) match { case Some(file_model: File_Model) => Some(file_model) case Some(buffer_model: Buffer_Model) => Some(buffer_model.exit()) case _ => None } val buffer_model = Buffer_Model(session, node_name, buffer).init(old_model) (buffer_model, copy(models = models + (node_name -> buffer_model), buffer_models = buffer_models + (buffer -> buffer_model))) } def close_buffer(buffer: JEditBuffer): State = { buffer_models.get(buffer) match { case None => this case Some(buffer_model) => val file_model = buffer_model.exit() copy(models = models + (file_model.node_name -> file_model), buffer_models = buffer_models - buffer) } } def provide_file(session: Session, node_name: Document.Node.Name, text: String): State = if (models.isDefinedAt(node_name)) this else { val edit = Text.Edit.insert(0, text) val model = File_Model.init(session, node_name, text, pending_edits = List(edit)) copy(models = models + (node_name -> model)) } } private val state = Synchronized(State()) // owned by GUI thread def get_models(): Map[Document.Node.Name, Document_Model] = state.value.models def get(name: Document.Node.Name): Option[Document_Model] = get_models().get(name) def get(buffer: JEditBuffer): Option[Buffer_Model] = state.value.buffer_models.get(buffer) def document_blobs(): Document.Blobs = state.value.document_blobs /* bibtex */ def bibtex_entries_iterator(): Iterator[Text.Info[(String, Document_Model)]] = Bibtex.entries_iterator(state.value.models) def bibtex_completion(history: Completion.History, rendering: Rendering, caret: Text.Offset) : Option[Completion.Result] = Bibtex.completion(history, rendering, caret, state.value.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(st => st.copy(overlays = st.overlays.insert(command, fn, args))) def remove_overlay(command: Command, fn: String, args: List[String]): Unit = state.change(st => st.copy(overlays = st.overlays.remove(command, fn, args))) /* sync external files */ def sync_files(changed_files: Set[JFile]): Boolean = { state.change_result(st => { val changed_models = (for { (node_name, model) <- st.file_models_iterator file <- model.file if changed_files(file) text <- PIDE.resources.read_file_content(node_name) if model.content.text != text } yield { val content = Document_Model.File_Content(text) val edits = Text.Edit.replace(0, model.content.text, text) (node_name, model.copy(content = content, pending_edits = model.pending_edits ::: edits)) }).toList if (changed_models.isEmpty) (false, st) else (true, st.copy(models = changed_models.foldLeft(st.models)(_ + _))) }) } /* syntax */ def syntax_changed(names: List[Document.Node.Name]): Unit = { GUI_Thread.require {} val models = state.value.models for (name <- names.iterator; model <- models.get(name)) { model match { case buffer_model: Buffer_Model => buffer_model.syntax_changed() case _ => } } } /* init and exit */ def init(session: Session, node_name: Document.Node.Name, buffer: Buffer): Buffer_Model = { GUI_Thread.require {} state.change_result(st => st.buffer_models.get(buffer) match { case Some(buffer_model) if buffer_model.node_name == node_name => buffer_model.init_token_marker() (buffer_model, st) case _ => val res = st.close_buffer(buffer).open_buffer(session, node_name, buffer) buffer.propertiesChanged() res }) } def exit(buffer: Buffer): Unit = { GUI_Thread.require {} state.change(st => if (st.buffer_models.isDefinedAt(buffer)) { val res = st.close_buffer(buffer) buffer.propertiesChanged() res } else st) } def provide_files(session: Session, files: List[(Document.Node.Name, String)]): Unit = { GUI_Thread.require {} state.change(st => files.foldLeft(st) { case (st1, (node_name, text)) => st1.provide_file(session, node_name, text) }) } /* required nodes */ def required_nodes(): Set[Document.Node.Name] = (for { (node_name, model) <- state.value.models.iterator if model.node_required } yield node_name).toSet def node_required( name: Document.Node.Name, toggle: Boolean = false, set: Boolean = false): Unit = { GUI_Thread.require {} val changed = state.change_result(st => st.models.get(name) match { case None => (false, st) case Some(model) => val required = if (toggle) !model.node_required else set model match { case model1: File_Model if required != model1.node_required => (true, st.copy(models = st.models + (name -> model1.copy(node_required = required)))) case model1: Buffer_Model if required != model1.node_required => model1.set_node_required(required); (true, st) case _ => (false, st) } }) if (changed) { PIDE.plugin.options_changed() PIDE.editor.flush() } } def view_node_required(view: View, toggle: Boolean = false, set: Boolean = false): Unit = Document_Model.get(view.getBuffer).foreach(model => node_required(model.node_name, toggle = toggle, set = set)) /* flushed edits */ def flush_edits(hidden: Boolean, purge: Boolean): (Document.Blobs, List[Document.Edit_Text]) = { GUI_Thread.require {} state.change_result(st => { val doc_blobs = st.document_blobs val buffer_edits = (for { (_, model) <- st.buffer_models.iterator edit <- model.flush_edits(doc_blobs, hidden).iterator } yield edit).toList val file_edits = (for { (node_name, model) <- st.file_models_iterator (edits, model1) <- model.flush_edits(doc_blobs, hidden) } yield (edits, node_name -> model1)).toList val model_edits = buffer_edits ::: file_edits.flatMap(_._1) val purge_edits = if (purge) { val purged = (for ((node_name, model) <- st.file_models_iterator) yield (node_name -> model.purge_edits(doc_blobs))).toList val imports = { val open_nodes = (for ((_, model) <- st.buffer_models.iterator) yield model.node_name).toList val touched_nodes = model_edits.map(_._1) val pending_nodes = for ((node_name, None) <- purged) yield node_name (open_nodes ::: touched_nodes ::: pending_nodes).map((_, Position.none)) } val retain = PIDE.resources.dependencies(imports).theories.toSet for ((node_name, Some(edits)) <- purged if !retain(node_name); edit <- edits) yield edit } else Nil val st1 = st.copy(models = st.models ++ file_edits.map(_._2) -- purge_edits.map(_._1)) PIDE.plugin.file_watcher.purge( (for { (_, model) <- st1.file_models_iterator file <- model.file } yield file.getParentFile).toSet) ((doc_blobs, model_edits ::: purge_edits), st1) }) } /* file content */ sealed case class File_Content(text: String) { lazy val bytes: Bytes = Bytes(Symbol.encode(text)) lazy val chunk: Symbol.Text_Chunk = Symbol.Text_Chunk(text) lazy val bibtex_entries: List[Text.Info[String]] = try { Bibtex.entries(text) } catch { case ERROR(_) => Nil } } /* HTTP preview */ private val plain_text_prefix = "plain_text=" def open_preview(view: View, plain_text: Boolean): Unit = { Document_Model.get(view.getBuffer) match { case Some(model) => val name = model.node_name val url = PIDE.plugin.http_server.url + PIDE.plugin.http_root + "/preview?" + (if (plain_text) plain_text_prefix else "") + Url.encode(name.node) PIDE.editor.hyperlink_url(url).follow(view) case _ => } } def http_handlers(http_root: String): List[HTTP.Handler] = { val fonts_root = http_root + "/fonts" val preview_root = http_root + "/preview" val html = HTTP.get(preview_root, arg => for { query <- Library.try_unprefix(preview_root + "?", arg.uri.toString).map(Url.decode) name = Library.perhaps_unprefix(plain_text_prefix, query) model <- get(PIDE.resources.node_name(name)) } yield { val snapshot = model.await_stable_snapshot() val html_context = Presentation.html_context(fonts_url = HTML.fonts_dir(fonts_root)) val document = Presentation.html_document( PIDE.resources, snapshot, html_context, Presentation.elements2, plain_text = query.startsWith(plain_text_prefix)) HTTP.Response.html(document.content) }) List(HTTP.fonts(fonts_root), html) } } sealed abstract class Document_Model extends Document.Model { /* perspective */ def document_view_ranges(snapshot: Document.Snapshot): List[Text.Range] = Nil def node_perspective( doc_blobs: Document.Blobs, hidden: Boolean): (Boolean, Document.Node.Perspective_Text) = { GUI_Thread.require {} if (Isabelle.continuous_checking && is_theory) { val snapshot = this.snapshot() val reparse = snapshot.node.load_commands_changed(doc_blobs) val perspective = if (hidden) Text.Perspective.empty else { val view_ranges = document_view_ranges(snapshot) val load_ranges = snapshot.commands_loading_ranges(PIDE.editor.visible_node) Text.Perspective(view_ranges ::: load_ranges) } val overlays = PIDE.editor.node_overlays(node_name) (reparse, Document.Node.Perspective(node_required, perspective, overlays)) } else (false, Document.Node.no_perspective_text) } /* snapshot */ @tailrec final def await_stable_snapshot(): Document.Snapshot = { val snapshot = this.snapshot() if (snapshot.is_outdated) { PIDE.options.seconds("editor_output_delay").sleep await_stable_snapshot() } else snapshot } } object File_Model { def empty(session: Session): File_Model = File_Model(session, Document.Node.Name.empty, None, Document_Model.File_Content(""), false, Document.Node.no_perspective_text, Nil) def init(session: Session, node_name: Document.Node.Name, text: String, node_required: Boolean = false, last_perspective: Document.Node.Perspective_Text = Document.Node.no_perspective_text, pending_edits: List[Text.Edit] = Nil): File_Model = { val file = JEdit_Lib.check_file(node_name.node) file.foreach(PIDE.plugin.file_watcher.register_parent(_)) val content = Document_Model.File_Content(text) val node_required1 = node_required || File_Format.registry.is_theory(node_name) File_Model(session, node_name, file, content, node_required1, last_perspective, pending_edits) } } case class File_Model( session: Session, node_name: Document.Node.Name, file: Option[JFile], content: Document_Model.File_Content, node_required: Boolean, last_perspective: Document.Node.Perspective_Text, pending_edits: List[Text.Edit]) extends Document_Model { /* text */ def get_text(range: Text.Range): Option[String] = range.try_substring(content.text) /* header */ def node_header: Document.Node.Header = PIDE.resources.special_header(node_name) getOrElse PIDE.resources.check_thy(node_name, Scan.char_reader(content.text), strict = false) /* content */ def node_position(offset: Text.Offset): Line.Node_Position = Line.Node_Position(node_name.node, Line.Position.zero.advance(content.text.substring(0, offset))) def get_blob: Option[Document.Blob] = if (is_theory) None else Some(Document.Blob(content.bytes, content.text, content.chunk, pending_edits.nonEmpty)) def bibtex_entries: List[Text.Info[String]] = if (Bibtex.is_bibtex(node_name.node)) content.bibtex_entries else Nil /* edits */ def update_text(text: String): Option[File_Model] = Text.Edit.replace(0, content.text, text) match { case Nil => None case edits => val content1 = Document_Model.File_Content(text) val pending_edits1 = pending_edits ::: edits Some(copy(content = content1, pending_edits = pending_edits1)) } def flush_edits(doc_blobs: Document.Blobs, hidden: Boolean) : Option[(List[Document.Edit_Text], File_Model)] = { val (reparse, perspective) = node_perspective(doc_blobs, hidden) if (reparse || pending_edits.nonEmpty || last_perspective != perspective) { val edits = node_edits(node_header, pending_edits, perspective) Some((edits, copy(last_perspective = perspective, pending_edits = Nil))) } else None } def purge_edits(doc_blobs: Document.Blobs): Option[List[Document.Edit_Text]] = if (pending_edits.nonEmpty || !File_Format.registry.is_theory(node_name) && (node_required || !Document.Node.is_no_perspective_text(last_perspective))) None else { val text_edits = List(Text.Edit.remove(0, content.text)) Some(node_edits(Document.Node.no_header, text_edits, Document.Node.no_perspective_text)) } /* snapshot */ def is_stable: Boolean = pending_edits.isEmpty def snapshot(): Document.Snapshot = session.snapshot(node_name, pending_edits) } case class Buffer_Model(session: Session, node_name: Document.Node.Name, buffer: Buffer) extends Document_Model { /* text */ def get_text(range: Text.Range): Option[String] = JEdit_Lib.get_text(buffer, range) /* header */ def node_header(): Document.Node.Header = { GUI_Thread.require {} PIDE.resources.special_header(node_name) getOrElse JEdit_Lib.buffer_lock(buffer) { PIDE.resources.check_thy(node_name, JEdit_Lib.buffer_reader(buffer), strict = false) } } /* perspective */ // owned by GUI thread private var _node_required = false def node_required: Boolean = _node_required def set_node_required(b: Boolean): Unit = GUI_Thread.require { _node_required = b } def document_view_iterator: Iterator[Document_View] = for { text_area <- JEdit_Lib.jedit_text_areas(buffer) doc_view <- Document_View.get(text_area) } yield doc_view override def document_view_ranges(snapshot: Document.Snapshot): List[Text.Range] = { GUI_Thread.require {} (for { doc_view <- document_view_iterator range <- doc_view.perspective(snapshot).ranges.iterator } yield range).toList } /* blob */ // owned by GUI thread private var _blob: Option[(Bytes, String, Symbol.Text_Chunk)] = None private def reset_blob(): Unit = GUI_Thread.require { _blob = None } def get_blob: Option[Document.Blob] = GUI_Thread.require { if (is_theory) None else { val (bytes, text, chunk) = _blob match { case Some(x) => x case None => val bytes = PIDE.resources.make_file_content(buffer) val text = buffer.getText(0, buffer.getLength) val chunk = Symbol.Text_Chunk(text) val x = (bytes, text, chunk) _blob = Some(x) x } val changed = pending_edits.nonEmpty Some(Document.Blob(bytes, text, chunk, changed)) } } /* bibtex entries */ // owned by GUI thread private var _bibtex_entries: Option[List[Text.Info[String]]] = None private def reset_bibtex_entries(): Unit = GUI_Thread.require { _bibtex_entries = None } def bibtex_entries: List[Text.Info[String]] = GUI_Thread.require { if (Bibtex.is_bibtex(node_name.node)) { _bibtex_entries match { case Some(entries) => entries case None => val text = JEdit_Lib.buffer_text(buffer) val entries = try { Bibtex.entries(text) } catch { case ERROR(msg) => Output.warning(msg); Nil } _bibtex_entries = Some(entries) entries } } else Nil } /* pending edits */ private object pending_edits { private val pending = new mutable.ListBuffer[Text.Edit] private var last_perspective = Document.Node.no_perspective_text def nonEmpty: Boolean = synchronized { pending.nonEmpty } def get_edits: List[Text.Edit] = synchronized { pending.toList } def get_last_perspective: Document.Node.Perspective_Text = synchronized { last_perspective } def set_last_perspective(perspective: Document.Node.Perspective_Text): Unit = synchronized { last_perspective = perspective } def flush_edits(doc_blobs: Document.Blobs, hidden: Boolean): List[Document.Edit_Text] = synchronized { GUI_Thread.require {} val edits = get_edits val (reparse, perspective) = node_perspective(doc_blobs, hidden) if (reparse || edits.nonEmpty || last_perspective != perspective) { - pending.clear + pending.clear() last_perspective = perspective node_edits(node_header(), edits, perspective) } else Nil } def edit(edits: List[Text.Edit]): Unit = synchronized { GUI_Thread.require {} reset_blob() reset_bibtex_entries() for (doc_view <- document_view_iterator) doc_view.rich_text_area.active_reset() pending ++= edits PIDE.editor.invoke() } } def is_stable: Boolean = !pending_edits.nonEmpty def snapshot(): Document.Snapshot = session.snapshot(node_name, pending_edits.get_edits) def flush_edits(doc_blobs: Document.Blobs, hidden: Boolean): List[Document.Edit_Text] = pending_edits.flush_edits(doc_blobs, hidden) /* buffer listener */ private val buffer_listener: BufferListener = new BufferAdapter { override def contentInserted(buffer: JEditBuffer, start_line: Int, offset: Int, num_lines: Int, length: Int): Unit = { pending_edits.edit(List(Text.Edit.insert(offset, buffer.getText(offset, length)))) } override def preContentRemoved(buffer: JEditBuffer, start_line: Int, offset: Int, num_lines: Int, removed_length: Int): Unit = { pending_edits.edit(List(Text.Edit.remove(offset, buffer.getText(offset, removed_length)))) } } /* syntax */ def syntax_changed(): Unit = { JEdit_Lib.buffer_line_manager(buffer).setFirstInvalidLineContext(0) for (text_area <- JEdit_Lib.jedit_text_areas(buffer)) Untyped.method(Class.forName("org.gjt.sp.jedit.textarea.TextArea"), "foldStructureChanged"). invoke(text_area) buffer.invalidateCachedFoldLevels() } def init_token_marker(): Unit = { Isabelle.buffer_token_marker(buffer) match { case Some(marker) if marker != buffer.getTokenMarker => buffer.setTokenMarker(marker) syntax_changed() case _ => } } /* init */ def init(old_model: Option[File_Model]): Buffer_Model = { GUI_Thread.require {} old_model match { case None => pending_edits.edit(List(Text.Edit.insert(0, JEdit_Lib.buffer_text(buffer)))) case Some(file_model) => set_node_required(file_model.node_required) pending_edits.set_last_perspective(file_model.last_perspective) pending_edits.edit( file_model.pending_edits ::: Text.Edit.replace(0, file_model.content.text, JEdit_Lib.buffer_text(buffer))) } buffer.addBufferListener(buffer_listener) init_token_marker() this } /* exit */ def exit(): File_Model = { GUI_Thread.require {} buffer.removeBufferListener(buffer_listener) init_token_marker() File_Model.init(session, node_name, JEdit_Lib.buffer_text(buffer), node_required, pending_edits.get_last_perspective, pending_edits.get_edits) } } diff --git a/src/Tools/jEdit/src/isabelle.scala b/src/Tools/jEdit/src/isabelle.scala --- a/src/Tools/jEdit/src/isabelle.scala +++ b/src/Tools/jEdit/src/isabelle.scala @@ -1,620 +1,620 @@ /* Title: Tools/jEdit/src/isabelle.scala Author: Makarius Global configuration and convenience operations for Isabelle/jEdit. */ package isabelle.jedit import isabelle._ import java.awt.{Point, Frame, Rectangle} import scala.swing.CheckBox import scala.swing.event.ButtonClicked import org.gjt.sp.jedit.{jEdit, View, Buffer, EditBus} import org.gjt.sp.jedit.msg.ViewUpdate import org.gjt.sp.jedit.buffer.JEditBuffer import org.gjt.sp.jedit.textarea.{JEditTextArea, TextArea, StructureMatcher, Selection} import org.gjt.sp.jedit.syntax.TokenMarker import org.gjt.sp.jedit.indent.IndentRule import org.gjt.sp.jedit.gui.{DockableWindowManager, CompleteWord} import org.jedit.options.CombinedOptions object Isabelle { /* editor modes */ val modes = List( "isabelle", // theory source "isabelle-ml", // ML source "isabelle-news", // NEWS "isabelle-options", // etc/options "isabelle-output", // pretty text area output "isabelle-root", // session ROOT "sml") // Standard ML (not Isabelle/ML) private val ml_syntax: Outer_Syntax = Outer_Syntax.empty.no_tokens. set_language_context(Completion.Language_Context.ML_outer) private val sml_syntax: Outer_Syntax = Outer_Syntax.empty.no_tokens. set_language_context(Completion.Language_Context.SML_outer) private val news_syntax: Outer_Syntax = Outer_Syntax.empty.no_tokens def mode_syntax(mode: String): Option[Outer_Syntax] = mode match { case "isabelle" => Some(PIDE.resources.session_base.overall_syntax) case "isabelle-options" => Some(Options.options_syntax) case "isabelle-root" => Some(Sessions.root_syntax) case "isabelle-ml" => Some(ml_syntax) case "isabelle-news" => Some(news_syntax) case "isabelle-output" => None case "sml" => Some(sml_syntax) case _ => None } def buffer_syntax(buffer: JEditBuffer): Option[Outer_Syntax] = if (buffer == null) None else (JEdit_Lib.buffer_mode(buffer), Document_Model.get(buffer)) match { case ("isabelle", Some(model)) => Some(PIDE.session.recent_syntax(model.node_name)) case (mode, _) => mode_syntax(mode) } /* token markers */ private val mode_markers: Map[String, TokenMarker] = Map(modes.map(mode => (mode, new Token_Markup.Marker(mode, None))): _*) + ("bibtex" -> new JEdit_Bibtex.Token_Marker) def mode_token_marker(mode: String): Option[TokenMarker] = mode_markers.get(mode) def buffer_token_marker(buffer: Buffer): Option[TokenMarker] = { val mode = JEdit_Lib.buffer_mode(buffer) if (mode == "isabelle") Some(new Token_Markup.Marker(mode, Some(buffer))) else mode_token_marker(mode) } /* text structure */ def indent_rule(mode: String): Option[IndentRule] = mode match { case "isabelle" | "isabelle-options" | "isabelle-root" => Some(Text_Structure.Indent_Rule) case _ => None } def structure_matchers(mode: String): List[StructureMatcher] = if (mode == "isabelle") List(Text_Structure.Matcher) else Nil /* dockable windows */ private def wm(view: View): DockableWindowManager = view.getDockableWindowManager def debugger_dockable(view: View): Option[Debugger_Dockable] = wm(view).getDockableWindow("isabelle-debugger") match { case dockable: Debugger_Dockable => Some(dockable) case _ => None } def documentation_dockable(view: View): Option[Documentation_Dockable] = wm(view).getDockableWindow("isabelle-documentation") match { case dockable: Documentation_Dockable => Some(dockable) case _ => None } def monitor_dockable(view: View): Option[Monitor_Dockable] = wm(view).getDockableWindow("isabelle-monitor") match { case dockable: Monitor_Dockable => Some(dockable) case _ => None } def output_dockable(view: View): Option[Output_Dockable] = wm(view).getDockableWindow("isabelle-output") match { case dockable: Output_Dockable => Some(dockable) case _ => None } def protocol_dockable(view: View): Option[Protocol_Dockable] = wm(view).getDockableWindow("isabelle-protocol") match { case dockable: Protocol_Dockable => Some(dockable) case _ => None } def query_dockable(view: View): Option[Query_Dockable] = wm(view).getDockableWindow("isabelle-query") match { case dockable: Query_Dockable => Some(dockable) case _ => None } def raw_output_dockable(view: View): Option[Raw_Output_Dockable] = wm(view).getDockableWindow("isabelle-raw-output") match { case dockable: Raw_Output_Dockable => Some(dockable) case _ => None } def simplifier_trace_dockable(view: View): Option[Simplifier_Trace_Dockable] = wm(view).getDockableWindow("isabelle-simplifier-trace") match { case dockable: Simplifier_Trace_Dockable => Some(dockable) case _ => None } def sledgehammer_dockable(view: View): Option[Sledgehammer_Dockable] = wm(view).getDockableWindow("isabelle-sledgehammer") match { case dockable: Sledgehammer_Dockable => Some(dockable) case _ => None } def state_dockable(view: View): Option[State_Dockable] = wm(view).getDockableWindow("isabelle-state") match { case dockable: State_Dockable => Some(dockable) case _ => None } def symbols_dockable(view: View): Option[Symbols_Dockable] = wm(view).getDockableWindow("isabelle-symbols") match { case dockable: Symbols_Dockable => Some(dockable) case _ => None } def syslog_dockable(view: View): Option[Syslog_Dockable] = wm(view).getDockableWindow("isabelle-syslog") match { case dockable: Syslog_Dockable => Some(dockable) case _ => None } def theories_dockable(view: View): Option[Theories_Dockable] = wm(view).getDockableWindow("isabelle-theories") match { case dockable: Theories_Dockable => Some(dockable) case _ => None } def timing_dockable(view: View): Option[Timing_Dockable] = wm(view).getDockableWindow("isabelle-timing") match { case dockable: Timing_Dockable => Some(dockable) case _ => None } /* continuous checking */ private val CONTINUOUS_CHECKING = "editor_continuous_checking" def continuous_checking: Boolean = PIDE.options.bool(CONTINUOUS_CHECKING) def continuous_checking_=(b: Boolean): Unit = GUI_Thread.require { if (continuous_checking != b) { PIDE.options.bool(CONTINUOUS_CHECKING) = b PIDE.session.update_options(PIDE.options.value) } } def set_continuous_checking(): Unit = { continuous_checking = true } def reset_continuous_checking(): Unit = { continuous_checking = false } def toggle_continuous_checking(): Unit = { continuous_checking = !continuous_checking } class Continuous_Checking extends CheckBox("Continuous checking") { tooltip = "Continuous checking of proof document (visible and required parts)" reactions += { case ButtonClicked(_) => continuous_checking = selected } def load(): Unit = { selected = continuous_checking } load() } /* update state */ def update_state(view: View): Unit = state_dockable(view).foreach(_.update_request()) /* required document nodes */ def set_node_required(view: View): Unit = Document_Model.view_node_required(view, set = true) def reset_node_required(view: View): Unit = Document_Model.view_node_required(view, set = false) def toggle_node_required(view: View): Unit = Document_Model.view_node_required(view, toggle = true) /* full screen */ // see toggleFullScreen() method in jEdit/org/gjt/sp/jedit/View.java def toggle_full_screen(view: View): Unit = { if (PIDE.options.bool("jedit_toggle_full_screen") || Untyped.get[Boolean](view, "fullScreenMode")) view.toggleFullScreen() else { Untyped.set[Boolean](view, "fullScreenMode", true) val screen = GUI.screen_size(view) view.dispose() view.updateFullScreenProps() Untyped.set[Rectangle](view, "windowedBounds", view.getBounds) view.setUndecorated(true) view.setBounds(screen.full_screen_bounds) view.validate() view.setVisible(true) view.toFront() view.closeAllMenus() view.getEditPane.getTextArea.requestFocus() EditBus.send(new ViewUpdate(view, ViewUpdate.FULL_SCREEN_TOGGLED)) } } /* font size */ def reset_font_size(): Unit = Font_Info.main_change.reset(PIDE.options.int("jedit_reset_font_size").toFloat) def increase_font_size(): Unit = Font_Info.main_change.step(1) def decrease_font_size(): Unit = Font_Info.main_change.step(-1) /* structured edits */ def indent_enabled(buffer: JEditBuffer, option: String): Boolean = indent_rule(JEdit_Lib.buffer_mode(buffer)).isDefined && buffer.getStringProperty("autoIndent") == "full" && PIDE.options.bool(option) def indent_input(text_area: TextArea): Unit = { val buffer = text_area.getBuffer val line = text_area.getCaretLine val caret = text_area.getCaretPosition if (text_area.isEditable && indent_enabled(buffer, "jedit_indent_input")) { buffer_syntax(buffer) match { case Some(syntax) => val nav = new Text_Structure.Navigator(syntax, buffer, true) nav.iterator(line, 1).nextOption() match { case Some(Text.Info(range, tok)) if range.stop == caret && syntax.keywords.is_indent_command(tok) => buffer.indentLine(line, true) case _ => } case None => } } } def newline(text_area: TextArea): Unit = { if (!text_area.isEditable()) text_area.getToolkit().beep() else { val buffer = text_area.getBuffer val line = text_area.getCaretLine val caret = text_area.getCaretPosition def nl: Unit = text_area.userInput('\n') if (indent_enabled(buffer, "jedit_indent_newline")) { buffer_syntax(buffer) match { case Some(syntax) => val keywords = syntax.keywords val (toks1, toks2) = Text_Structure.split_line_content(buffer, keywords, line, caret) if (toks1.isEmpty) buffer.removeTrailingWhiteSpace(Array(line)) else if (keywords.is_indent_command(toks1.head)) buffer.indentLine(line, true) if (toks2.isEmpty || keywords.is_indent_command(toks2.head)) { text_area.setSelectedText("\n") if (!buffer.indentLine(line + 1, true)) text_area.goToStartOfWhiteSpace(false) } else nl case None => nl } } else nl } } def insert_line_padding(text_area: JEditTextArea, text: String): Unit = { val buffer = text_area.getBuffer JEdit_Lib.buffer_edit(buffer) { val text1 = if (text_area.getSelectionCount == 0) { def pad(range: Text.Range): String = if (JEdit_Lib.get_text(buffer, range) == Some("\n")) "" else "\n" val caret = JEdit_Lib.caret_range(text_area) val before_caret = JEdit_Lib.point_range(buffer, caret.start - 1) pad(before_caret) + text + pad(caret) } else text text_area.setSelectedText(text1) } } def edit_command( snapshot: Document.Snapshot, text_area: TextArea, padding: Boolean, id: Document_ID.Generic, text: String): Unit = { val buffer = text_area.getBuffer if (!snapshot.is_outdated && text != "") { (snapshot.find_command(id), Document_Model.get(buffer)) match { case (Some((node, command)), Some(model)) if command.node_name == model.node_name => node.command_start(command) match { case Some(start) => JEdit_Lib.buffer_edit(buffer) { val range = command.core_range + start JEdit_Lib.buffer_edit(buffer) { if (padding) { text_area.moveCaretPosition(start + range.length) val start_line = text_area.getCaretLine + 1 text_area.setSelectedText("\n" + text) val end_line = text_area.getCaretLine buffer.indentLines(start_line, end_line) } else { buffer.remove(start, range.length) text_area.moveCaretPosition(start) text_area.setSelectedText(text) } } } case None => } case _ => } } } /* formal entities */ def goto_entity(view: View): Unit = { val text_area = view.getTextArea for { doc_view <- Document_View.get(text_area) rendering = doc_view.get_rendering() caret_range = JEdit_Lib.caret_range(text_area) link <- rendering.hyperlink_entity(caret_range) } link.info.follow(view) } def select_entity(text_area: JEditTextArea): Unit = { for (doc_view <- Document_View.get(text_area)) { val rendering = doc_view.get_rendering() val caret_range = JEdit_Lib.caret_range(text_area) val buffer_range = JEdit_Lib.buffer_range(text_area.getBuffer) val active_focus = rendering.caret_focus_ranges(caret_range, buffer_range) if (active_focus.nonEmpty) { text_area.selectNone() for (r <- active_focus) text_area.addToSelection(new Selection.Range(r.start, r.stop)) } } } /* completion */ def complete(view: View, word_only: Boolean): Unit = Completion_Popup.Text_Area.action(view.getTextArea, word_only) /* control styles */ def control_sub(text_area: JEditTextArea): Unit = Syntax_Style.edit_control_style(text_area, Symbol.sub) def control_sup(text_area: JEditTextArea): Unit = Syntax_Style.edit_control_style(text_area, Symbol.sup) def control_bold(text_area: JEditTextArea): Unit = Syntax_Style.edit_control_style(text_area, Symbol.bold) def control_emph(text_area: JEditTextArea): Unit = Syntax_Style.edit_control_style(text_area, Symbol.emph) def control_reset(text_area: JEditTextArea): Unit = Syntax_Style.edit_control_style(text_area, "") /* block styles */ private def enclose_input(text_area: JEditTextArea, s1: String, s2: String): Unit = { s1.foreach(text_area.userInput) s2.foreach(text_area.userInput) s2.foreach(_ => text_area.goToPrevCharacter(false)) } def input_bsub(text_area: JEditTextArea): Unit = enclose_input(text_area, Symbol.bsub_decoded, Symbol.esub_decoded) def input_bsup(text_area: JEditTextArea): Unit = enclose_input(text_area, Symbol.bsup_decoded, Symbol.esup_decoded) /* antiquoted cartouche */ def antiquoted_cartouche(text_area: TextArea): Unit = { val buffer = text_area.getBuffer for { doc_view <- Document_View.get(text_area) rendering = doc_view.get_rendering() caret_range = JEdit_Lib.caret_range(text_area) antiq_range <- rendering.antiquoted(caret_range) antiq_text <- JEdit_Lib.get_text(buffer, antiq_range) body_text <- Antiquote.read_antiq_body(antiq_text) (name, arg) <- Token.read_antiq_arg(Keyword.Keywords.empty, body_text) if Symbol.is_ascii_identifier(name) } { val op_text = Isabelle_Encoding.perhaps_decode(buffer, Symbol.control_prefix + name + Symbol.control_suffix) val arg_text = if (arg.isEmpty) "" else if (Isabelle_Encoding.is_active(buffer)) Symbol.cartouche_decoded(arg.get) else Symbol.cartouche(arg.get) buffer.remove(antiq_range.start, antiq_range.length) text_area.moveCaretPosition(antiq_range.start) text_area.selectNone text_area.setSelectedText(op_text + arg_text) } } /* spell-checker dictionary */ def update_dictionary(text_area: JEditTextArea, include: Boolean, permanent: Boolean): Unit = { for { spell_checker <- PIDE.plugin.spell_checker.get doc_view <- Document_View.get(text_area) rendering = doc_view.get_rendering() range = JEdit_Lib.caret_range(text_area) Text.Info(_, word) <- Spell_Checker.current_word(rendering, range) } { spell_checker.update(word, include, permanent) JEdit_Lib.jedit_views().foreach(_.repaint()) } } def reset_dictionary(): Unit = { for (spell_checker <- PIDE.plugin.spell_checker.get) { spell_checker.reset() JEdit_Lib.jedit_views().foreach(_.repaint()) } } /* debugger */ def toggle_breakpoint(text_area: JEditTextArea): Unit = { GUI_Thread.require {} if (PIDE.session.debugger.is_active()) { Debugger_Dockable.get_breakpoint(text_area, text_area.getCaretPosition) match { case Some((command, breakpoint)) => PIDE.session.debugger.toggle_breakpoint(command, breakpoint) jEdit.propertiesChanged() case None => } } } /* plugin options */ def plugin_options(frame: Frame): Unit = { GUI_Thread.require {} jEdit.setProperty("Plugin Options.last", "isabelle-general") new CombinedOptions(frame, 1) } /* popups */ def dismissed_popups(view: View): Boolean = { var dismissed = false JEdit_Lib.jedit_text_areas(view).foreach(text_area => if (Completion_Popup.Text_Area.dismissed(text_area)) dismissed = true) if (Pretty_Tooltip.dismissed_all()) dismissed = true dismissed } /* tooltips */ def show_tooltip(view: View, control: Boolean): Unit = { GUI_Thread.require {} val text_area = view.getTextArea val painter = text_area.getPainter val caret_range = JEdit_Lib.caret_range(text_area) for { doc_view <- Document_View.get(text_area) rendering = doc_view.get_rendering() tip <- rendering.tooltip(caret_range, control) loc0 <- Option(text_area.offsetToXY(caret_range.start)) } { val loc = new Point(loc0.x, loc0.y + painter.getLineHeight * 3 / 4) val results = rendering.snapshot.command_results(tip.range) Pretty_Tooltip(view, painter, loc, rendering, results, tip) } } /* error navigation */ private def goto_error( view: View, range: Text.Range, avoid_range: Text.Range = Text.Range.offside, which: String = "")(get: List[Text.Markup] => Option[Text.Markup]): Unit = { GUI_Thread.require {} val text_area = view.getTextArea for (doc_view <- Document_View.get(text_area)) { val rendering = doc_view.get_rendering() val errs = rendering.errors(range).filterNot(_.range.overlaps(avoid_range)) get(errs) match { case Some(err) => PIDE.editor.goto_buffer(false, view, view.getBuffer, err.range.start) case None => view.getStatus.setMessageAndClear("No " + which + "error in current document snapshot") } } } def goto_first_error(view: View): Unit = goto_error(view, JEdit_Lib.buffer_range(view.getBuffer))(_.headOption) def goto_last_error(view: View): Unit = goto_error(view, JEdit_Lib.buffer_range(view.getBuffer))(_.lastOption) def goto_prev_error(view: View): Unit = { val caret_range = JEdit_Lib.caret_range(view.getTextArea) val range = Text.Range(0, caret_range.stop) goto_error(view, range, avoid_range = caret_range, which = "previous ")(_.lastOption) } def goto_next_error(view: View): Unit = { val caret_range = JEdit_Lib.caret_range(view.getTextArea) val range = Text.Range(caret_range.start, view.getBuffer.getLength) goto_error(view, range, avoid_range = caret_range, which = "next ")(_.headOption) } /* java monitor */ def java_monitor(view: View): Unit = - Java_Monitor.java_monitor_external(view, look_and_feel = GUI.current_laf()) + Java_Monitor.java_monitor_external(view, look_and_feel = GUI.current_laf) } diff --git a/src/Tools/jEdit/src/isabelle_sidekick.scala b/src/Tools/jEdit/src/isabelle_sidekick.scala --- a/src/Tools/jEdit/src/isabelle_sidekick.scala +++ b/src/Tools/jEdit/src/isabelle_sidekick.scala @@ -1,268 +1,268 @@ /* Title: Tools/jEdit/src/isabelle_sidekick.scala Author: Fabian Immler, TU Munich Author: Makarius SideKick parsers for Isabelle proof documents. */ package isabelle.jedit import isabelle._ import javax.swing.tree.DefaultMutableTreeNode import javax.swing.text.Position import javax.swing.{JLabel, Icon} import org.gjt.sp.jedit.Buffer import sidekick.{SideKickParser, SideKickParsedData, IAsset} object Isabelle_Sidekick { def int_to_pos(offset: Text.Offset): Position = new Position { def getOffset = offset; override def toString: String = offset.toString } def root_data(buffer: Buffer): SideKickParsedData = { val data = new SideKickParsedData(buffer.getName) data.getAsset(data.root).setEnd(int_to_pos(buffer.getLength)) data } class Keyword_Asset(keyword: String, text: String, range: Text.Range) extends IAsset { private val css = GUI.imitate_font_css(GUI.label_font()) protected var _name = text protected var _start = int_to_pos(range.start) protected var _end = int_to_pos(range.stop) override def getIcon: Icon = null override def getShortString: String = { val n = keyword.length val s = _name.indexOf(keyword) match { case i if i >= 0 && n > 0 => HTML.output(_name.substring(0, i)) + "" + HTML.output(keyword) + "" + HTML.output(_name.substring(i + n)) case _ => HTML.output(_name) } "" + s + "" } override def getLongString: String = _name override def getName: String = _name override def setName(name: String): Unit = _name = name override def getStart: Position = _start override def setStart(start: Position): Unit = _start = start override def getEnd: Position = _end override def setEnd(end: Position): Unit = _end = end override def toString: String = _name } class Asset(name: String, range: Text.Range) extends Keyword_Asset("", name, range) def swing_markup_tree(tree: Markup_Tree, parent: DefaultMutableTreeNode, swing_node: Text.Info[List[XML.Elem]] => DefaultMutableTreeNode): Unit = { for ((_, entry) <- tree.branches) { val node = swing_node(Text.Info(entry.range, entry.markup)) swing_markup_tree(entry.subtree, node, swing_node) parent.add(node) } } } class Isabelle_Sidekick(name: String) extends SideKickParser(name) { override def supportsCompletion = false /* parsing */ @volatile protected var stopped = false override def stop() = { stopped = true } def parser(buffer: Buffer, syntax: Outer_Syntax, data: SideKickParsedData): Boolean = false def parse(buffer: Buffer, error_source: errorlist.DefaultErrorSource): SideKickParsedData = { stopped = false // FIXME lock buffer (!??) val data = Isabelle_Sidekick.root_data(buffer) val syntax = Isabelle.buffer_syntax(buffer) val ok = if (syntax.isDefined) { val ok = parser(buffer, syntax.get, data) if (stopped) { data.root.add(new DefaultMutableTreeNode("")); true } else ok } else false if (!ok) data.root.add(new DefaultMutableTreeNode("")) data } } class Isabelle_Sidekick_Structure( name: String, node_name: Buffer => Option[Document.Node.Name], parse: (Outer_Syntax, Document.Node.Name, CharSequence) => List[Document_Structure.Document]) extends Isabelle_Sidekick(name) { override def parser(buffer: Buffer, syntax: Outer_Syntax, data: SideKickParsedData): Boolean = { def make_tree( parent: DefaultMutableTreeNode, offset: Text.Offset, documents: List[Document_Structure.Document]): Unit = { documents.foldLeft(offset) { case (i, document) => document match { case Document_Structure.Block(name, text, body) => val range = Text.Range(i, i + document.length) val node = new DefaultMutableTreeNode( new Isabelle_Sidekick.Keyword_Asset(name, Library.first_line(text), range)) parent.add(node) make_tree(node, i, body) case _ => } i + document.length } } node_name(buffer) match { case Some(name) => make_tree(data.root, 0, parse(syntax, name, JEdit_Lib.buffer_text(buffer))) true case None => false } } } class Isabelle_Sidekick_Default extends Isabelle_Sidekick_Structure("isabelle", PIDE.resources.theory_node_name, Document_Structure.parse_sections) class Isabelle_Sidekick_Context extends Isabelle_Sidekick_Structure("isabelle-context", PIDE.resources.theory_node_name, Document_Structure.parse_blocks) class Isabelle_Sidekick_Options extends Isabelle_Sidekick_Structure("isabelle-options", _ => Some(Document.Node.Name("options")), Document_Structure.parse_sections) class Isabelle_Sidekick_Root extends Isabelle_Sidekick_Structure("isabelle-root", _ => Some(Document.Node.Name("ROOT")), Document_Structure.parse_sections) class Isabelle_Sidekick_ML extends Isabelle_Sidekick_Structure("isabelle-ml", buffer => Some(PIDE.resources.node_name(buffer)), (_, _, text) => Document_Structure.parse_ml_sections(false, text)) class Isabelle_Sidekick_SML extends Isabelle_Sidekick_Structure("isabelle-sml", buffer => Some(PIDE.resources.node_name(buffer)), (_, _, text) => Document_Structure.parse_ml_sections(true, text)) class Isabelle_Sidekick_Markup extends Isabelle_Sidekick("isabelle-markup") { override def parser(buffer: Buffer, syntax: Outer_Syntax, data: SideKickParsedData): Boolean = { val opt_snapshot = Document_Model.get(buffer) match { - case Some(model) if model.is_theory => Some(model.snapshot) + case Some(model) if model.is_theory => Some(model.snapshot()) case _ => None } opt_snapshot match { case Some(snapshot) => for ((command, command_start) <- snapshot.node.command_iterator() if !stopped) { val markup = snapshot.state.command_markup( snapshot.version, command, Command.Markup_Index.markup, command.range, Markup.Elements.full) Isabelle_Sidekick.swing_markup_tree(markup, data.root, (info: Text.Info[List[XML.Elem]]) => { val range = info.range + command_start val content = command.source(info.range).replace('\n', ' ') val info_text = Pretty.formatted(Pretty.fbreaks(info.info), margin = 40.0).mkString new DefaultMutableTreeNode( new Isabelle_Sidekick.Asset(command.toString, range) { override def getShortString: String = content override def getLongString: String = info_text override def toString: String = quote(content) + " " + range.toString }) }) } true case None => false } } } class Isabelle_Sidekick_News extends Isabelle_Sidekick("isabelle-news") { private val Heading1 = """^New in (.*)\w*$""".r private val Heading2 = """^\*\*\*\w*(.*)\w*\*\*\*\w*$""".r private def make_node(s: String, start: Text.Offset, stop: Text.Offset): DefaultMutableTreeNode = new DefaultMutableTreeNode(new Isabelle_Sidekick.Asset(s, Text.Range(start, stop))) override def parser(buffer: Buffer, syntax: Outer_Syntax, data: SideKickParsedData): Boolean = { var offset = 0 var end_offset = 0 var start1: Option[(Int, String, Vector[DefaultMutableTreeNode])] = None var start2: Option[(Int, String)] = None def close1: Unit = start1 match { case Some((start_offset, s, body)) => val node = make_node(s, start_offset, end_offset) body.foreach(node.add(_)) data.root.add(node) start1 = None case None => } def close2: Unit = start2 match { case Some((start_offset, s)) => start1 match { case Some((start_offset1, s1, body)) => val node = make_node(s, start_offset, end_offset) start1 = Some((start_offset1, s1, body :+ node)) case None => } start2 = None case None => } for (line <- split_lines(JEdit_Lib.buffer_text(buffer)) if !stopped) { line match { case Heading1(s) => close2; close1; start1 = Some((offset, s, Vector())) case Heading2(s) => close2; start2 = Some((offset, s)) case _ => } offset += line.length + 1 if (!line.forall(Character.isSpaceChar)) end_offset = offset } if (!stopped) { close2; close1 } true } } diff --git a/src/Tools/jEdit/src/jedit_editor.scala b/src/Tools/jEdit/src/jedit_editor.scala --- a/src/Tools/jEdit/src/jedit_editor.scala +++ b/src/Tools/jEdit/src/jedit_editor.scala @@ -1,322 +1,322 @@ /* Title: Tools/jEdit/src/jedit_editor.scala Author: Makarius PIDE editor operations for jEdit. */ package isabelle.jedit import isabelle._ import java.io.{File => JFile} import org.gjt.sp.jedit.{jEdit, View, Buffer} import org.gjt.sp.jedit.browser.VFSBrowser import org.gjt.sp.jedit.io.{VFSManager, VFSFile} class JEdit_Editor extends Editor[View] { /* session */ override def session: Session = PIDE.session def flush_edits(hidden: Boolean = false, purge: Boolean = false): Unit = GUI_Thread.require { val (doc_blobs, edits) = Document_Model.flush_edits(hidden, purge) session.update(doc_blobs, edits) } override def flush(): Unit = flush_edits() def purge(): Unit = flush_edits(purge = true) private val delay1_flush = Delay.last(PIDE.options.seconds("editor_input_delay"), gui = true) { flush() } private val delay2_flush = Delay.first(PIDE.options.seconds("editor_generated_input_delay"), gui = true) { flush() } def invoke(): Unit = delay1_flush.invoke() def invoke_generated(): Unit = { delay1_flush.invoke(); delay2_flush.invoke() } def shutdown(): Unit = GUI_Thread.require { delay1_flush.revoke() delay2_flush.revoke() Document_Model.flush_edits(hidden = false, purge = false) } def visible_node(name: Document.Node.Name): Boolean = (for { text_area <- JEdit_Lib.jedit_text_areas() doc_view <- Document_View.get(text_area) } yield doc_view.model.node_name).contains(name) /* current situation */ override def current_node(view: View): Option[Document.Node.Name] = GUI_Thread.require { Document_Model.get(view.getBuffer).map(_.node_name) } override def current_node_snapshot(view: View): Option[Document.Snapshot] = GUI_Thread.require { Document_Model.get(view.getBuffer).map(_.snapshot()) } override def node_snapshot(name: Document.Node.Name): Document.Snapshot = { GUI_Thread.require {} Document_Model.get(name) match { - case Some(model) => model.snapshot + case Some(model) => model.snapshot() case None => session.snapshot(name) } } override def current_command(view: View, snapshot: Document.Snapshot): Option[Command] = { GUI_Thread.require {} val text_area = view.getTextArea val buffer = view.getBuffer Document_View.get(text_area) match { case Some(doc_view) if doc_view.model.is_theory => snapshot.current_command(doc_view.model.node_name, text_area.getCaretPosition) case _ => Document_Model.get(buffer) match { case Some(model) if !model.is_theory => snapshot.version.nodes.commands_loading(model.node_name).headOption case _ => None } } } /* overlays */ override def node_overlays(name: Document.Node.Name): Document.Node.Overlays = Document_Model.node_overlays(name) override def insert_overlay(command: Command, fn: String, args: List[String]): Unit = Document_Model.insert_overlay(command, fn, args) override def remove_overlay(command: Command, fn: String, args: List[String]): Unit = Document_Model.remove_overlay(command, fn, args) /* navigation */ def push_position(view: View): Unit = { val navigator = jEdit.getPlugin("ise.plugin.nav.NavigatorPlugin") if (navigator != null) { try { Untyped.method(navigator.getClass, "pushPosition", view.getClass).invoke(null, view) } catch { case _: NoSuchMethodException => } } } def goto_buffer(focus: Boolean, view: View, buffer: Buffer, offset: Text.Offset): Unit = { GUI_Thread.require {} push_position(view) if (focus) view.goToBuffer(buffer) else view.showBuffer(buffer) try { view.getTextArea.moveCaretPosition(offset) } catch { case _: ArrayIndexOutOfBoundsException => case _: IllegalArgumentException => } } def goto_file(focus: Boolean, view: View, name: String): Unit = goto_file(focus, view, Line.Node_Position.offside(name)) def goto_file(focus: Boolean, view: View, pos: Line.Node_Position): Unit = { GUI_Thread.require {} push_position(view) val name = pos.name val line = pos.line val column = pos.column JEdit_Lib.jedit_buffer(name) match { case Some(buffer) => if (focus) view.goToBuffer(buffer) else view.showBuffer(buffer) val text_area = view.getTextArea try { val line_start = text_area.getBuffer.getLineStartOffset(line) text_area.moveCaretPosition(line_start) if (column > 0) text_area.moveCaretPosition(line_start + column) } catch { case _: ArrayIndexOutOfBoundsException => case _: IllegalArgumentException => } case None => val is_dir = try { val vfs = VFSManager.getVFSForPath(name) val vfs_file = vfs._getFile((), name, view) vfs_file != null && vfs_file.getType == VFSFile.DIRECTORY } catch { case ERROR(_) => false } if (is_dir) VFSBrowser.browseDirectory(view, name) else if (!Isabelle_System.open_external_file(name)) { val args = if (line <= 0) Array(name) else if (column <= 0) Array(name, "+line:" + (line + 1)) else Array(name, "+line:" + (line + 1) + "," + (column + 1)) jEdit.openFiles(view, null, args) } } } def goto_doc(view: View, path: Path): Unit = { if (path.is_file) goto_file(true, view, File.platform_path(path)) else { Isabelle_Thread.fork(name = "documentation") { try { Doc.view(path) } catch { case exn: Throwable => GUI_Thread.later { GUI.error_dialog(view, "Documentation error", GUI.scrollable_text(Exn.message(exn))) } } } } } /* hyperlinks */ def hyperlink_doc(name: String): Option[Hyperlink] = Doc.contents().collectFirst({ case doc: Doc.Text_File if doc.name == name => doc.path case doc: Doc.Doc if doc.name == name => doc.path}).map(path => new Hyperlink { override val external: Boolean = !path.is_file def follow(view: View): Unit = goto_doc(view, path) override def toString: String = "doc " + quote(name) }) def hyperlink_url(name: String): Hyperlink = new Hyperlink { override val external = true def follow(view: View): Unit = Isabelle_Thread.fork(name = "hyperlink_url") { try { Isabelle_System.open(Url.escape_name(name)) } catch { case exn: Throwable => GUI_Thread.later { GUI.error_dialog(view, "System error", GUI.scrollable_text(Exn.message(exn))) } } } override def toString: String = "URL " + quote(name) } def hyperlink_file(focus: Boolean, name: String): Hyperlink = hyperlink_file(focus, Line.Node_Position.offside(name)) def hyperlink_file(focus: Boolean, pos: Line.Node_Position): Hyperlink = new Hyperlink { def follow(view: View): Unit = goto_file(focus, view, pos) override def toString: String = "file " + quote(pos.name) } def hyperlink_model(focus: Boolean, model: Document_Model, offset: Text.Offset): Hyperlink = model match { case file_model: File_Model => val pos = try { file_model.node_position(offset) } catch { case ERROR(_) => Line.Node_Position(file_model.node_name.node) } hyperlink_file(focus, pos) case buffer_model: Buffer_Model => new Hyperlink { def follow(view: View): Unit = goto_buffer(focus, view, buffer_model.buffer, offset) override def toString: String = "buffer " + quote(model.node_name.node) } } def hyperlink_source_file(focus: Boolean, source_name: String, line1: Int, offset: Symbol.Offset) : Option[Hyperlink] = { for (name <- PIDE.resources.source_file(source_name)) yield { def hyperlink(pos: Line.Position) = hyperlink_file(focus, Line.Node_Position(name, pos)) if (offset > 0) { PIDE.resources.get_file_content(PIDE.resources.node_name(name)) match { case Some(text) => hyperlink( Symbol.iterator(text).zipWithIndex.takeWhile(p => p._2 < offset - 1).map(_._1). foldLeft(Line.Position.zero)(_.advance(_))) case None => hyperlink(Line.Position((line1 - 1) max 0)) } } else hyperlink(Line.Position((line1 - 1) max 0)) } } 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(hyperlink_file(focus, _)) } def is_hyperlink_position(snapshot: Document.Snapshot, text_offset: Text.Offset, pos: Position.T): Boolean = { pos match { case Position.Item_Id(id, range) if range.start > 0 => snapshot.find_command(id) match { case Some((node, command)) if snapshot.get_node(command.node_name) eq node => node.command_start(command) match { case Some(start) => text_offset == start + command.chunk.decode(range.start) case None => false } case _ => false } case _ => false } } def hyperlink_position(focus: Boolean, snapshot: Document.Snapshot, pos: Position.T) : Option[Hyperlink] = pos match { case Position.Item_File(name, line, range) => hyperlink_source_file(focus, name, line, range.start) case Position.Item_Id(id, range) => hyperlink_command(focus, snapshot, id, range.start) case _ => None } def hyperlink_def_position(focus: Boolean, snapshot: Document.Snapshot, pos: Position.T) : Option[Hyperlink] = pos match { case Position.Item_Def_File(name, line, range) => hyperlink_source_file(focus, name, line, range.start) case Position.Item_Def_Id(id, range) => hyperlink_command(focus, snapshot, id, range.start) case _ => None } /* dispatcher thread */ override def assert_dispatcher[A](body: => A): A = GUI_Thread.assert(body) override def require_dispatcher[A](body: => A): A = GUI_Thread.require(body) override def send_dispatcher(body: => Unit): Unit = GUI_Thread.later(body) override def send_wait_dispatcher(body: => Unit): Unit = GUI_Thread.now(body) } diff --git a/src/Tools/jEdit/src/jedit_lib.scala b/src/Tools/jEdit/src/jedit_lib.scala --- a/src/Tools/jEdit/src/jedit_lib.scala +++ b/src/Tools/jEdit/src/jedit_lib.scala @@ -1,371 +1,371 @@ /* Title: Tools/jEdit/src/jedit_lib.scala Author: Makarius Misc library functions for jEdit. */ package isabelle.jedit import isabelle._ import java.io.{File => JFile} import java.awt.{Component, Container, GraphicsEnvironment, Point, Rectangle, Dimension, Toolkit} import java.awt.event.{InputEvent, KeyEvent, KeyListener} import javax.swing.{Icon, ImageIcon, JWindow, SwingUtilities} import scala.util.parsing.input.CharSequenceReader import scala.jdk.CollectionConverters._ import org.gjt.sp.jedit.{jEdit, Buffer, View, GUIUtilities, Debug, EditPane} import org.gjt.sp.jedit.io.{FileVFS, VFSManager} import org.gjt.sp.jedit.gui.{KeyEventWorkaround, KeyEventTranslator} import org.gjt.sp.jedit.buffer.{JEditBuffer, LineManager} import org.gjt.sp.jedit.textarea.{JEditTextArea, TextArea, TextAreaPainter} object JEdit_Lib { /* jEdit directories */ def directories: List[JFile] = (Option(jEdit.getSettingsDirectory).toList ::: List(jEdit.getJEditHome)).map(new JFile(_)) /* window geometry measurement */ private lazy val dummy_window = new JWindow final case class Window_Geometry(width: Int, height: Int, inner_width: Int, inner_height: Int) { def deco_width: Int = width - inner_width def deco_height: Int = height - inner_height } def window_geometry(outer: Container, inner: Component): Window_Geometry = { GUI_Thread.require {} val old_content = dummy_window.getContentPane dummy_window.setContentPane(outer) dummy_window.pack - dummy_window.revalidate + dummy_window.revalidate() val geometry = Window_Geometry( dummy_window.getWidth, dummy_window.getHeight, inner.getWidth, inner.getHeight) dummy_window.setContentPane(old_content) geometry } /* files */ def is_file(name: String): Boolean = VFSManager.getVFSForPath(name).isInstanceOf[FileVFS] def check_file(name: String): Option[JFile] = if (is_file(name)) Some(new JFile(name)) else None /* buffers */ def buffer_text(buffer: JEditBuffer): String = buffer_lock(buffer) { buffer.getText(0, buffer.getLength) } def buffer_reader(buffer: JEditBuffer): CharSequenceReader = Scan.char_reader(buffer.getSegment(0, buffer.getLength)) def buffer_mode(buffer: JEditBuffer): String = { val mode = buffer.getMode if (mode == null) "" else { val name = mode.getName if (name == null) "" else name } } def buffer_line_manager(buffer: JEditBuffer): LineManager = Untyped.get[LineManager](buffer, "lineMgr") def buffer_name(buffer: Buffer): String = buffer.getSymlinkPath def buffer_file(buffer: Buffer): Option[JFile] = check_file(buffer_name(buffer)) def buffer_undo_in_progress[A](buffer: JEditBuffer, body: => A): A = { val undo_in_progress = buffer.isUndoInProgress def set(b: Boolean): Unit = Untyped.set[Boolean](buffer, "undoInProgress", b) try { set(true); body } finally { set(undo_in_progress) } } /* main jEdit components */ def jedit_buffers(): Iterator[Buffer] = jEdit.getBufferManager().getBuffers().asScala.iterator def jedit_buffer(name: String): Option[Buffer] = jedit_buffers().find(buffer => buffer_name(buffer) == name) def jedit_buffer(name: Document.Node.Name): Option[Buffer] = jedit_buffer(name.node) def jedit_views(): Iterator[View] = jEdit.getViewManager().getViews().asScala.iterator def jedit_view(view: View = null): View = if (view == null) jEdit.getActiveView() else view def jedit_edit_panes(view: View): Iterator[EditPane] = if (view == null) Iterator.empty else view.getEditPanes().iterator.filter(_ != null) def jedit_text_areas(view: View): Iterator[JEditTextArea] = if (view == null) Iterator.empty else view.getEditPanes().iterator.filter(_ != null).map(_.getTextArea).filter(_ != null) def jedit_text_areas(): Iterator[JEditTextArea] = jedit_views().flatMap(jedit_text_areas) def jedit_text_areas(buffer: JEditBuffer): Iterator[JEditTextArea] = jedit_text_areas().filter(_.getBuffer == buffer) def buffer_lock[A](buffer: JEditBuffer)(body: => A): A = { try { buffer.readLock(); body } finally { buffer.readUnlock() } } def buffer_edit[A](buffer: JEditBuffer)(body: => A): A = { try { buffer.beginCompoundEdit(); body } finally { buffer.endCompoundEdit() } } /* get text */ def get_text(buffer: JEditBuffer, range: Text.Range): Option[String] = try { Some(buffer.getText(range.start, range.length)) } catch { case _: ArrayIndexOutOfBoundsException => None } /* point range */ def point_range(buffer: JEditBuffer, offset: Text.Offset): Text.Range = if (offset < 0) Text.Range.offside else buffer_lock(buffer) { def text(i: Text.Offset): Char = buffer.getText(i, 1).charAt(0) try { val c = text(offset) if (Character.isHighSurrogate(c) && Character.isLowSurrogate(text(offset + 1))) Text.Range(offset, offset + 2) else if (Character.isLowSurrogate(c) && Character.isHighSurrogate(text(offset - 1))) Text.Range(offset - 1, offset + 1) else Text.Range(offset, offset + 1) } catch { case _: ArrayIndexOutOfBoundsException => Text.Range(offset, offset + 1) } } /* text ranges */ def buffer_range(buffer: JEditBuffer): Text.Range = Text.Range(0, buffer.getLength) def line_range(buffer: JEditBuffer, line: Int): Text.Range = Text.Range(buffer.getLineStartOffset(line), buffer.getLineEndOffset(line) min buffer.getLength) def caret_range(text_area: TextArea): Text.Range = point_range(text_area.getBuffer, text_area.getCaretPosition) def visible_range(text_area: TextArea): Option[Text.Range] = { val buffer = text_area.getBuffer val n = text_area.getVisibleLines if (n > 0) { val start = text_area.getScreenLineStartOffset(0) val raw_end = text_area.getScreenLineEndOffset(n - 1) val end = if (raw_end >= 0) raw_end min buffer.getLength else buffer.getLength Some(Text.Range(start, end)) } else None } def invalidate_range(text_area: TextArea, range: Text.Range): Unit = { val buffer = text_area.getBuffer buffer_range(buffer).try_restrict(range) match { case Some(range1) if !range1.is_singularity => try { text_area.invalidateLineRange( buffer.getLineOfOffset(range1.start), buffer.getLineOfOffset(range1.stop)) } catch { case _: ArrayIndexOutOfBoundsException => } case _ => } } def invalidate(text_area: TextArea): Unit = { val visible_lines = text_area.getVisibleLines if (visible_lines > 0) text_area.invalidateScreenLineRange(0, visible_lines) } /* graphics range */ case class Gfx_Range(x: Int, y: Int, length: Int) // NB: jEdit always normalizes \r\n and \r to \n // NB: last line lacks \n def gfx_range(text_area: TextArea, range: Text.Range): Option[Gfx_Range] = { val metric = pretty_metric(text_area.getPainter) val char_width = (metric.unit * metric.average).round.toInt val buffer = text_area.getBuffer val end = buffer.getLength val stop = range.stop val (p, q, r) = try { val p = text_area.offsetToXY(range.start) val (q, r) = if (get_text(buffer, Text.Range(stop - 1, stop)) == Some("\n")) (text_area.offsetToXY(stop - 1), char_width) else if (stop >= end) (text_area.offsetToXY(end), char_width * (stop - end)) else (text_area.offsetToXY(stop), 0) (p, q, r) } catch { case _: ArrayIndexOutOfBoundsException => (null, null, 0) } if (p != null && q != null && p.x < q.x + r && p.y == q.y) Some(Gfx_Range(p.x, p.y, q.x + r - p.x)) else None } /* pixel range */ def pixel_range(text_area: TextArea, x: Int, y: Int): Option[Text.Range] = { // coordinates wrt. inner painter component val painter = text_area.getPainter if (0 <= x && x < painter.getWidth && 0 <= y && y < painter.getHeight) { val offset = text_area.xyToOffset(x, y, false) if (offset >= 0) { val range = point_range(text_area.getBuffer, offset) gfx_range(text_area, range) match { case Some(g) if g.x <= x && x < g.x + g.length => Some(range) case _ => None } } else None } else None } /* pretty text metric */ abstract class Pretty_Metric extends Pretty.Metric { def average: Double } def pretty_metric(painter: TextAreaPainter): Pretty_Metric = new Pretty_Metric { def string_width(s: String): Double = painter.getFont.getStringBounds(s, painter.getFontRenderContext).getWidth val unit: Double = string_width(Symbol.space) max 1.0 val average: Double = string_width("mix") / (3 * unit) def apply(s: String): Double = if (s == "\n") 1.0 else string_width(s) / unit } /* icons */ def load_icon(name: String): Icon = { val name1 = if (name.startsWith("idea-icons/")) { val file = Path.explode("$JEDIT_HOME/dist/jars/idea-icons.jar").file.toURI.toASCIIString "jar:" + file + "!/" + name } else name val icon = GUIUtilities.loadIcon(name1) if (icon.getIconWidth < 0 || icon.getIconHeight < 0) error("Bad icon: " + name) else icon } def load_image_icon(name: String): ImageIcon = load_icon(name) match { case icon: ImageIcon => icon case _ => error("Bad image icon: " + name) } /* key event handling */ def request_focus_view(alt_view: View = null): Unit = isabelle.jedit_base.JEdit_Lib.request_focus_view(alt_view) def propagate_key(view: View, evt: KeyEvent): Unit = { if (view != null && !evt.isConsumed) view.getInputHandler().processKeyEvent(evt, View.ACTION_BAR, false) } def key_listener( key_typed: KeyEvent => Unit = _ => (), key_pressed: KeyEvent => Unit = _ => (), key_released: KeyEvent => Unit = _ => ()): KeyListener = { def process_key_event(evt0: KeyEvent, handle: KeyEvent => Unit): Unit = { val evt = KeyEventWorkaround.processKeyEvent(evt0) if (evt != null) handle(evt) } new KeyListener { def keyTyped(evt: KeyEvent): Unit = process_key_event(evt, key_typed) def keyPressed(evt: KeyEvent): Unit = process_key_event(evt, key_pressed) def keyReleased(evt: KeyEvent): Unit = process_key_event(evt, key_released) } } def special_key(evt: KeyEvent): Boolean = { // cf. 5.2.0/jEdit/org/gjt/sp/jedit/gui/KeyEventWorkaround.java val mod = evt.getModifiersEx (mod & InputEvent.CTRL_DOWN_MASK) != 0 && (mod & InputEvent.ALT_DOWN_MASK) == 0 || (mod & InputEvent.CTRL_DOWN_MASK) == 0 && (mod & InputEvent.ALT_DOWN_MASK) != 0 && !Debug.ALT_KEY_PRESSED_DISABLED || (mod & InputEvent.META_DOWN_MASK) != 0 } def command_modifier(evt: InputEvent): Boolean = (evt.getModifiersEx & Toolkit.getDefaultToolkit.getMenuShortcutKeyMaskEx) != 0 def shift_modifier(evt: InputEvent): Boolean = (evt.getModifiersEx & InputEvent.SHIFT_DOWN_MASK) != 0 def modifier_string(evt: InputEvent): String = KeyEventTranslator.getModifierString(evt) match { case null => "" case s => s } } diff --git a/src/Tools/jEdit/src/jedit_sessions.scala b/src/Tools/jEdit/src/jedit_sessions.scala --- a/src/Tools/jEdit/src/jedit_sessions.scala +++ b/src/Tools/jEdit/src/jedit_sessions.scala @@ -1,152 +1,152 @@ /* Title: Tools/jEdit/src/jedit_sessions.scala Author: Makarius Isabelle/jEdit session information, based on implicit process environment and explicit options. */ package isabelle.jedit import isabelle._ import scala.swing.ComboBox import scala.swing.event.SelectionChanged object JEdit_Sessions { /* session options */ - def session_dirs(): List[Path] = + def session_dirs: List[Path] = Path.split(Isabelle_System.getenv("JEDIT_SESSION_DIRS")).filterNot(p => p.implode == "-") - def session_no_build(): Boolean = + def session_no_build: Boolean = Isabelle_System.getenv("JEDIT_NO_BUILD") == "true" def session_options(options: Options): Options = { val options1 = Isabelle_System.getenv("JEDIT_BUILD_MODE") match { case "default" => options case mode => options.bool.update("system_heaps", mode == "system") } val options2 = Isabelle_System.getenv("JEDIT_ML_PROCESS_POLICY") match { case "" => options1 case s => options1.string.update("ML_process_policy", s) } options2 } - def sessions_structure(options: Options, dirs: List[Path] = session_dirs()): Sessions.Structure = + def sessions_structure(options: Options, dirs: List[Path] = session_dirs): Sessions.Structure = Sessions.load_structure(session_options(options), dirs = dirs) /* raw logic info */ private val jedit_logic_option = "jedit_logic" def logic_name(options: Options): String = Isabelle_System.default_logic( Isabelle_System.getenv("JEDIT_LOGIC"), options.string(jedit_logic_option)) def logic_ancestor: Option[String] = proper_string(Isabelle_System.getenv("JEDIT_LOGIC_ANCESTOR")) def logic_requirements: Boolean = Isabelle_System.getenv("JEDIT_LOGIC_REQUIREMENTS") == "true" def logic_include_sessions: List[String] = space_explode(':', Isabelle_System.getenv("JEDIT_INCLUDE_SESSIONS")) def logic_info(options: Options): Option[Sessions.Info] = try { sessions_structure(options).get(logic_name(options)) } catch { case ERROR(_) => None } def logic_root(options: Options): Position.T = if (logic_requirements) logic_info(options).map(_.pos) getOrElse Position.none else Position.none /* logic selector */ private class Logic_Entry(val name: String, val description: String) { override def toString: String = description } def logic_selector(options: Options_Variable, autosave: Boolean): Option_Component = { GUI_Thread.require {} val session_list = { val sessions = sessions_structure(options.value) val (main_sessions, other_sessions) = sessions.imports_topological_order.partition(name => sessions(name).groups.contains("main")) main_sessions.sorted ::: other_sessions.sorted } val entries = new Logic_Entry("", "default (" + logic_name(options.value) + ")") :: session_list.map(name => new Logic_Entry(name, name)) val component = new ComboBox(entries) with Option_Component { name = jedit_logic_option val title = "Logic" def load(): Unit = { val logic = options.string(jedit_logic_option) entries.find(_.name == logic) match { case Some(entry) => selection.item = entry case None => } } def save(): Unit = options.string(jedit_logic_option) = selection.item.name } component.load() if (autosave) { component.listenTo(component.selection) component.reactions += { case SelectionChanged(_) => component.save() } } component.tooltip = "Logic session name (change requires restart)" component } /* session build process */ def session_base_info(options: Options): Sessions.Base_Info = Sessions.base_info(options, - dirs = JEdit_Sessions.session_dirs(), + dirs = JEdit_Sessions.session_dirs, include_sessions = logic_include_sessions, session = logic_name(options), session_ancestor = logic_ancestor, session_requirements = logic_requirements) def session_build( options: Options, progress: Progress = new Progress, no_build: Boolean = false): Int = { Build.build(session_options(options), selection = Sessions.Selection.session(PIDE.resources.session_name), - progress = progress, build_heap = true, no_build = no_build, dirs = session_dirs(), + progress = progress, build_heap = true, no_build = no_build, dirs = session_dirs, infos = PIDE.resources.session_base_info.infos).rc } def session_start(options0: Options): Unit = { val session = PIDE.session val options = session_options(options0) val sessions_structure = PIDE.resources.session_base_info.sessions_structure val store = Sessions.store(options) session.phase_changed += PIDE.plugin.session_phase_changed Isabelle_Process(session, options, sessions_structure, store, logic = PIDE.resources.session_name, modes = (space_explode(',', options.string("jedit_print_mode")) ::: space_explode(',', Isabelle_System.getenv("JEDIT_PRINT_MODE"))).reverse) } } diff --git a/src/Tools/jEdit/src/jedit_spell_checker.scala b/src/Tools/jEdit/src/jedit_spell_checker.scala --- a/src/Tools/jEdit/src/jedit_spell_checker.scala +++ b/src/Tools/jEdit/src/jedit_spell_checker.scala @@ -1,109 +1,109 @@ /* Title: Tools/jEdit/src/jedit_spell_checker.scala Author: Makarius Specific spell-checker support for Isabelle/jEdit. */ package isabelle.jedit import isabelle._ import javax.swing.JMenuItem import scala.swing.ComboBox import org.gjt.sp.jedit.menu.EnhancedMenuItem import org.gjt.sp.jedit.jEdit import org.gjt.sp.jedit.textarea.{JEditTextArea, TextArea} object JEdit_Spell_Checker { /* completion */ def completion( rendering: JEdit_Rendering, explicit: Boolean, caret: Text.Offset): Option[Completion.Result] = { PIDE.plugin.spell_checker.get match { case Some(spell_checker) if explicit => spell_checker.completion(rendering, caret) case _ => None } } /* context menu */ def context_menu(text_area: JEditTextArea, offset: Text.Offset): List[JMenuItem] = { val result = for { spell_checker <- PIDE.plugin.spell_checker.get doc_view <- Document_View.get(text_area) rendering = doc_view.get_rendering() range = JEdit_Lib.point_range(text_area.getBuffer, offset) Text.Info(_, word) <- Spell_Checker.current_word(rendering, range) } yield (spell_checker, word) result match { case Some((spell_checker, word)) => val context = jEdit.getActionContext() def item(name: String): JMenuItem = new EnhancedMenuItem(context.getAction(name).getLabel, name, context) val complete_items = if (spell_checker.complete(word).nonEmpty) List(item("isabelle.complete-word")) else Nil val update_items = if (spell_checker.check(word)) List(item("isabelle.exclude-word"), item("isabelle.exclude-word-permanently")) else List(item("isabelle.include-word"), item("isabelle.include-word-permanently")) val reset_items = spell_checker.reset_enabled() match { case 0 => Nil case n => val name = "isabelle.reset-words" val label = context.getAction(name).getLabel List(new EnhancedMenuItem(label + " (" + n + ")", name, context)) } complete_items ::: update_items ::: reset_items case None => Nil } } /* dictionaries */ def dictionaries_selector(): Option_Component = { GUI_Thread.require {} val option_name = "spell_checker_dictionary" val opt = PIDE.options.value.check_name(option_name) - val entries = Spell_Checker.dictionaries() + val entries = Spell_Checker.dictionaries val component = new ComboBox(entries) with Option_Component { name = option_name val title = opt.title() def load(): Unit = { val lang = PIDE.options.string(option_name) entries.find(_.lang == lang) match { case Some(entry) => selection.item = entry case None => } } def save(): Unit = PIDE.options.string(option_name) = selection.item.lang } component.load() component.tooltip = GUI.tooltip_lines(opt.print_default) component } } diff --git a/src/Tools/jEdit/src/plugin.scala b/src/Tools/jEdit/src/plugin.scala --- a/src/Tools/jEdit/src/plugin.scala +++ b/src/Tools/jEdit/src/plugin.scala @@ -1,504 +1,504 @@ /* Title: Tools/jEdit/src/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) + 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: Plugin = null def plugin: 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 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.last(options.seconds("editor_load_delay"), gui = true) { init_models() } private val delay_load_active = Synchronized(false) private def delay_load_activated(): Boolean = delay_load_active.guarded_access(a => Some((!a, true))) private def delay_load_action(): Unit = { if (Isabelle.continuous_checking && delay_load_activated() && PerspectiveManager.isPerspectiveEnabled) { if (JEdit_Lib.jedit_buffers().exists(_.isLoading)) delay_load.invoke() else { val required_files = { val models = Document_Model.get_models() val thys = (for ((node_name, model) <- models.iterator if model.is_theory) yield (node_name, Position.none)).toList val thy_files1 = resources.dependencies(thys).theories val thy_files2 = (for { (name, _) <- models.iterator thy_name <- resources.make_theory_name(name) } yield thy_name).toList val aux_files = if (options.bool("jedit_auto_resolve")) { val stable_tip_version = if (models.forall(p => p._2.is_stable)) session.get_state().stable_tip_version else None stable_tip_version match { case Some(version) => resources.undefined_blobs(version.nodes) case None => delay_load.invoke(); Nil } } else Nil (thy_files1 ::: thy_files2 ::: 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_active.change(_ => false) } } } } catch { case _: Throwable => delay_load_active.change(_ => false) } } else delay_load_active.change(_ => false) } } } private lazy val delay_load = Delay.last(options.seconds("editor_load_delay"), gui = true) { delay_load_action() } 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 (!Isabelle.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) Isabelle.continuous_checking = true } } 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) != Some(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_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 msg: 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 msg: EditorStarted => if (Distribution.is_identified && !Distribution.is_official) { GUI.warning_dialog(jEdit.getActiveView, "Isabelle version for testing", "This is " + Distribution.version + ".", "It is for testing only, not for production use.") } 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 msg: PropertiesChanged => for { - view <- JEdit_Lib.jedit_views + 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() val http_server: HTTP.Server = HTTP.server(Document_Model.http_handlers(http_root)) /* 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) + JEdit_Lib.jedit_views().foreach(init_title) isabelle.jedit_base.Syntax_Style.set_style_extender(Syntax_Style.Extender) init_mode_provider() - JEdit_Lib.jedit_text_areas.foreach(Completion_Popup.Text_Area.init) + JEdit_Lib.jedit_text_areas().foreach(Completion_Popup.Text_Area.init) - http_server.start + 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 + http_server.stop() isabelle.jedit_base.Syntax_Style.dummy_style_extender() exit_mode_provider() - JEdit_Lib.jedit_text_areas.foreach(Completion_Popup.Text_Area.exit) + 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() } } diff --git a/src/Tools/jEdit/src/pretty_text_area.scala b/src/Tools/jEdit/src/pretty_text_area.scala --- a/src/Tools/jEdit/src/pretty_text_area.scala +++ b/src/Tools/jEdit/src/pretty_text_area.scala @@ -1,246 +1,246 @@ /* Title: Tools/jEdit/src/pretty_text_area.scala Author: Makarius GUI component for pretty-printed text with markup, rendered like jEdit text area. */ package isabelle.jedit import isabelle._ import java.awt.{Color, Font, Toolkit, Window} import java.awt.event.{InputEvent, KeyEvent} import java.awt.im.InputMethodRequests import javax.swing.JTextField import javax.swing.event.{DocumentListener, DocumentEvent} import scala.swing.{Label, Component} import scala.util.matching.Regex import org.gjt.sp.jedit.{jEdit, View, Registers, JEditBeanShellAction} import org.gjt.sp.jedit.input.{DefaultInputHandlerProvider, TextAreaInputHandler} import org.gjt.sp.jedit.textarea.{AntiAlias, JEditEmbeddedTextArea} import org.gjt.sp.jedit.syntax.SyntaxStyle import org.gjt.sp.jedit.gui.KeyEventTranslator import org.gjt.sp.util.{SyntaxUtilities, Log} class Pretty_Text_Area( view: View, close_action: () => Unit = () => (), propagate_keys: Boolean = false) extends JEditEmbeddedTextArea { text_area => GUI_Thread.require {} private var current_font_info: Font_Info = Font_Info.main() private var current_body: XML.Body = Nil private var current_base_snapshot = Document.Snapshot.init private var current_base_results = Command.Results.empty private var current_rendering: JEdit_Rendering = JEdit_Rendering.text(current_base_snapshot, Nil)._2 private var future_refresh: Option[Future[Unit]] = None private val rich_text_area = new Rich_Text_Area(view, text_area, () => current_rendering, close_action, get_search_pattern _, () => (), caret_visible = false, enable_hovering = true) private var current_search_pattern: Option[Regex] = None def get_search_pattern(): Option[Regex] = GUI_Thread.require { current_search_pattern } def get_background(): Option[Color] = None def refresh(): Unit = { GUI_Thread.require {} val font = current_font_info.font getPainter.setFont(font) getPainter.setAntiAlias(new AntiAlias(jEdit.getProperty("view.antiAlias"))) getPainter.setFractionalFontMetricsEnabled(jEdit.getBooleanProperty("view.fracFontMetrics")) getPainter.setStyles( SyntaxUtilities.loadStyles(current_font_info.family, current_font_info.size.round)) getPainter.setLineExtraSpacing(jEdit.getIntegerProperty("options.textarea.lineSpacing", 0)) val fold_line_style = new Array[SyntaxStyle](4) for (i <- 0 to 3) { fold_line_style(i) = SyntaxUtilities.parseStyle( jEdit.getProperty("view.style.foldLine." + i), current_font_info.family, current_font_info.size.round, true) } getPainter.setFoldLineStyle(fold_line_style) getGutter.setForeground(jEdit.getColorProperty("view.gutter.fgColor")) getGutter.setBackground(jEdit.getColorProperty("view.gutter.bgColor")) get_background().foreach(bg => { getPainter.setBackground(bg); getGutter.setBackground(bg) }) getGutter.setHighlightedForeground(jEdit.getColorProperty("view.gutter.highlightColor")) getGutter.setFoldColor(jEdit.getColorProperty("view.gutter.foldColor")) getGutter.setFont(jEdit.getFontProperty("view.gutter.font")) getGutter.setBorder(0, jEdit.getColorProperty("view.gutter.focusBorderColor"), jEdit.getColorProperty("view.gutter.noFocusBorderColor"), getPainter.getBackground) getGutter.setFoldPainter(view.getTextArea.getFoldPainter) getGutter.setGutterEnabled(jEdit.getBooleanProperty("view.gutter.enabled")) if (getWidth > 0) { val metric = JEdit_Lib.pretty_metric(getPainter) val margin = ((getPainter.getWidth.toDouble / metric.unit) max 20.0).floor val snapshot = current_base_snapshot val results = current_base_results val formatted_body = Pretty.formatted(current_body, margin = margin, metric = metric) - future_refresh.foreach(_.cancel) + future_refresh.foreach(_.cancel()) future_refresh = Some(Future.fork { val (text, rendering) = try { JEdit_Rendering.text(snapshot, formatted_body, results = results) } catch { case exn: Throwable => Log.log(Log.ERROR, this, exn); throw exn } Exn.Interrupt.expose() GUI_Thread.later { current_rendering = rendering JEdit_Lib.buffer_edit(getBuffer) { rich_text_area.active_reset() getBuffer.setFoldHandler(new Fold_Handling.Document_Fold_Handler(rendering)) JEdit_Lib.buffer_undo_in_progress(getBuffer, setText(text)) setCaretPosition(0) } } }) } } def resize(font_info: Font_Info): Unit = { GUI_Thread.require {} current_font_info = font_info refresh() } def update( base_snapshot: Document.Snapshot, base_results: Command.Results, body: XML.Body): Unit = { GUI_Thread.require {} require(!base_snapshot.is_outdated, "document snapshot outdated") current_base_snapshot = base_snapshot current_base_results = base_results current_body = body refresh() } def detach: Unit = { GUI_Thread.require {} Info_Dockable(view, current_base_snapshot, current_base_results, current_body) } def detach_operation: Option[() => Unit] = if (current_body.isEmpty) None else Some(() => detach) /* common GUI components */ val search_label: Component = new Label("Search:") { tooltip = "Search and highlight output via regular expression" } val search_field: Component = Component.wrap(new Completion_Popup.History_Text_Field("isabelle-search") { private val input_delay = Delay.last(PIDE.options.seconds("editor_input_delay"), gui = true) { search_action(this) } getDocument.addDocumentListener(new DocumentListener { def changedUpdate(e: DocumentEvent): Unit = input_delay.invoke() def insertUpdate(e: DocumentEvent): Unit = input_delay.invoke() def removeUpdate(e: DocumentEvent): Unit = input_delay.invoke() }) setColumns(20) setToolTipText(search_label.tooltip) setFont(GUI.imitate_font(getFont, scale = 1.2)) }) private val search_field_foreground = search_field.foreground private def search_action(text_field: JTextField): Unit = { val (pattern, ok) = text_field.getText match { case null | "" => (None, true) case s => val re = Library.make_regex(s) (re, re.isDefined) } if (current_search_pattern != pattern) { current_search_pattern = pattern text_area.getPainter.repaint() } text_field.setForeground( if (ok) search_field_foreground else current_rendering.color(Rendering.Color.error)) } /* key handling */ override def getInputMethodRequests: InputMethodRequests = null inputHandlerProvider = new DefaultInputHandlerProvider(new TextAreaInputHandler(text_area) { override def getAction(action: String): JEditBeanShellAction = text_area.getActionContext.getAction(action) override def processKeyEvent(evt: KeyEvent, from: Int, global: Boolean): Unit = {} override def handleKey(key: KeyEventTranslator.Key, dry_run: Boolean): Boolean = false }) addKeyListener(JEdit_Lib.key_listener( key_pressed = (evt: KeyEvent) => { val strict_control = JEdit_Lib.command_modifier(evt) && !JEdit_Lib.shift_modifier(evt) evt.getKeyCode match { case KeyEvent.VK_C | KeyEvent.VK_INSERT if strict_control && text_area.getSelectionCount != 0 => Registers.copy(text_area, '$') evt.consume case KeyEvent.VK_A if strict_control => text_area.selectAll evt.consume case KeyEvent.VK_ESCAPE => if (Isabelle.dismissed_popups(view)) evt.consume case _ => } if (propagate_keys) JEdit_Lib.propagate_key(view, evt) }, key_typed = (evt: KeyEvent) => { if (propagate_keys) JEdit_Lib.propagate_key(view, evt) } ) ) /* init */ getPainter.setStructureHighlightEnabled(false) getPainter.setLineHighlightEnabled(false) getBuffer.setTokenMarker(Isabelle.mode_token_marker("isabelle-output").get) getBuffer.setStringProperty("noWordSep", "_'?⇩") rich_text_area.activate() } diff --git a/src/Tools/jEdit/src/pretty_tooltip.scala b/src/Tools/jEdit/src/pretty_tooltip.scala --- a/src/Tools/jEdit/src/pretty_tooltip.scala +++ b/src/Tools/jEdit/src/pretty_tooltip.scala @@ -1,291 +1,291 @@ /* Title: Tools/jEdit/src/pretty_tooltip.scala Author: Makarius Tooltip based on Pretty_Text_Area. */ package isabelle.jedit import isabelle._ import java.awt.{Color, Point, BorderLayout, Dimension} import java.awt.event.{FocusAdapter, FocusEvent} import javax.swing.{JPanel, JComponent, SwingUtilities, JLayeredPane} import javax.swing.border.LineBorder import scala.swing.{FlowPanel, Label} import scala.swing.event.MouseClicked import org.gjt.sp.jedit.View import org.gjt.sp.jedit.textarea.TextArea object Pretty_Tooltip { /* tooltip hierarchy */ // owned by GUI thread private var stack: List[Pretty_Tooltip] = Nil private def hierarchy(tip: Pretty_Tooltip): Option[(List[Pretty_Tooltip], List[Pretty_Tooltip])] = { GUI_Thread.require {} if (stack.contains(tip)) Some(stack.span(_ != tip)) else None } private def descendant(parent: JComponent): Option[Pretty_Tooltip] = GUI_Thread.require { stack.find(tip => tip.original_parent == parent) } def apply( view: View, parent: JComponent, location: Point, rendering: JEdit_Rendering, results: Command.Results, info: Text.Info[XML.Body]): Unit = { GUI_Thread.require {} stack match { case top :: _ if top.results == results && top.info == info => case _ => GUI.layered_pane(parent) match { case None => case Some(layered) => val (old, rest) = GUI.ancestors(parent).collectFirst({ case x: Pretty_Tooltip => x }) match { case Some(tip) => hierarchy(tip).getOrElse((stack, Nil)) case None => (stack, Nil) } old.foreach(_.hide_popup) val loc = SwingUtilities.convertPoint(parent, location, layered) val tip = new Pretty_Tooltip(view, layered, parent, loc, rendering, results, info) stack = tip :: rest tip.show_popup } } } /* pending event and active state */ // owned by GUI thread private var pending: Option[() => Unit] = None private var active = true private val pending_delay = Delay.last(PIDE.options.seconds("jedit_tooltip_delay"), gui = true) { pending match { case Some(body) => pending = None; body() case None => } } def invoke(body: () => Unit): Unit = GUI_Thread.require { if (active) { pending = Some(body) pending_delay.invoke() } } def revoke(): Unit = GUI_Thread.require { pending = None pending_delay.revoke() } private lazy val reactivate_delay = Delay.last(PIDE.options.seconds("jedit_tooltip_delay"), gui = true) { active = true } private def deactivate(): Unit = GUI_Thread.require { revoke() active = false reactivate_delay.invoke() } /* dismiss */ private lazy val focus_delay = Delay.last(PIDE.options.seconds("editor_input_delay"), gui = true) { dismiss_unfocused() } def dismiss_unfocused(): Unit = { stack.span(tip => !tip.pretty_text_area.isFocusOwner) match { case (Nil, _) => case (unfocused, rest) => deactivate() unfocused.foreach(_.hide_popup) stack = rest } } def dismiss(tip: Pretty_Tooltip): Unit = { deactivate() hierarchy(tip) match { case Some((old, _ :: rest)) => rest match { case top :: _ => top.request_focus case Nil => JEdit_Lib.request_focus_view() } old.foreach(_.hide_popup) tip.hide_popup stack = rest case _ => } } def dismiss_descendant(parent: JComponent): Unit = descendant(parent).foreach(dismiss) def dismissed_all(): Boolean = { deactivate() if (stack.isEmpty) false else { JEdit_Lib.request_focus_view() stack.foreach(_.hide_popup) stack = Nil true } } } class Pretty_Tooltip private( view: View, layered: JLayeredPane, val original_parent: JComponent, location: Point, rendering: JEdit_Rendering, private val results: Command.Results, private val info: Text.Info[XML.Body]) extends JPanel(new BorderLayout) { tip => GUI_Thread.require {} /* controls */ private val close = new Label { icon = rendering.tooltip_close_icon tooltip = "Close tooltip window" listenTo(mouse.clicks) reactions += { case _: MouseClicked => Pretty_Tooltip.dismiss(tip) } } private val detach = new Label { icon = rendering.tooltip_detach_icon tooltip = "Detach tooltip window" listenTo(mouse.clicks) reactions += { case _: MouseClicked => Info_Dockable(view, rendering.snapshot, results, info.info) Pretty_Tooltip.dismiss(tip) } } private val controls = new FlowPanel(FlowPanel.Alignment.Left)(close, detach) { background = rendering.tooltip_color } /* text area */ val pretty_text_area: Pretty_Text_Area = new Pretty_Text_Area(view, () => Pretty_Tooltip.dismiss(tip), true) { override def get_background() = Some(rendering.tooltip_color) } pretty_text_area.addFocusListener(new FocusAdapter { override def focusGained(e: FocusEvent): Unit = { tip_border(true) Pretty_Tooltip.focus_delay.invoke() } override def focusLost(e: FocusEvent): Unit = { tip_border(false) Pretty_Tooltip.focus_delay.invoke() } }) pretty_text_area.resize(Font_Info.main(PIDE.options.real("jedit_popup_font_scale"))) /* main content */ def tip_border(has_focus: Boolean): Unit = { tip.setBorder(new LineBorder(if (has_focus) Color.BLACK else Color.GRAY)) tip.repaint() } tip_border(true) override def getFocusTraversalKeysEnabled = false tip.setBackground(rendering.tooltip_color) tip.add(controls.peer, BorderLayout.NORTH) tip.add(pretty_text_area) /* popup */ private val popup = { val screen = GUI.screen_location(layered, location) val size = { val bounds = JEdit_Rendering.popup_bounds val w_max = layered.getWidth min (screen.bounds.width * bounds).toInt val h_max = layered.getHeight min (screen.bounds.height * bounds).toInt val painter = pretty_text_area.getPainter val geometry = JEdit_Lib.window_geometry(tip, painter) val metric = JEdit_Lib.pretty_metric(painter) val margin = ((rendering.tooltip_margin * metric.average) min ((w_max - geometry.deco_width) / metric.unit).toInt) max 20 val formatted = Pretty.formatted(info.info, margin = margin, metric = metric) val lines = XML.traverse_text(formatted)(if (XML.text_length(formatted) == 0) 0 else 1)( (n: Int, s: String) => n + s.iterator.count(_ == '\n')) val h = painter.getLineHeight * lines + geometry.deco_height val margin1 = if (h <= h_max) { split_lines(XML.content(formatted)).foldLeft(0.0) { case (m, line) => m max metric(line) } } else margin val w = (metric.unit * (margin1 + metric.average)).round.toInt + geometry.deco_width new Dimension(w min w_max, h min h_max) } new Popup(layered, tip, screen.relative(layered, size), size) } private def show_popup: Unit = { popup.show - pretty_text_area.requestFocus + pretty_text_area.requestFocus() pretty_text_area.update(rendering.snapshot, results, info.info) } private def hide_popup: Unit = popup.hide - private def request_focus: Unit = pretty_text_area.requestFocus + private def request_focus: Unit = pretty_text_area.requestFocus() } diff --git a/src/Tools/jEdit/src/process_indicator.scala b/src/Tools/jEdit/src/process_indicator.scala --- a/src/Tools/jEdit/src/process_indicator.scala +++ b/src/Tools/jEdit/src/process_indicator.scala @@ -1,69 +1,69 @@ /* Title: Tools/jEdit/src/process_indicator.scala Author: Makarius Process indicator with animated icon. */ package isabelle.jedit import isabelle._ import java.awt.Image import java.awt.event.{ActionListener, ActionEvent} import javax.swing.{ImageIcon, Timer} import scala.swing.{Label, Component} class Process_Indicator { private val label = new Label private val passive_icon = JEdit_Lib.load_image_icon(PIDE.options.string("process_passive_icon")).getImage private val active_icons = space_explode(':', PIDE.options.string("process_active_icons")).map(name => JEdit_Lib.load_image_icon(name).getImage) private class Animation extends ImageIcon(passive_icon) { private var current_frame = 0 private val timer = new Timer(0, new ActionListener { override def actionPerformed(e: ActionEvent): Unit = { current_frame = (current_frame + 1) % active_icons.length setImage(active_icons(current_frame)) - label.repaint + label.repaint() } }) timer.setRepeats(true) def update(rate: Int): Unit = { if (rate == 0) { setImage(passive_icon) timer.stop - label.repaint + label.repaint() } else { val delay = 1000 / rate timer.setInitialDelay(delay) timer.setDelay(delay) timer.restart } } } private val animation = new Animation label.icon = animation def component: Component = label def update(tip: String, rate: Int): Unit = { label.tooltip = tip animation.update(rate) } } diff --git a/src/Tools/jEdit/src/query_dockable.scala b/src/Tools/jEdit/src/query_dockable.scala --- a/src/Tools/jEdit/src/query_dockable.scala +++ b/src/Tools/jEdit/src/query_dockable.scala @@ -1,353 +1,353 @@ /* Title: Tools/jEdit/src/query_dockable.scala Author: Makarius Dockable window for query operations. */ package isabelle.jedit import isabelle._ import isabelle.jedit_base.Dockable import java.awt.event.{ComponentEvent, ComponentAdapter, KeyEvent} import javax.swing.{JComponent, JTextField} import scala.swing.{Button, Component, TextField, CheckBox, Label, ListView, ComboBox, TabbedPane, BorderPanel} import scala.swing.event.{SelectionChanged, ButtonClicked, Key, KeyPressed} import org.gjt.sp.jedit.View object Query_Dockable { private abstract class Operation(view: View) { val pretty_text_area = new Pretty_Text_Area(view) def query_operation: Query_Operation[View] def query: JComponent def select: Unit def page: TabbedPane.Page } } class Query_Dockable(view: View, position: String) extends Dockable(view, position) { /* common GUI components */ private val zoom = new Font_Info.Zoom_Box { def changed = handle_resize() } private def make_query(property: String, tooltip: String, apply_query: () => Unit) : Completion_Popup.History_Text_Field = new Completion_Popup.History_Text_Field(property) { override def processKeyEvent(evt: KeyEvent): Unit = { if (evt.getID == KeyEvent.KEY_PRESSED && evt.getKeyCode == KeyEvent.VK_ENTER) apply_query() super.processKeyEvent(evt) } { val max = getPreferredSize; max.width = Integer.MAX_VALUE; setMaximumSize(max) } setColumns(40) setToolTipText(tooltip) setFont(GUI.imitate_font(getFont, scale = 1.2)) } /* consume status */ def consume_status( process_indicator: Process_Indicator, status: Query_Operation.Status.Value): Unit = { status match { case Query_Operation.Status.WAITING => process_indicator.update("Waiting for evaluation of context ...", 5) case Query_Operation.Status.RUNNING => process_indicator.update("Running find operation ...", 15) case Query_Operation.Status.FINISHED => process_indicator.update(null, 0) } } /* find theorems */ private val find_theorems = new Query_Dockable.Operation(view) { /* query */ private val process_indicator = new Process_Indicator val query_operation = new Query_Operation(PIDE.editor, view, "find_theorems", consume_status(process_indicator, _), (snapshot, results, body) => pretty_text_area.update(snapshot, results, Pretty.separate(body))) private def apply_query(): Unit = { query.addCurrentToHistory() query_operation.apply_query(List(limit.text, allow_dups.selected.toString, query.getText)) } private val query_label = new Label("Find:") { tooltip = GUI.tooltip_lines( "Search criteria for find operation, e.g.\n\"_ = _\" \"(+)\" name: Group -name: monoid") } val query = make_query("isabelle-find-theorems", query_label.tooltip, apply_query _) /* GUI page */ private val limit = new TextField(PIDE.options.int("find_theorems_limit").toString, 5) { tooltip = "Limit of displayed results" verifier = (s: String) => s match { case Value.Int(x) => x >= 0 case _ => false } listenTo(keys) reactions += { case KeyPressed(_, Key.Enter, 0, _) => apply_query() } } private val allow_dups = new CheckBox("Duplicates") { tooltip = "Show all versions of matching theorems" selected = false reactions += { case ButtonClicked(_) => apply_query() } } private val apply_button = new Button("Apply") { tooltip = "Find theorems meeting specified criteria" reactions += { case ButtonClicked(_) => apply_query() } } private val control_panel = Wrap_Panel( List(query_label, Component.wrap(query), limit, allow_dups, process_indicator.component, apply_button, pretty_text_area.search_label, pretty_text_area.search_field)) def select: Unit = { control_panel.contents += zoom } val page = new TabbedPane.Page("Find Theorems", new BorderPanel { layout(control_panel) = BorderPanel.Position.North layout(Component.wrap(pretty_text_area)) = BorderPanel.Position.Center }, apply_button.tooltip) } /* find consts */ private val find_consts = new Query_Dockable.Operation(view) { /* query */ private val process_indicator = new Process_Indicator val query_operation = new Query_Operation(PIDE.editor, view, "find_consts", consume_status(process_indicator, _), (snapshot, results, body) => pretty_text_area.update(snapshot, results, Pretty.separate(body))) private def apply_query(): Unit = { query.addCurrentToHistory() query_operation.apply_query(List(query.getText)) } private val query_label = new Label("Find:") { tooltip = GUI.tooltip_lines("Name / type patterns for constants") } val query = make_query("isabelle-find-consts", query_label.tooltip, apply_query _) /* GUI page */ private val apply_button = new Button("Apply") { tooltip = "Find constants by name / type patterns" reactions += { case ButtonClicked(_) => apply_query() } } private val control_panel = Wrap_Panel( List( query_label, Component.wrap(query), process_indicator.component, apply_button, pretty_text_area.search_label, pretty_text_area.search_field)) def select: Unit = { control_panel.contents += zoom } val page = new TabbedPane.Page("Find Constants", new BorderPanel { layout(control_panel) = BorderPanel.Position.North layout(Component.wrap(pretty_text_area)) = BorderPanel.Position.Center }, apply_button.tooltip) } /* print operation */ private val print_operation = new Query_Dockable.Operation(view) { /* items */ private class Item(val name: String, description: String, sel: Boolean) { val checkbox = new CheckBox(name) { tooltip = "Print " + description selected = sel reactions += { case ButtonClicked(_) => apply_query() } } } private var _items: List[Item] = Nil private def selected_items(): List[String] = for (item <- _items if item.checkbox.selected) yield item.name private def update_items(): List[Item] = { val old_items = _items def was_selected(name: String): Boolean = old_items.find(item => item.name == name) match { case None => false case Some(item) => item.checkbox.selected } _items = for ((name, description) <- Print_Operation.get(PIDE.session)) yield new Item(name, description, was_selected(name)) _items } /* query */ private val process_indicator = new Process_Indicator val query_operation = new Query_Operation(PIDE.editor, view, "print_operation", consume_status(process_indicator, _), (snapshot, results, body) => pretty_text_area.update(snapshot, results, Pretty.separate(body))) private def apply_query(): Unit = query_operation.apply_query(selected_items()) private val query_label = new Label("Print:") def query: JComponent = apply_button.peer update_items() /* GUI page */ private val apply_button = new Button("Apply") { tooltip = "Apply to current context" listenTo(keys) reactions += { case ButtonClicked(_) => apply_query() case evt @ KeyPressed(_, Key.Enter, 0, _) => evt.peer.consume apply_query() } } private val control_panel = Wrap_Panel() def select: Unit = { - control_panel.contents.clear + control_panel.contents.clear() control_panel.contents += query_label update_items().foreach(item => control_panel.contents += item.checkbox) control_panel.contents ++= List(process_indicator.component, apply_button, pretty_text_area.search_label, pretty_text_area.search_field, zoom) } val page = new TabbedPane.Page("Print Context", new BorderPanel { layout(control_panel) = BorderPanel.Position.North layout(Component.wrap(pretty_text_area)) = BorderPanel.Position.Center }, "Print information from context") } /* operations */ private val operations = List(find_theorems, find_consts, print_operation) private val operations_pane = new TabbedPane { pages ++= operations.map(_.page) listenTo(selection) reactions += { case SelectionChanged(_) => select_operation() } } private def get_operation(): Option[Query_Dockable.Operation] = try { Some(operations(operations_pane.selection.index)) } catch { case _: IndexOutOfBoundsException => None } private def select_operation(): Unit = { - for (op <- get_operation()) { op.select; op.query.requestFocus } - operations_pane.revalidate + for (op <- get_operation()) { op.select; op.query.requestFocus() } + operations_pane.revalidate() } override def focusOnDefaultComponent(): Unit = { - for (op <- get_operation()) op.query.requestFocus + for (op <- get_operation()) op.query.requestFocus() } select_operation() set_content(operations_pane) override def detach_operation: Option[() => Unit] = get_operation() match { case None => None case Some(op) => op.pretty_text_area.detach_operation } /* resize */ private def handle_resize(): Unit = GUI_Thread.require { for (op <- operations) { op.pretty_text_area.resize( Font_Info.main(PIDE.options.real("jedit_font_scale") * zoom.factor / 100)) } } private val delay_resize = Delay.first(PIDE.options.seconds("editor_update_delay"), gui = true) { handle_resize() } addComponentListener(new ComponentAdapter { override def componentResized(e: ComponentEvent): Unit = delay_resize.invoke() override def componentShown(e: ComponentEvent): Unit = delay_resize.invoke() }) /* main */ private val main = Session.Consumer[Session.Global_Options](getClass.getName) { case _: Session.Global_Options => GUI_Thread.later { handle_resize() } } override def init(): Unit = { PIDE.session.global_options += main handle_resize() operations.foreach(op => op.query_operation.activate()) } override def exit(): Unit = { operations.foreach(op => op.query_operation.deactivate()) PIDE.session.global_options -= main delay_resize.revoke() } } diff --git a/src/Tools/jEdit/src/rich_text_area.scala b/src/Tools/jEdit/src/rich_text_area.scala --- a/src/Tools/jEdit/src/rich_text_area.scala +++ b/src/Tools/jEdit/src/rich_text_area.scala @@ -1,732 +1,732 @@ /* Title: Tools/jEdit/src/rich_text_area.scala Author: Makarius Enhanced version of jEdit text area, with rich text rendering, tooltips, hyperlinks etc. */ package isabelle.jedit import isabelle._ import java.awt.{Graphics2D, Shape, Color, Point, Cursor, MouseInfo} import java.awt.event.{MouseMotionAdapter, MouseAdapter, MouseEvent, FocusAdapter, FocusEvent, WindowEvent, WindowAdapter, KeyEvent} import java.awt.font.TextAttribute import javax.swing.SwingUtilities import java.text.AttributedString import scala.util.matching.Regex import org.gjt.sp.util.Log import org.gjt.sp.jedit.View import org.gjt.sp.jedit.syntax.Chunk import org.gjt.sp.jedit.textarea.{TextAreaExtension, TextAreaPainter, TextArea} class Rich_Text_Area( view: View, text_area: TextArea, get_rendering: () => JEdit_Rendering, close_action: () => Unit, get_search_pattern: () => Option[Regex], caret_update: () => Unit, caret_visible: Boolean, enable_hovering: Boolean) { private val buffer = text_area.getBuffer /* robust extension body */ def check_robust_body: Boolean = GUI_Thread.require { buffer == text_area.getBuffer } def robust_body[A](default: A)(body: => A): A = { try { if (check_robust_body) body else { Log.log(Log.ERROR, this, ERROR("Implicit change of text area buffer")) default } } catch { case exn: Throwable => Log.log(Log.ERROR, this, exn); default } } /* original painters */ private def pick_extension(name: String): TextAreaExtension = { text_area.getPainter.getExtensions.iterator.filter(x => x.getClass.getName == name).toList match { case List(x) => x case _ => error("Expected exactly one " + name) } } private val orig_text_painter = pick_extension("org.gjt.sp.jedit.textarea.TextAreaPainter$PaintText") /* caret focus modifier */ @volatile private var caret_focus_modifier = false def caret_focus_range: Text.Range = if (caret_focus_modifier) Text.Range.full else JEdit_Lib.visible_range(text_area) getOrElse Text.Range.offside private val key_listener = JEdit_Lib.key_listener( key_pressed = (evt: KeyEvent) => { val mod = PIDE.options.string("jedit_focus_modifier") val old = caret_focus_modifier caret_focus_modifier = (mod.nonEmpty && mod == JEdit_Lib.modifier_string(evt)) if (caret_focus_modifier != old) caret_update() }, key_released = _ => { if (caret_focus_modifier) { caret_focus_modifier = false caret_update() } }) /* common painter state */ @volatile private var painter_rendering: JEdit_Rendering = null @volatile private var painter_clip: Shape = null @volatile private var caret_focus = Rendering.Focus.empty private val set_state = new TextAreaExtension { override def paintScreenLineRange(gfx: Graphics2D, first_line: Int, last_line: Int, physical_lines: Array[Int], start: Array[Int], end: Array[Int], y: Int, line_height: Int): Unit = { painter_rendering = get_rendering() painter_clip = gfx.getClip caret_focus = if (caret_enabled && !painter_rendering.snapshot.is_outdated) { painter_rendering.caret_focus(JEdit_Lib.caret_range(text_area), caret_focus_range) } else Rendering.Focus.empty } } private val reset_state = new TextAreaExtension { override def paintScreenLineRange(gfx: Graphics2D, first_line: Int, last_line: Int, physical_lines: Array[Int], start: Array[Int], end: Array[Int], y: Int, line_height: Int): Unit = { painter_rendering = null painter_clip = null caret_focus = Rendering.Focus.empty } } def robust_rendering(body: JEdit_Rendering => Unit): Unit = { robust_body(()) { body(painter_rendering) } } /* active areas within the text */ private class Active_Area[A]( rendering: JEdit_Rendering => Text.Range => Option[Text.Info[A]], cursor: Option[Int]) { private var the_text_info: Option[(String, Text.Info[A])] = None def is_active: Boolean = the_text_info.isDefined def text_info: Option[(String, Text.Info[A])] = the_text_info def info: Option[Text.Info[A]] = the_text_info.map(_._2) def update(new_info: Option[Text.Info[A]]): Unit = { val old_text_info = the_text_info val new_text_info = new_info.map(info => (text_area.getText(info.range.start, info.range.length), info)) if (new_text_info != old_text_info) { caret_update() if (cursor.isDefined) { if (new_text_info.isDefined) text_area.getPainter.setCursor(Cursor.getPredefinedCursor(cursor.get)) else text_area.getPainter.resetCursor() } for { r0 <- JEdit_Lib.visible_range(text_area) opt <- List(old_text_info, new_text_info) (_, Text.Info(r1, _)) <- opt r2 <- r1.try_restrict(r0) // FIXME more precise?! } JEdit_Lib.invalidate_range(text_area, r2) the_text_info = new_text_info } } def update_rendering(r: JEdit_Rendering, range: Text.Range): Unit = update(rendering(r)(range)) def reset: Unit = update(None) } // owned by GUI thread private val highlight_area = new Active_Area[Color]( (rendering: JEdit_Rendering) => rendering.highlight, None) private val hyperlink_area = new Active_Area[PIDE.editor.Hyperlink]( (rendering: JEdit_Rendering) => rendering.hyperlink, Some(Cursor.HAND_CURSOR)) private val active_area = new Active_Area[XML.Elem]( (rendering: JEdit_Rendering) => rendering.active, Some(Cursor.DEFAULT_CURSOR)) private val active_areas = List((highlight_area, true), (hyperlink_area, true), (active_area, false)) def active_reset(): Unit = active_areas.foreach(_._1.reset) private def area_active(): Boolean = active_areas.exists({ case (area, _) => area.is_active }) private val focus_listener = new FocusAdapter { override def focusLost(e: FocusEvent): Unit = { robust_body(()) { active_reset() } } } private val window_listener = new WindowAdapter { override def windowIconified(e: WindowEvent): Unit = { robust_body(()) { active_reset() } } override def windowDeactivated(e: WindowEvent): Unit = { robust_body(()) { active_reset() } } } private val mouse_listener = new MouseAdapter { override def mouseClicked(e: MouseEvent): Unit = { robust_body(()) { hyperlink_area.info match { case Some(Text.Info(range, link)) => if (!link.external) { try { text_area.moveCaretPosition(range.start) } catch { case _: ArrayIndexOutOfBoundsException => case _: IllegalArgumentException => } - text_area.requestFocus + text_area.requestFocus() } link.follow(view) case None => } active_area.text_info match { case Some((text, Text.Info(_, markup))) => Active.action(view, text, markup) close_action() case None => } } } } private def mouse_inside_painter(): Boolean = MouseInfo.getPointerInfo match { case null => false case info => val point = info.getLocation val painter = text_area.getPainter SwingUtilities.convertPointFromScreen(point, painter) painter.contains(point) } private val mouse_motion_listener = new MouseMotionAdapter { override def mouseDragged(evt: MouseEvent): Unit = { robust_body(()) { active_reset() Completion_Popup.Text_Area.dismissed(text_area) Pretty_Tooltip.dismiss_descendant(text_area.getPainter) } } override def mouseMoved(evt: MouseEvent): Unit = { robust_body(()) { val x = evt.getX val y = evt.getY val control = JEdit_Lib.command_modifier(evt) if ((control || enable_hovering) && !buffer.isLoading) { JEdit_Lib.buffer_lock(buffer) { JEdit_Lib.pixel_range(text_area, x, y) match { case None => active_reset() case Some(range) => val rendering = get_rendering() for ((area, require_control) <- active_areas) { if (control == require_control && !rendering.snapshot.is_outdated) area.update_rendering(rendering, range) else area.reset } } } } else active_reset() if (evt.getSource == text_area.getPainter) { Pretty_Tooltip.invoke(() => robust_body(()) { if (mouse_inside_painter()) { val rendering = get_rendering() val snapshot = rendering.snapshot if (!snapshot.is_outdated) { JEdit_Lib.pixel_range(text_area, x, y) match { case None => case Some(range) => rendering.tooltip(range, control) match { case None => case Some(tip) => val painter = text_area.getPainter val loc = new Point(x, y + painter.getLineHeight / 2) val results = snapshot.command_results(tip.range) Pretty_Tooltip(view, painter, loc, rendering, results, tip) } } } } }) } } } } /* text background */ private val background_painter = new TextAreaExtension { override def paintScreenLineRange(gfx: Graphics2D, first_line: Int, last_line: Int, physical_lines: Array[Int], start: Array[Int], end: Array[Int], y: Int, line_height: Int): Unit = { robust_rendering { rendering => val fm = text_area.getPainter.getFontMetrics for (i <- physical_lines.indices) { if (physical_lines(i) != -1) { val line_range = Text.Range(start(i), end(i) min buffer.getLength) // line background color for { (c, separator) <- rendering.line_background(line_range) } { gfx.setColor(rendering.color(c)) val sep = if (separator) 2 min (line_height / 2) else 0 gfx.fillRect(0, y + i * line_height, text_area.getWidth, line_height - sep) } // background color for { Text.Info(range, c) <- rendering.background(Rendering.background_elements, line_range, caret_focus) r <- JEdit_Lib.gfx_range(text_area, range) } { gfx.setColor(rendering.color(c)) gfx.fillRect(r.x, y + i * line_height, r.length, line_height) } // active area -- potentially from other snapshot for { info <- active_area.info Text.Info(range, _) <- info.try_restrict(line_range) r <- JEdit_Lib.gfx_range(text_area, range) } { gfx.setColor(rendering.active_hover_color) gfx.fillRect(r.x, y + i * line_height, r.length, line_height) } // squiggly underline for { Text.Info(range, c) <- rendering.squiggly_underline(line_range) r <- JEdit_Lib.gfx_range(text_area, range) } { gfx.setColor(rendering.color(c)) val x0 = (r.x / 2) * 2 val y0 = r.y + fm.getAscent + 1 for (x1 <- Range(x0, x0 + r.length, 2)) { val y1 = if (x1 % 4 < 2) y0 else y0 + 1 gfx.drawLine(x1, y1, x1 + 1, y1) } } // spell checker for { spell_checker <- PIDE.plugin.spell_checker.get spell <- rendering.spell_checker(line_range) text <- JEdit_Lib.get_text(buffer, spell.range) info <- spell_checker.marked_words(spell.range.start, text) r <- JEdit_Lib.gfx_range(text_area, info.range) } { gfx.setColor(rendering.spell_checker_color) val y0 = r.y + ((fm.getAscent + 4) min (line_height - 2)) gfx.drawLine(r.x, y0, r.x + r.length, y0) } } } } } } /* text */ private def caret_enabled: Boolean = caret_visible && (!text_area.hasFocus || text_area.isCaretVisible) private def caret_color(rendering: JEdit_Rendering, offset: Text.Offset): Color = { if (text_area.isCaretVisible) text_area.getPainter.getCaretColor else { val debug_positions = (for { c <- PIDE.session.debugger.focus().iterator pos <- c.debug_position.iterator } yield pos).toList if (debug_positions.exists(PIDE.editor.is_hyperlink_position(rendering.snapshot, offset, _))) rendering.caret_debugger_color else rendering.caret_invisible_color } } private def paint_chunk_list(rendering: JEdit_Rendering, gfx: Graphics2D, line_start: Text.Offset, head: Chunk, x: Float, y: Float): Float = { val clip_rect = gfx.getClipBounds val painter = text_area.getPainter val font_context = painter.getFontRenderContext val caret_range = if (caret_enabled) JEdit_Lib.caret_range(text_area) else Text.Range.offside var w = 0.0f var chunk = head while (chunk != null) { val chunk_offset = line_start + chunk.offset if (x + w + chunk.width > clip_rect.x && x + w < clip_rect.x + clip_rect.width && chunk.length > 0) { val chunk_range = Text.Range(chunk_offset, chunk_offset + chunk.length) val chunk_str = if (chunk.chars == null) Symbol.spaces(chunk.length) else { if (chunk.str == null) { chunk.str = new String(chunk.chars) } chunk.str } val chunk_font = chunk.style.getFont val chunk_color = chunk.style.getForegroundColor def string_width(s: String): Float = if (s.isEmpty) 0.0f else chunk_font.getStringBounds(s, font_context).getWidth.toFloat val markup = for { r1 <- rendering.text_color(chunk_range, chunk_color) r2 <- r1.try_restrict(chunk_range) } yield r2 val padded_markup_iterator = if (markup.isEmpty) Iterator(Text.Info(chunk_range, chunk_color)) else Iterator( Text.Info(Text.Range(chunk_range.start, markup.head.range.start), chunk_color)) ++ markup.iterator ++ Iterator(Text.Info(Text.Range(markup.last.range.stop, chunk_range.stop), chunk_color)) var x1 = x + w gfx.setFont(chunk_font) for (Text.Info(range, color) <- padded_markup_iterator if !range.is_singularity) { val str = chunk_str.substring(range.start - chunk_offset, range.stop - chunk_offset) gfx.setColor(color) range.try_restrict(caret_range) match { case Some(r) if !r.is_singularity => val i = r.start - range.start val j = r.stop - range.start val s1 = str.substring(0, i) val s2 = str.substring(i, j) val s3 = str.substring(j) if (s1.nonEmpty) gfx.drawString(Word.bidi_override(s1), x1, y) val astr = new AttributedString(Word.bidi_override(s2)) astr.addAttribute(TextAttribute.FONT, chunk_font) astr.addAttribute(TextAttribute.FOREGROUND, caret_color(rendering, r.start)) astr.addAttribute(TextAttribute.SWAP_COLORS, TextAttribute.SWAP_COLORS_ON) gfx.drawString(astr.getIterator, x1 + string_width(s1), y) if (s3.nonEmpty) gfx.drawString(Word.bidi_override(s3), x1 + string_width(str.substring(0, j)), y) case _ => gfx.drawString(Word.bidi_override(str), x1, y) } x1 += string_width(str) } } w += chunk.width chunk = chunk.next.asInstanceOf[Chunk] } w } private val text_painter = new TextAreaExtension { override def paintScreenLineRange(gfx: Graphics2D, first_line: Int, last_line: Int, physical_lines: Array[Int], start: Array[Int], end: Array[Int], y: Int, line_height: Int): Unit = { robust_rendering { rendering => val painter = text_area.getPainter val fm = painter.getFontMetrics val lm = painter.getFont.getLineMetrics(" ", painter.getFontRenderContext) val clip = gfx.getClip val x0 = text_area.getHorizontalOffset var y0 = y + painter.getLineHeight - (fm.getLeading + 1) - fm.getDescent val (bullet_x, bullet_y, bullet_w, bullet_h) = { val w = fm.charWidth(' ') val b = (w / 2) max 1 val c = (lm.getAscent + lm.getStrikethroughOffset).round.toInt ((w - b + 1) / 2, c - b / 2, w - b, line_height - b) } for (i <- physical_lines.indices) { val line = physical_lines(i) if (line != -1) { val line_range = Text.Range(start(i), end(i) min buffer.getLength) // text chunks val screen_line = first_line + i val chunks = text_area.getChunksOfScreenLine(screen_line) if (chunks != null) { try { val line_start = buffer.getLineStartOffset(line) gfx.clipRect(x0, y + line_height * i, Integer.MAX_VALUE, line_height) val w = paint_chunk_list(rendering, gfx, line_start, chunks, x0.toFloat, y0.toFloat).toInt gfx.clipRect(x0 + w.toInt, 0, Integer.MAX_VALUE, Integer.MAX_VALUE) orig_text_painter.paintValidLine(gfx, screen_line, line, start(i), end(i), y + line_height * i) } finally { gfx.setClip(clip) } } // bullet bar for { Text.Info(range, color) <- rendering.bullet(line_range) r <- JEdit_Lib.gfx_range(text_area, range) } { gfx.setColor(color) gfx.fillRect(r.x + bullet_x, y + i * line_height + bullet_y, r.length - bullet_w, line_height - bullet_h) } } y0 += line_height } } } } /* foreground */ private val foreground_painter = new TextAreaExtension { override def paintScreenLineRange(gfx: Graphics2D, first_line: Int, last_line: Int, physical_lines: Array[Int], start: Array[Int], end: Array[Int], y: Int, line_height: Int): Unit = { robust_rendering { rendering => val search_pattern = get_search_pattern() for (i <- physical_lines.indices) { if (physical_lines(i) != -1) { val line_range = Text.Range(start(i), end(i) min buffer.getLength) // foreground color for { Text.Info(range, c) <- rendering.foreground(line_range) r <- JEdit_Lib.gfx_range(text_area, range) } { gfx.setColor(rendering.color(c)) gfx.fillRect(r.x, y + i * line_height, r.length, line_height) } // search pattern for { regex <- search_pattern text <- JEdit_Lib.get_text(buffer, line_range) m <- regex.findAllMatchIn(text) range = Text.Range(m.start, m.end) + line_range.start r <- JEdit_Lib.gfx_range(text_area, range) } { gfx.setColor(rendering.search_color) gfx.fillRect(r.x, y + i * line_height, r.length, line_height) } // highlight range -- potentially from other snapshot for { info <- highlight_area.info Text.Info(range, color) <- info.try_restrict(line_range) r <- JEdit_Lib.gfx_range(text_area, range) } { gfx.setColor(color) gfx.fillRect(r.x, y + i * line_height, r.length, line_height) } // hyperlink range -- potentially from other snapshot for { info <- hyperlink_area.info Text.Info(range, _) <- info.try_restrict(line_range) r <- JEdit_Lib.gfx_range(text_area, range) } { gfx.setColor(rendering.hyperlink_color) gfx.drawRect(r.x, y + i * line_height, r.length - 1, line_height - 1) } // entity def range if (!area_active() && caret_visible) { for { Text.Info(range, color) <- rendering.entity_ref(line_range, caret_focus) r <- JEdit_Lib.gfx_range(text_area, range) } { gfx.setColor(color) gfx.drawRect(r.x, y + i * line_height, r.length - 1, line_height - 1) } } // completion range if (!area_active() && caret_visible) { for { completion <- Completion_Popup.Text_Area(text_area) Text.Info(range, color) <- completion.rendering(rendering, line_range) r <- JEdit_Lib.gfx_range(text_area, range) } { gfx.setColor(color) gfx.drawRect(r.x, y + i * line_height, r.length - 1, line_height - 1) } } } } } } } /* caret -- outside of text range */ private class Caret_Painter(before: Boolean) extends TextAreaExtension { override def paintValidLine(gfx: Graphics2D, screen_line: Int, physical_line: Int, start: Int, end: Int, y: Int): Unit = { robust_rendering { _ => if (before) gfx.clipRect(0, 0, 0, 0) else gfx.setClip(painter_clip) } } } private val before_caret_painter1 = new Caret_Painter(true) private val after_caret_painter1 = new Caret_Painter(false) private val before_caret_painter2 = new Caret_Painter(true) private val after_caret_painter2 = new Caret_Painter(false) private val caret_painter = new TextAreaExtension { override def paintValidLine(gfx: Graphics2D, screen_line: Int, physical_line: Int, start: Int, end: Int, y: Int): Unit = { robust_rendering { rendering => if (caret_visible) { val caret = text_area.getCaretPosition if (caret_enabled && start <= caret && caret == end - 1) { val painter = text_area.getPainter val fm = painter.getFontMetrics val offset = caret - text_area.getLineStartOffset(physical_line) val x = text_area.offsetToXY(physical_line, offset).x val y1 = y + painter.getLineHeight - (fm.getLeading + 1) - fm.getDescent val astr = new AttributedString(" ") astr.addAttribute(TextAttribute.FONT, painter.getFont) astr.addAttribute(TextAttribute.FOREGROUND, caret_color(rendering, caret)) astr.addAttribute(TextAttribute.SWAP_COLORS, TextAttribute.SWAP_COLORS_ON) val clip = gfx.getClip try { gfx.clipRect(x, y, Integer.MAX_VALUE, painter.getLineHeight) gfx.drawString(astr.getIterator, x, y1) } finally { gfx.setClip(clip) } } } } } } /* activation */ def activate(): Unit = { val painter = text_area.getPainter painter.addExtension(TextAreaPainter.LOWEST_LAYER, set_state) painter.addExtension(TextAreaPainter.LINE_BACKGROUND_LAYER + 1, background_painter) painter.addExtension(TextAreaPainter.TEXT_LAYER, text_painter) painter.addExtension(TextAreaPainter.CARET_LAYER - 1, before_caret_painter1) painter.addExtension(TextAreaPainter.CARET_LAYER + 1, after_caret_painter1) painter.addExtension(TextAreaPainter.BLOCK_CARET_LAYER - 1, before_caret_painter2) painter.addExtension(TextAreaPainter.BLOCK_CARET_LAYER + 1, after_caret_painter2) painter.addExtension(TextAreaPainter.BLOCK_CARET_LAYER + 2, caret_painter) painter.addExtension(500, foreground_painter) painter.addExtension(TextAreaPainter.HIGHEST_LAYER, reset_state) painter.removeExtension(orig_text_painter) painter.addMouseListener(mouse_listener) painter.addMouseMotionListener(mouse_motion_listener) text_area.addKeyListener(key_listener) text_area.addFocusListener(focus_listener) view.addWindowListener(window_listener) } def deactivate(): Unit = { active_reset() val painter = text_area.getPainter view.removeWindowListener(window_listener) text_area.removeFocusListener(focus_listener) text_area.removeKeyListener(key_listener) painter.removeMouseMotionListener(mouse_motion_listener) painter.removeMouseListener(mouse_listener) painter.addExtension(TextAreaPainter.TEXT_LAYER, orig_text_painter) painter.removeExtension(reset_state) painter.removeExtension(foreground_painter) painter.removeExtension(caret_painter) painter.removeExtension(after_caret_painter2) painter.removeExtension(before_caret_painter2) painter.removeExtension(after_caret_painter1) painter.removeExtension(before_caret_painter1) painter.removeExtension(text_painter) painter.removeExtension(background_painter) painter.removeExtension(set_state) } } diff --git a/src/Tools/jEdit/src/scala_console.scala b/src/Tools/jEdit/src/scala_console.scala --- a/src/Tools/jEdit/src/scala_console.scala +++ b/src/Tools/jEdit/src/scala_console.scala @@ -1,181 +1,181 @@ /* Title: Tools/jEdit/src/scala_console.scala Author: Makarius Scala instance of Console plugin. */ package isabelle.jedit import isabelle._ import console.{Console, ConsolePane, Shell, Output} import org.gjt.sp.jedit.JARClassLoader import java.io.{OutputStream, Writer, PrintWriter} class Scala_Console extends Shell("Scala") { /* global state -- owned by GUI thread */ @volatile private var interpreters = Map.empty[Console, Interpreter] @volatile private var global_console: Console = null @volatile private var global_out: Output = null @volatile private var global_err: Output = null private val console_stream = new OutputStream { val buf = new StringBuilder(100) override def flush(): Unit = { - val s = buf.synchronized { val s = buf.toString; buf.clear; s } + val s = buf.synchronized { val s = buf.toString; buf.clear(); s } val str = UTF8.decode_permissive(s) GUI_Thread.later { if (global_out == null) System.out.print(str) else global_out.writeAttrs(null, str) } Time.seconds(0.01).sleep // FIXME adhoc delay to avoid loosing output } override def close(): Unit = flush() def write(byte: Int): Unit = { val c = byte.toChar buf.synchronized { buf.append(c) } if (c == '\n') flush() } } private val console_writer = new Writer { def flush(): Unit = console_stream.flush() def close(): Unit = console_stream.flush() def write(cbuf: Array[Char], off: Int, len: Int): Unit = { if (len > 0) { UTF8.bytes(new String(cbuf.slice(off, off + len))).foreach(console_stream.write(_)) } } } private def with_console[A](console: Console, out: Output, err: Output)(e: => A): A = { global_console = console global_out = out global_err = if (err == null) out else err try { scala.Console.withErr(console_stream) { scala.Console.withOut(console_stream) { e } } } finally { console_stream.flush global_console = null global_out = null global_err = null } } private def report_error(str: String): Unit = { if (global_console == null || global_err == null) isabelle.Output.writeln(str) else GUI_Thread.later { global_err.print(global_console.getErrorColor, str) } } /* interpreter thread */ private abstract class Request private case class Start(console: Console) extends Request private case class Execute(console: Console, out: Output, err: Output, command: String) extends Request private class Interpreter { private val running = Synchronized[Option[Thread]](None) - def interrupt: Unit = running.change(opt => { opt.foreach(_.interrupt); opt }) + def interrupt(): Unit = running.change(opt => { opt.foreach(_.interrupt()); opt }) private val interp = Scala.Compiler.context(error = report_error, jar_dirs = JEdit_Lib.directories). interpreter( print_writer = new PrintWriter(console_writer, true), class_loader = new JARClassLoader) val thread: Consumer_Thread[Request] = Consumer_Thread.fork("Scala_Console") { case Start(console) => interp.bind("view", "org.gjt.sp.jedit.View", console.getView) interp.bind("console", "console.Console", console) interp.interpret("import isabelle._; import isabelle.jedit._") true case Execute(console, out, err, command) => with_console(console, out, err) { try { running.change(_ => Some(Thread.currentThread())) interp.interpret(command) } finally { running.change(_ => None) Exn.Interrupt.dispose() } GUI_Thread.later { if (err != null) err.commandDone() out.commandDone() } true } } } /* jEdit console methods */ override def openConsole(console: Console): Unit = { val interp = new Interpreter interp.thread.send(Start(console)) interpreters += (console -> interp) } override def closeConsole(console: Console): Unit = { interpreters.get(console) match { case Some(interp) => - interp.interrupt - interp.thread.shutdown + interp.interrupt() + interp.thread.shutdown() interpreters -= console case None => } } override def printInfoMessage(out: Output): Unit = { out.print(null, "This shell evaluates Isabelle/Scala expressions.\n\n" + "The contents of package isabelle and isabelle.jedit are imported.\n" + "The following special toplevel bindings are provided:\n" + " view -- current jEdit/Swing view (e.g. view.getBuffer, view.getTextArea)\n" + " console -- jEdit Console plugin\n" + " PIDE -- Isabelle/PIDE plugin (e.g. PIDE.session, PIDE.snapshot, PIDE.rendering)\n") } override def printPrompt(console: Console, out: Output): Unit = { out.writeAttrs(ConsolePane.colorAttributes(console.getInfoColor), "scala>") out.writeAttrs(ConsolePane.colorAttributes(console.getPlainColor), " ") } override def execute( console: Console, input: String, out: Output, err: Output, command: String): Unit = { interpreters(console).thread.send(Execute(console, out, err, command)) } override def stop(console: Console): Unit = - interpreters.get(console).foreach(_.interrupt) + interpreters.get(console).foreach(_.interrupt()) } diff --git a/src/Tools/jEdit/src/session_build.scala b/src/Tools/jEdit/src/session_build.scala --- a/src/Tools/jEdit/src/session_build.scala +++ b/src/Tools/jEdit/src/session_build.scala @@ -1,183 +1,183 @@ /* Title: Tools/jEdit/src/session_build.scala Author: Makarius Session build management. */ package isabelle.jedit import isabelle._ import java.awt.event.{WindowEvent, WindowAdapter} import javax.swing.{WindowConstants, JDialog} import scala.swing.{ScrollPane, Button, CheckBox, FlowPanel, BorderPanel, TextArea, Component, Label} import scala.swing.event.ButtonClicked import org.gjt.sp.jedit.View object Session_Build { def check_dialog(view: View): Unit = { val options = PIDE.options.value Isabelle_Thread.fork() { try { if (JEdit_Sessions.session_no_build || JEdit_Sessions.session_build(options, no_build = true) == 0) JEdit_Sessions.session_start(options) else GUI_Thread.later { new Dialog(view) } } catch { case exn: Throwable => GUI.dialog(view, "Isabelle", GUI.scrollable_text(Exn.message(exn))) } } } private class Dialog(view: View) extends JDialog(view) { val options: Options = PIDE.options.value /* text */ private val text = new TextArea { editable = false columns = 60 rows = 24 } text.font = GUI.copy_font((new Label).font) private val scroll_text = new ScrollPane(text) /* progress */ private val progress = new Progress { override def echo(txt: String): Unit = GUI_Thread.later { text.append(txt + "\n") val vertical = scroll_text.peer.getVerticalScrollBar vertical.setValue(vertical.getMaximum) } override def theory(theory: Progress.Theory): Unit = echo(theory.message) } /* layout panel with dynamic actions */ private val action_panel = new FlowPanel(FlowPanel.Alignment.Center)() private val layout_panel = new BorderPanel layout_panel.layout(scroll_text) = BorderPanel.Position.Center layout_panel.layout(action_panel) = BorderPanel.Position.South setContentPane(layout_panel.peer) private def set_actions(cs: Component*): Unit = { - action_panel.contents.clear + action_panel.contents.clear() action_panel.contents ++= cs - layout_panel.revalidate - layout_panel.repaint + layout_panel.revalidate() + layout_panel.repaint() } /* return code and exit */ private var _return_code: Option[Int] = None private def return_code(rc: Int): Unit = GUI_Thread.later { _return_code = Some(rc) - delay_exit.invoke + delay_exit.invoke() } private val delay_exit = Delay.first(Time.seconds(1.0), gui = true) { if (can_auto_close) conclude() else { val button = new Button("Close") { reactions += { case ButtonClicked(_) => conclude() } } set_actions(button) button.peer.getRootPane.setDefaultButton(button.peer) } } private def conclude(): Unit = { setVisible(false) dispose() } /* close */ setDefaultCloseOperation(WindowConstants.DO_NOTHING_ON_CLOSE) addWindowListener(new WindowAdapter { override def windowClosing(e: WindowEvent): Unit = { if (_return_code.isDefined) conclude() else stopping() } }) private def stopping(): Unit = { - progress.stop + progress.stop() set_actions(new Label("Stopping ...")) } private val stop_button = new Button("Stop") { reactions += { case ButtonClicked(_) => stopping() } } private var do_auto_close = true private def can_auto_close: Boolean = do_auto_close && _return_code == Some(0) private val auto_close = new CheckBox("Auto close") { reactions += { case ButtonClicked(_) => do_auto_close = this.selected if (can_auto_close) conclude() } } auto_close.selected = do_auto_close auto_close.tooltip = "Automatically close dialog when finished" set_actions(stop_button, auto_close) /* main */ setTitle("Isabelle build (" + Isabelle_System.getenv("ML_IDENTIFIER") + " / " + "jdk-" + Platform.jvm_version + "_" + Platform.jvm_platform + ")") pack() setLocationRelativeTo(view) setVisible(true) Isabelle_Thread.fork(name = "session_build") { progress.echo("Build started for Isabelle/" + PIDE.resources.session_name + " ...") val (out, rc) = try { ("", JEdit_Sessions.session_build(options, progress = progress)) } catch { case exn: Throwable => (Output.error_message_text(Exn.message(exn)) + "\n", Exn.return_code(exn, 2)) } progress.echo(out + (if (rc == 0) "OK" else Process_Result.print_return_code(rc)) + "\n") if (rc == 0) JEdit_Sessions.session_start(options) else progress.echo("Session build failed -- prover process remains inactive!") return_code(rc) } } } diff --git a/src/Tools/jEdit/src/sledgehammer_dockable.scala b/src/Tools/jEdit/src/sledgehammer_dockable.scala --- a/src/Tools/jEdit/src/sledgehammer_dockable.scala +++ b/src/Tools/jEdit/src/sledgehammer_dockable.scala @@ -1,172 +1,172 @@ /* Title: Tools/jEdit/src/sledgehammer_dockable.scala Author: Makarius Dockable window for Sledgehammer. */ package isabelle.jedit import isabelle._ import isabelle.jedit_base.Dockable import scala.swing.{Button, Component, Label, CheckBox} import scala.swing.event.ButtonClicked import java.awt.BorderLayout import java.awt.event.{ComponentEvent, ComponentAdapter, KeyEvent} import org.gjt.sp.jedit.View import org.gjt.sp.jedit.gui.HistoryTextField class Sledgehammer_Dockable(view: View, position: String) extends Dockable(view, position) { GUI_Thread.require {} /* text area */ val pretty_text_area = new Pretty_Text_Area(view) set_content(pretty_text_area) override def detach_operation: Option[() => Unit] = pretty_text_area.detach_operation /* query operation */ private val process_indicator = new Process_Indicator private def consume_status(status: Query_Operation.Status.Value): Unit = { status match { case Query_Operation.Status.WAITING => process_indicator.update("Waiting for evaluation of context ...", 5) case Query_Operation.Status.RUNNING => process_indicator.update("Sledgehammering ...", 15) case Query_Operation.Status.FINISHED => process_indicator.update(null, 0) } } private val sledgehammer = new Query_Operation(PIDE.editor, view, "sledgehammer", consume_status, (snapshot, results, body) => pretty_text_area.update(snapshot, results, Pretty.separate(body))) /* resize */ private val delay_resize = Delay.first(PIDE.options.seconds("editor_update_delay"), gui = true) { handle_resize() } addComponentListener(new ComponentAdapter { override def componentResized(e: ComponentEvent): Unit = delay_resize.invoke() override def componentShown(e: ComponentEvent): Unit = delay_resize.invoke() }) private def handle_resize(): Unit = { GUI_Thread.require {} pretty_text_area.resize( Font_Info.main(PIDE.options.real("jedit_font_scale") * zoom.factor / 100)) } /* controls */ private def clicked: Unit = { provers.addCurrentToHistory() PIDE.options.string("sledgehammer_provers") = provers.getText sledgehammer.apply_query( List(provers.getText, isar_proofs.selected.toString, try0.selected.toString)) } private val provers_label = new Label("Provers:") { tooltip = GUI.tooltip_lines( "Automatic provers as space-separated list, e.g.\n" + PIDE.options.value.check_name("sledgehammer_provers").default_value) } private val provers = new HistoryTextField("isabelle-sledgehammer-provers") { override def processKeyEvent(evt: KeyEvent): Unit = { if (evt.getID == KeyEvent.KEY_PRESSED && evt.getKeyCode == KeyEvent.VK_ENTER) clicked super.processKeyEvent(evt) } setToolTipText(provers_label.tooltip) setColumns(30) } private def update_provers(): Unit = { val new_provers = PIDE.options.string("sledgehammer_provers") if (new_provers != provers.getText) { provers.setText(new_provers) if (provers.getCaret != null) provers.getCaret.setDot(0) } } private val isar_proofs = new CheckBox("Isar proofs") { tooltip = "Specify whether Isar proofs should be output in addition to \"by\" one-liner" selected = false } private val try0 = new CheckBox("Try methods") { tooltip = "Try standard proof methods like \"auto\" and \"blast\" as alternatives to \"metis\"" selected = true } private val apply_query = new Button("Apply") { tooltip = "Search for first-order proof using automatic theorem provers" reactions += { case ButtonClicked(_) => clicked } } private val cancel_query = new Button("Cancel") { tooltip = "Interrupt unfinished sledgehammering" reactions += { case ButtonClicked(_) => sledgehammer.cancel_query() } } private val locate_query = new Button("Locate") { tooltip = "Locate context of current query within source text" reactions += { case ButtonClicked(_) => sledgehammer.locate_query() } } private val zoom = new Font_Info.Zoom_Box { def changed = handle_resize() } private val controls = Wrap_Panel( List(provers_label, Component.wrap(provers), isar_proofs, try0, process_indicator.component, apply_query, cancel_query, locate_query, zoom)) add(controls.peer, BorderLayout.NORTH) - override def focusOnDefaultComponent(): Unit = provers.requestFocus + override def focusOnDefaultComponent(): Unit = provers.requestFocus() /* main */ private val main = Session.Consumer[Session.Global_Options](getClass.getName) { case _: Session.Global_Options => GUI_Thread.later { update_provers(); handle_resize() } } override def init(): Unit = { PIDE.session.global_options += main update_provers() handle_resize() sledgehammer.activate() } override def exit(): Unit = { sledgehammer.deactivate() PIDE.session.global_options -= main delay_resize.revoke() } } diff --git a/src/Tools/jEdit/src/symbols_dockable.scala b/src/Tools/jEdit/src/symbols_dockable.scala --- a/src/Tools/jEdit/src/symbols_dockable.scala +++ b/src/Tools/jEdit/src/symbols_dockable.scala @@ -1,222 +1,222 @@ /* Title: Tools/jEdit/src/symbols_dockable.scala Author: Fabian Immler Dockable window for Symbol Palette. */ package isabelle.jedit import isabelle._ import isabelle.jedit_base.Dockable import scala.swing.event.{SelectionChanged, ValueChanged} import scala.swing.{Component, Action, Button, TabbedPane, TextField, BorderPanel, Label, ScrollPane} import org.gjt.sp.jedit.{EditBus, EBComponent, EBMessage, View} class Symbols_Dockable(view: View, position: String) extends Dockable(view, position) { private def font_size: Int = Font_Info.main_size(PIDE.options.real("jedit_font_scale")).round /* abbrevs */ private val abbrevs_panel = new Abbrevs_Panel private val abbrevs_refresh_delay = Delay.last(PIDE.options.seconds("editor_update_delay"), gui = true) { abbrevs_panel.refresh } private class Abbrev_Component(txt: String, abbrs: List[String]) extends Button { def update_font: Unit = { font = GUI.font(size = font_size) } update_font text = "" + HTML.output(Symbol.decode(txt)) + "" action = Action(text) { val text_area = view.getTextArea val (s1, s2) = Completion.split_template(Isabelle_Encoding.perhaps_decode(text_area.getBuffer, txt)) text_area.setSelectedText(s1 + s2) text_area.moveCaretPosition(text_area.getCaretPosition - s2.length) - text_area.requestFocus + text_area.requestFocus() } tooltip = GUI.tooltip_lines(cat_lines(txt :: abbrs.map(a => "abbrev: " + a))) } private class Abbrevs_Panel extends Wrap_Panel(Nil, Wrap_Panel.Alignment.Center) { private var abbrevs: Thy_Header.Abbrevs = Nil def refresh: Unit = GUI_Thread.require { val abbrevs1 = Isabelle.buffer_syntax(view.getBuffer).getOrElse(Outer_Syntax.empty).abbrevs if (abbrevs != abbrevs1) { abbrevs = abbrevs1 val entries: List[(String, List[String])] = Multi_Map( (for { (abbr, txt0) <- abbrevs txt = Symbol.encode(txt0) if !Symbol.iterator(txt). forall(s => s.length == 1 && s(0) != Completion.caret_indicator) } yield (txt, abbr)): _*).iterator_list.toList - contents.clear + contents.clear() for ((txt, abbrs) <- entries.sortBy(_._1)) contents += new Abbrev_Component(txt, abbrs.filterNot(_ == "").sorted) - revalidate - repaint + revalidate() + repaint() } } refresh } /* symbols */ private class Symbol_Component(val symbol: String, is_control: Boolean) extends Button { def update_font: Unit = { font = Symbol.fonts.get(symbol) match { case None => GUI.font(size = font_size) case Some(font_family) => GUI.font(family = font_family, size = font_size) } } update_font action = Action(Symbol.decode(symbol)) { val text_area = view.getTextArea if (is_control && HTML.is_control(symbol)) Syntax_Style.edit_control_style(text_area, symbol) else text_area.setSelectedText(Isabelle_Encoding.perhaps_decode(text_area.getBuffer, symbol)) - text_area.requestFocus + text_area.requestFocus() } tooltip = GUI.tooltip_lines( cat_lines(symbol :: Symbol.abbrevs.get_list(symbol).map(a => "abbrev: " + a))) } private class Reset_Component extends Button { action = Action("Reset") { val text_area = view.getTextArea Syntax_Style.edit_control_style(text_area, "") - text_area.requestFocus + text_area.requestFocus() } tooltip = "Reset control symbols within text" } /* search */ private class Search_Panel extends BorderPanel { val search_field = new TextField(10) val results_panel = Wrap_Panel(Nil, Wrap_Panel.Alignment.Center) layout(search_field) = BorderPanel.Position.North layout(new ScrollPane(results_panel)) = BorderPanel.Position.Center val search_space = for ((sym, _) <- Symbol.codes) yield (sym, Word.lowercase(sym)) val search_delay = Delay.last(PIDE.options.seconds("editor_input_delay"), gui = true) { val search_words = Word.explode(Word.lowercase(search_field.text)) val search_limit = PIDE.options.int("jedit_symbols_search_limit") max 0 val results = for ((sym, s) <- search_space; if search_words.forall(s.contains(_))) yield sym - results_panel.contents.clear + results_panel.contents.clear() for (sym <- results.take(search_limit)) results_panel.contents += new Symbol_Component(sym, false) if (results.length > search_limit) results_panel.contents += new Label("...") { tooltip = "(" + (results.length - search_limit) + " more)" } - revalidate - repaint + revalidate() + repaint() } search_field.reactions += { case ValueChanged(_) => search_delay.invoke() } } /* tabs */ private val group_tabs: TabbedPane = new TabbedPane { pages += new TabbedPane.Page("abbrevs", abbrevs_panel) pages ++= Symbol.groups_code.map({ case (group, symbols) => val control = group == "control" new TabbedPane.Page(group, new ScrollPane(Wrap_Panel( symbols.map(new Symbol_Component(_, control)) ::: (if (control) List(new Reset_Component) else Nil), Wrap_Panel.Alignment.Center)), null) }) val search_panel = new Search_Panel val search_page = new TabbedPane.Page("search", search_panel, "Search Symbols") pages += search_page listenTo(selection) reactions += { case SelectionChanged(_) if selection.page == search_page => - search_panel.search_field.requestFocus + search_panel.search_field.requestFocus() } for (page <- pages) page.title = Word.implode(Word.explode('_', page.title).map(Word.perhaps_capitalize)) } set_content(group_tabs) /* main */ private val edit_bus_handler: EBComponent = new EBComponent { def handleMessage(msg: EBMessage): Unit = abbrevs_refresh_delay.invoke() } private val main = Session.Consumer[Any](getClass.getName) { case _: Session.Global_Options => GUI_Thread.later { val comp = group_tabs.peer GUI.traverse_components(comp, { case c0: javax.swing.JComponent => Component.wrap(c0) match { case c: Abbrev_Component => c.update_font case c: Symbol_Component => c.update_font case _ => } case _ => }) comp.revalidate() comp.repaint() } case _: Session.Commands_Changed => abbrevs_refresh_delay.invoke() } override def init(): Unit = { EditBus.addToBus(edit_bus_handler) PIDE.session.global_options += main PIDE.session.commands_changed += main } override def exit(): Unit = { EditBus.removeFromBus(edit_bus_handler) PIDE.session.global_options -= main PIDE.session.commands_changed -= main } } diff --git a/src/Tools/jEdit/src/text_overview.scala b/src/Tools/jEdit/src/text_overview.scala --- a/src/Tools/jEdit/src/text_overview.scala +++ b/src/Tools/jEdit/src/text_overview.scala @@ -1,187 +1,187 @@ /* Title: Tools/jEdit/src/text_overview.scala Author: Makarius GUI component for text status overview. */ package isabelle.jedit import isabelle._ import scala.annotation.tailrec import java.awt.{Graphics, Graphics2D, BorderLayout, Dimension, Color} import java.awt.event.{MouseAdapter, MouseEvent} import javax.swing.{JPanel, ToolTipManager} class Text_Overview(doc_view: Document_View) extends JPanel(new BorderLayout) { /* GUI components */ private val text_area = doc_view.text_area private val buffer = doc_view.model.buffer private def lines(): Int = buffer.getLineCount max text_area.getVisibleLines private val WIDTH = text_area.getPainter.getFontMetrics.stringWidth("A") max 10 private val HEIGHT = 4 setPreferredSize(new Dimension(WIDTH, 0)) setRequestFocusEnabled(false) addMouseListener(new MouseAdapter { override def mousePressed(event: MouseEvent): Unit = { val line = (event.getY * lines()) / getHeight if (line >= 0 && line < text_area.getLineCount) text_area.setCaretPosition(text_area.getLineStartOffset(line)) } }) override def addNotify(): Unit = { super.addNotify() ToolTipManager.sharedInstance.registerComponent(this) } override def removeNotify(): Unit = { ToolTipManager.sharedInstance.unregisterComponent(this) super.removeNotify } /* overview */ private case class Overview( line_count: Int = 0, char_count: Int = 0, L: Int = 0, H: Int = 0) private def get_overview(): Overview = Overview( line_count = buffer.getLineCount, char_count = buffer.getLength, L = lines(), H = getHeight()) /* synchronous painting */ type Color_Info = (Rendering.Color.Value, Int, Int) private var current_colors: List[Color_Info] = Nil private var current_overview = Overview() private val delay_repaint = Delay.first(PIDE.options.seconds("editor_update_delay"), gui = true) { repaint() } override def paintComponent(gfx: Graphics): Unit = { super.paintComponent(gfx) GUI_Thread.assert {} doc_view.rich_text_area.robust_body(()) { JEdit_Lib.buffer_lock(buffer) { val rendering = doc_view.get_rendering() val overview = get_overview() if (rendering.snapshot.is_outdated || overview != current_overview) { invoke() delay_repaint.invoke() gfx.setColor(rendering.outdated_color) gfx.asInstanceOf[Graphics2D].fill(gfx.getClipBounds) } else { gfx.setColor(getBackground) gfx.asInstanceOf[Graphics2D].fill(gfx.getClipBounds) for ((color, h, h1) <- current_colors) { gfx.setColor(rendering.color(color)) gfx.fillRect(0, h, getWidth, h1 - h) } } } } } /* asynchronous refresh */ private val future_refresh = Synchronized(Future.value(())) private def is_running(): Boolean = !future_refresh.value.is_finished def invoke(): Unit = delay_refresh.invoke() - def revoke(): Unit = { delay_refresh.revoke(); future_refresh.change(x => { x.cancel; x }) } + def revoke(): Unit = { delay_refresh.revoke(); future_refresh.change(x => { x.cancel(); x }) } private val delay_refresh = Delay.first(PIDE.options.seconds("editor_update_delay"), gui = true) { if (!doc_view.rich_text_area.check_robust_body) invoke() else { JEdit_Lib.buffer_lock(buffer) { val rendering = doc_view.get_rendering() val overview = get_overview() if (rendering.snapshot.is_outdated || is_running()) invoke() else { val line_offsets = { val line_manager = JEdit_Lib.buffer_line_manager(buffer) val a = new Array[Int](line_manager.getLineCount) for (i <- 1 until a.length) a(i) = line_manager.getLineEndOffset(i - 1) a } future_refresh.change(_ => Future.fork { val line_count = overview.line_count val char_count = overview.char_count val L = overview.L val H = overview.H @tailrec def loop(l: Int, h: Int, p: Int, q: Int, colors: List[Color_Info]) : List[Color_Info] = { Exn.Interrupt.expose() if (l < line_count && h < H) { val p1 = p + H val q1 = q + HEIGHT * L val (l1, h1) = if (p1 >= q1) (l + 1, h + (p1 - q) / L) else (l + (q1 - p) / H, h + HEIGHT) val start = line_offsets(l) val end = if (l1 < line_count) line_offsets(l1) else char_count val colors1 = (rendering.overview_color(Text.Range(start, end)), colors) match { case (Some(color), (old_color, old_h, old_h1) :: rest) if color == old_color && old_h1 == h => (color, old_h, h1) :: rest case (Some(color), _) => (color, h, h1) :: colors case (None, _) => colors } loop(l1, h1, p + (l1 - l) * H, q + (h1 - h) * L, colors1) } else colors.reverse } val new_colors = loop(0, 0, 0, 0, Nil) GUI_Thread.later { current_overview = overview current_colors = new_colors repaint() } } ) } } } } } diff --git a/src/Tools/jEdit/src/text_structure.scala b/src/Tools/jEdit/src/text_structure.scala --- a/src/Tools/jEdit/src/text_structure.scala +++ b/src/Tools/jEdit/src/text_structure.scala @@ -1,355 +1,355 @@ /* Title: Tools/jEdit/src/text_structure.scala Author: Makarius Text structure based on Isabelle/Isar outer syntax. */ package isabelle.jedit import isabelle._ import org.gjt.sp.jedit.indent.{IndentRule, IndentAction} import org.gjt.sp.jedit.textarea.{TextArea, StructureMatcher, Selection} import org.gjt.sp.jedit.buffer.JEditBuffer import org.gjt.sp.jedit.Buffer object Text_Structure { /* token navigator */ class Navigator(syntax: Outer_Syntax, buffer: JEditBuffer, comments: Boolean) { val limit: Int = PIDE.options.value.int("jedit_structure_limit") max 0 def iterator(line: Int, lim: Int = limit): Iterator[Text.Info[Token]] = { val it = Token_Markup.line_token_iterator(syntax, buffer, line, line + lim) if (comments) it.filterNot(_.info.is_space) else it.filter(_.info.is_proper) } def reverse_iterator(line: Int, lim: Int = limit): Iterator[Text.Info[Token]] = { val it = Token_Markup.line_token_reverse_iterator(syntax, buffer, line, line - lim) if (comments) it.filterNot(_.info.is_space) else it.filter(_.info.is_proper) } } /* indentation */ object Indent_Rule extends IndentRule { private val keyword_open = Keyword.theory_goal ++ Keyword.proof_open private val keyword_close = Keyword.proof_close def apply(buffer: JEditBuffer, current_line: Int, prev_line0: Int, prev_prev_line0: Int, actions: java.util.List[IndentAction]): Unit = { Isabelle.buffer_syntax(buffer) match { case Some(syntax) => val keywords = syntax.keywords val nav = new Navigator(syntax, buffer, true) val indent_size = buffer.getIndentSize def line_indent(line: Int): Int = if (line < 0 || line >= buffer.getLineCount) 0 else buffer.getCurrentIndentForLine(line, null) def line_head(line: Int): Option[Text.Info[Token]] = nav.iterator(line, 1).nextOption() def head_is_quasi_command(line: Int): Boolean = line_head(line) match { case None => false case Some(Text.Info(_, tok)) => keywords.is_quasi_command(tok) } val prev_line: Int = Range.inclusive(current_line - 1, 0, -1).find(line => Token_Markup.Line_Context.before(buffer, line).get_context == Scan.Finished && (!Token_Markup.Line_Context.after(buffer, line).structure.improper || Token_Markup.Line_Context.after(buffer, line).structure.blank)) getOrElse -1 def prev_line_command: Option[Token] = nav.reverse_iterator(prev_line, 1). collectFirst({ case Text.Info(_, tok) if tok.is_begin_or_command => tok }) def prev_line_span: Iterator[Token] = nav.reverse_iterator(prev_line, 1).map(_.info).takeWhile(tok => !tok.is_begin_or_command) def prev_span: Iterator[Token] = nav.reverse_iterator(prev_line).map(_.info).takeWhile(tok => !tok.is_begin_or_command) val script_indent: Text.Info[Token] => Int = { val opt_rendering: Option[JEdit_Rendering] = if (PIDE.options.value.bool("jedit_indent_script")) GUI_Thread.now { (for { text_area <- JEdit_Lib.jedit_text_areas(buffer) doc_view <- Document_View.get(text_area) - } yield doc_view.get_rendering).nextOption() + } yield doc_view.get_rendering()).nextOption() } else None val limit = PIDE.options.value.int("jedit_indent_script_limit") (info: Text.Info[Token]) => opt_rendering match { case Some(rendering) if keywords.is_command(info.info, Keyword.prf_script) => (rendering.indentation(info.range) min limit) max 0 case _ => 0 } } def indent_indent(tok: Token): Int = if (keywords.is_command(tok, keyword_open)) indent_size else if (keywords.is_command(tok, keyword_close)) - indent_size else 0 def indent_offset(tok: Token): Int = if (keywords.is_command(tok, Keyword.proof_enclose)) indent_size else 0 def indent_structure: Int = nav.reverse_iterator(current_line - 1).scanLeft((0, false))( { case ((ind, _), Text.Info(range, tok)) => val ind1 = ind + indent_indent(tok) if (tok.is_begin_or_command && !keywords.is_command(tok, Keyword.prf_script)) { val line = buffer.getLineOfOffset(range.start) line_head(line) match { case Some(info) if info.info == tok => (ind1 + indent_offset(tok) + line_indent(line), true) case _ => (ind1, false) } } else (ind1, false) }).collectFirst({ case (i, true) => i }).getOrElse(0) def indent_brackets: Int = prev_line_span.foldLeft(0) { case (i, tok) => if (tok.is_open_bracket) i + indent_size else if (tok.is_close_bracket) i - indent_size else i } def indent_extra: Int = if (prev_span.exists(keywords.is_quasi_command)) indent_size else 0 val indent = if (Token_Markup.Line_Context.before(buffer, current_line).get_context != Scan.Finished) line_indent(current_line) else if (Token_Markup.Line_Context.after(buffer, current_line).structure.blank) 0 else { line_head(current_line) match { case Some(info) => val tok = info.info if (tok.is_begin || keywords.is_before_command(tok) || keywords.is_command(tok, Keyword.theory)) 0 else if (keywords.is_command(tok, Keyword.proof_enclose)) indent_structure + script_indent(info) - indent_offset(tok) else if (keywords.is_command(tok, Keyword.proof)) (indent_structure + script_indent(info) - indent_offset(tok)) max indent_size else if (tok.is_command) indent_structure - indent_offset(tok) else { prev_line_command match { case None => val extra = (keywords.is_quasi_command(tok), head_is_quasi_command(prev_line)) match { case (true, true) | (false, false) => 0 case (true, false) => - indent_extra case (false, true) => indent_extra } line_indent(prev_line) + indent_brackets + extra - indent_offset(tok) case Some(prev_tok) => indent_structure + indent_brackets + indent_size - indent_offset(tok) - indent_offset(prev_tok) - indent_indent(prev_tok) } } case None => prev_line_command match { case None => val extra = if (head_is_quasi_command(prev_line)) indent_extra else 0 line_indent(prev_line) + indent_brackets + extra case Some(prev_tok) => indent_structure + indent_brackets + indent_size - indent_offset(prev_tok) - indent_indent(prev_tok) } } } actions.clear() actions.add(new IndentAction.AlignOffset(indent max 0)) case None => } } } def line_content(buffer: JEditBuffer, keywords: Keyword.Keywords, range: Text.Range, ctxt: Scan.Line_Context): (List[Token], Scan.Line_Context) = { val text = JEdit_Lib.get_text(buffer, range).getOrElse("") val (toks, ctxt1) = Token.explode_line(keywords, text, ctxt) val toks1 = toks.filterNot(_.is_space) (toks1, ctxt1) } def split_line_content(buffer: JEditBuffer, keywords: Keyword.Keywords, line: Int, caret: Int) : (List[Token], List[Token]) = { val line_range = JEdit_Lib.line_range(buffer, line) val ctxt0 = Token_Markup.Line_Context.before(buffer, line).get_context val (toks1, ctxt1) = line_content(buffer, keywords, Text.Range(line_range.start, caret), ctxt0) val (toks2, _) = line_content(buffer, keywords, Text.Range(caret, line_range.stop), ctxt1) (toks1, toks2) } /* structure matching */ object Matcher extends StructureMatcher { private def find_block( open: Token => Boolean, close: Token => Boolean, reset: Token => Boolean, restrict: Token => Boolean, it: Iterator[Text.Info[Token]]): Option[(Text.Range, Text.Range)] = { val range1 = it.next().range it.takeWhile(info => !info.info.is_command || restrict(info.info)). scanLeft((range1, 1))( { case ((r, d), Text.Info(range, tok)) => if (open(tok)) (range, d + 1) else if (close(tok)) (range, d - 1) else if (reset(tok)) (range, 0) else (r, d) } ).collectFirst({ case (range2, 0) => (range1, range2) }) } private def find_pair(text_area: TextArea): Option[(Text.Range, Text.Range)] = { val buffer = text_area.getBuffer val caret_line = text_area.getCaretLine val caret = text_area.getCaretPosition Isabelle.buffer_syntax(text_area.getBuffer) match { case Some(syntax) => val keywords = syntax.keywords val nav = new Navigator(syntax, buffer, false) def caret_iterator(): Iterator[Text.Info[Token]] = nav.iterator(caret_line).dropWhile(info => !info.range.touches(caret)) def reverse_caret_iterator(): Iterator[Text.Info[Token]] = nav.reverse_iterator(caret_line).dropWhile(info => !info.range.touches(caret)) nav.iterator(caret_line, 1).find(info => info.range.touches(caret)) match { case Some(Text.Info(range1, tok)) if keywords.is_command(tok, Keyword.theory_goal) => find_block( keywords.is_command(_, Keyword.proof_goal), keywords.is_command(_, Keyword.qed), keywords.is_command(_, Keyword.qed_global), t => keywords.is_command(t, Keyword.diag) || keywords.is_command(t, Keyword.proof), caret_iterator()) case Some(Text.Info(range1, tok)) if keywords.is_command(tok, Keyword.proof_goal) => find_block( keywords.is_command(_, Keyword.proof_goal), keywords.is_command(_, Keyword.qed), _ => false, t => keywords.is_command(t, Keyword.diag) || keywords.is_command(t, Keyword.proof), caret_iterator()) case Some(Text.Info(range1, tok)) if keywords.is_command(tok, Keyword.qed_global) => reverse_caret_iterator().find(info => keywords.is_command(info.info, Keyword.theory)) match { case Some(Text.Info(range2, tok)) if keywords.is_command(tok, Keyword.theory_goal) => Some((range1, range2)) case _ => None } case Some(Text.Info(range1, tok)) if keywords.is_command(tok, Keyword.qed) => find_block( keywords.is_command(_, Keyword.qed), t => keywords.is_command(t, Keyword.proof_goal) || keywords.is_command(t, Keyword.theory_goal), _ => false, t => keywords.is_command(t, Keyword.diag) || keywords.is_command(t, Keyword.proof) || keywords.is_command(t, Keyword.theory_goal), reverse_caret_iterator()) case Some(Text.Info(range1, tok)) if tok.is_begin => find_block(_.is_begin, _.is_end, _ => false, _ => true, caret_iterator()) case Some(Text.Info(range1, tok)) if tok.is_end => find_block(_.is_end, _.is_begin, _ => false, _ => true, reverse_caret_iterator()) match { case Some((_, range2)) => reverse_caret_iterator(). dropWhile(info => info.range != range2). dropWhile(info => info.range == range2). find(info => info.info.is_command || info.info.is_begin) match { case Some(Text.Info(range3, tok)) => if (keywords.is_command(tok, Keyword.theory_block)) Some((range1, range3)) else Some((range1, range2)) case None => None } case None => None } case _ => None } case None => None } } def getMatch(text_area: TextArea): StructureMatcher.Match = find_pair(text_area) match { case Some((_, range)) => val line = text_area.getBuffer.getLineOfOffset(range.start) new StructureMatcher.Match(Matcher, line, range.start, line, range.stop) case None => null } def selectMatch(text_area: TextArea): Unit = { def get_span(offset: Text.Offset): Option[Text.Range] = for { syntax <- Isabelle.buffer_syntax(text_area.getBuffer) span <- Token_Markup.command_span(syntax, text_area.getBuffer, offset) } yield span.range find_pair(text_area) match { case Some((r1, r2)) => (get_span(r1.start), get_span(r2.start)) match { case (Some(range1), Some(range2)) => val start = range1.start min range2.start val stop = range1.stop max range2.stop text_area.moveCaretPosition(stop, false) if (!text_area.isMultipleSelectionEnabled) text_area.selectNone text_area.addToSelection(new Selection.Range(start, stop)) case _ => } case None => } } } }