# HG changeset patch # User wenzelm # Date 1355606354 -3600 # Node ID e4dc37ec14273050942622ddd862bf45997e1826 # Parent 89c0d2f13ccad3d8bfa0563e7e1f7f4a4a07ddbf# Parent 0493efcc97e9a7494c9c4b176b0d84233e5fad87 merged diff -r 89c0d2f13cca -r e4dc37ec1427 src/Pure/PIDE/markup_tree.scala --- a/src/Pure/PIDE/markup_tree.scala Sat Dec 15 21:34:32 2012 +0100 +++ b/src/Pure/PIDE/markup_tree.scala Sat Dec 15 22:19:14 2012 +0100 @@ -18,6 +18,8 @@ object Markup_Tree { + /* construct trees */ + val empty: Markup_Tree = new Markup_Tree(Branches.empty) def merge_disjoint(trees: List[Markup_Tree]): Markup_Tree = @@ -35,26 +37,56 @@ }) } + + /* tree building blocks */ + + object Elements + { + val empty = new Elements(Set.empty) + } + + final class Elements private(private val rep: Set[String]) + { + def contains(name: String): Boolean = rep.contains(name) + + def + (name: String): Elements = + if (contains(name)) this + else new Elements(rep + name) + + def + (elem: XML.Elem): Elements = this + elem.markup.name + def ++ (elems: Iterable[XML.Elem]): Elements = (this /: elems.iterator)(_ + _) + + def ++ (other: Elements): Elements = + if (this eq other) this + else if (rep.isEmpty) other + else (this /: other.rep)(_ + _) + } + object Entry { def apply(markup: Text.Markup, subtree: Markup_Tree): Entry = - Entry(markup.range, List(markup.info), Set(markup.info.markup.name), subtree) + Entry(markup.range, List(markup.info), Elements.empty + markup.info, + subtree, subtree.make_elements) def apply(range: Text.Range, rev_markups: List[XML.Elem], subtree: Markup_Tree): Entry = - Entry(range, rev_markups, Set.empty ++ rev_markups.iterator.map(_.markup.name), subtree) + Entry(range, rev_markups, Elements.empty ++ rev_markups, + subtree, subtree.make_elements) } sealed case class Entry( range: Text.Range, rev_markup: List[XML.Elem], - elements: Set[String], - subtree: Markup_Tree) + elements: Elements, + subtree: Markup_Tree, + subtree_elements: Elements) { - def + (info: XML.Elem): Entry = - if (elements(info.markup.name)) copy(rev_markup = info :: rev_markup) - else copy(rev_markup = info :: rev_markup, elements = elements + info.markup.name) + def markup: List[XML.Elem] = rev_markup.reverse - def markup: List[XML.Elem] = rev_markup.reverse + def + (markup: Text.Markup): Entry = + copy(rev_markup = markup.info :: rev_markup, elements = elements + markup.info) + + def \ (markup: Text.Markup): Entry = + copy(subtree = subtree + markup, subtree_elements = subtree_elements + markup.info) } object Branches @@ -121,6 +153,10 @@ } } + def make_elements: Elements = + (Elements.empty /: branches)( + { case (elements, (_, entry)) => elements ++ entry.subtree_elements ++ entry.elements }) + def + (new_markup: Text.Markup): Markup_Tree = { val new_range = new_markup.range @@ -129,9 +165,9 @@ case None => new Markup_Tree(branches, Entry(new_markup, empty)) case Some(entry) => if (entry.range == new_range) - new Markup_Tree(branches, entry + new_markup.info) + new Markup_Tree(branches, entry + new_markup) else if (entry.range.contains(new_range)) - new Markup_Tree(branches, entry.copy(subtree = entry.subtree + new_markup)) + new Markup_Tree(branches, entry \ new_markup) else if (new_range.contains(branches.head._1) && new_range.contains(branches.last._1)) new Markup_Tree(Branches.empty, Entry(new_markup, this)) else { @@ -184,19 +220,25 @@ def cumulate[A](root_range: Text.Range, root_info: A, result_elements: Option[Set[String]], result: PartialFunction[(A, Text.Markup), A]): Stream[Text.Info[A]] = { + val notable: Elements => Boolean = + result_elements match { + case Some(res) => (elements: Elements) => res.exists(elements.contains) + case None => (elements: Elements) => true + } + def results(x: A, entry: Entry): Option[A] = - if (result_elements match { case Some(es) => es.exists(entry.elements) case None => true }) { - val (y, changed) = - ((x, false) /: entry.rev_markup)((res, info) => // FIXME proper order!? - { - val (y, changed) = res - val arg = (y, Text.Info(entry.range, info)) - if (result.isDefinedAt(arg)) (result(arg), true) - else res - }) - if (changed) Some(y) else None - } - else None + { + val (y, changed) = + // FIXME proper cumulation order (including status markup) (!?) + ((x, false) /: entry.rev_markup)((res, info) => + { + val (y, changed) = res + val arg = (y, Text.Info(entry.range, info)) + if (result.isDefinedAt(arg)) (result(arg), true) + else res + }) + if (changed) Some(y) else None + } def stream( last: Text.Offset, @@ -205,10 +247,13 @@ stack match { case (parent, (range, entry) #:: more) :: rest => val subrange = range.restrict(root_range) - val subtree = entry.subtree.overlapping(subrange).toStream + val subtree = + if (notable(entry.subtree_elements)) + entry.subtree.overlapping(subrange).toStream + else Stream.empty val start = subrange.start - results(parent.info, entry) match { + (if (notable(entry.elements)) results(parent.info, entry) else None) match { case Some(res) => val next = Text.Info(subrange, res) val nexts = stream(start, (next, subtree) :: (parent, more) :: rest) diff -r 89c0d2f13cca -r e4dc37ec1427 src/Tools/jEdit/etc/options --- a/src/Tools/jEdit/etc/options Sat Dec 15 21:34:32 2012 +0100 +++ b/src/Tools/jEdit/etc/options Sat Dec 15 22:19:14 2012 +0100 @@ -12,6 +12,9 @@ option jedit_tooltip_margin : int = 60 -- "margin for tooltip pretty-printing" +option jedit_tooltip_bounds : real = 0.5 + -- "relative bounds of tooltip window wrt. logical screen size" + option jedit_text_overview_limit : int = 100000 -- "maximum amount of text to visualize in overview column" diff -r 89c0d2f13cca -r e4dc37ec1427 src/Tools/jEdit/src/isabelle_options.scala --- a/src/Tools/jEdit/src/isabelle_options.scala Sat Dec 15 21:34:32 2012 +0100 +++ b/src/Tools/jEdit/src/isabelle_options.scala Sat Dec 15 22:19:14 2012 +0100 @@ -41,8 +41,8 @@ { // FIXME avoid hard-wired stuff private val relevant_options = - Set("jedit_logic", "jedit_font_scale", "jedit_text_overview_limit", - "jedit_tooltip_font_scale", "jedit_symbols_search_limit", "jedit_tooltip_margin", + Set("jedit_logic", "jedit_font_scale", "jedit_symbols_search_limit", "jedit_text_overview_limit", + "jedit_tooltip_bounds", "jedit_tooltip_font_scale", "jedit_tooltip_margin", "threads", "threads_trace", "parallel_proofs", "parallel_proofs_threshold", "ML_statistics", "editor_load_delay", "editor_input_delay", "editor_output_delay", "editor_reparse_limit", "editor_tracing_messages", "editor_update_delay") diff -r 89c0d2f13cca -r e4dc37ec1427 src/Tools/jEdit/src/jedit_lib.scala --- a/src/Tools/jEdit/src/jedit_lib.scala Sat Dec 15 21:34:32 2012 +0100 +++ b/src/Tools/jEdit/src/jedit_lib.scala Sat Dec 15 22:19:14 2012 +0100 @@ -9,7 +9,7 @@ import isabelle._ -import java.awt.{Component, Container, Frame, Window} +import java.awt.{Component, Container, Frame, Window, GraphicsEnvironment, Point, Rectangle} import scala.annotation.tailrec @@ -20,6 +20,19 @@ object JEdit_Lib { + /* bounds within multi-screen environment */ + + def screen_bounds(screen_point: Point): Rectangle = + { + val ge = GraphicsEnvironment.getLocalGraphicsEnvironment + (for { + device <- ge.getScreenDevices.iterator + config <- device.getConfigurations.iterator + bounds = config.getBounds + } yield bounds).find(_.contains(screen_point)) getOrElse ge.getMaximumWindowBounds + } + + /* GUI components */ def get_parent(component: Component): Option[Container] = diff -r 89c0d2f13cca -r e4dc37ec1427 src/Tools/jEdit/src/pretty_tooltip.scala --- a/src/Tools/jEdit/src/pretty_tooltip.scala Sat Dec 15 21:34:32 2012 +0100 +++ b/src/Tools/jEdit/src/pretty_tooltip.scala Sat Dec 15 22:19:14 2012 +0100 @@ -9,7 +9,7 @@ import isabelle._ -import java.awt.{Toolkit, Color, Point, BorderLayout, Window, Dimension} +import java.awt.{Color, Point, BorderLayout, Window, Dimension} import java.awt.event.{ActionListener, ActionEvent, KeyEvent, WindowEvent, WindowAdapter} import javax.swing.{SwingUtilities, JWindow, JPanel, JComponent, KeyStroke} import javax.swing.border.LineBorder @@ -107,27 +107,27 @@ /* window geometry */ + val screen_point = new Point(mouse_x, mouse_y) + SwingUtilities.convertPointToScreen(screen_point, parent) + val screen_bounds = JEdit_Lib.screen_bounds(screen_point) + { val font_metrics = pretty_text_area.getPainter.getFontMetrics - val margin = PIDE.options.int("jedit_tooltip_margin") // FIXME via rendering?! + val margin = rendering.tooltip_margin val lines = XML.traverse_text(Pretty.formatted(body, margin, Pretty.font_metric(font_metrics)))(0)( (n: Int, s: String) => n + s.iterator.filter(_ == '\n').length) - val screen = Toolkit.getDefaultToolkit.getScreenSize - val w = (font_metrics.charWidth(Pretty.spc) * (margin + 2)) min (screen.width / 2) - val h = (font_metrics.getHeight * (lines + 2)) min (screen.height / 2) + val bounds = rendering.tooltip_bounds + val w = + (font_metrics.charWidth(Pretty.spc) * (margin + 2)) min (screen_bounds.width * bounds).toInt + val h = + (font_metrics.getHeight * (lines + 2)) min (screen_bounds.height * bounds).toInt pretty_text_area.setPreferredSize(new Dimension(w, h)) window.pack - } - { - val point = new Point(mouse_x, mouse_y) - SwingUtilities.convertPointToScreen(point, parent) - - val screen = Toolkit.getDefaultToolkit.getScreenSize - val x = point.x min (screen.width - window.getWidth) - val y = point.y min (screen.height - window.getHeight) + val x = screen_point.x min (screen_bounds.x + screen_bounds.width - window.getWidth) + val y = screen_point.y min (screen_bounds.y + screen_bounds.height - window.getHeight) window.setLocation(x, y) } diff -r 89c0d2f13cca -r e4dc37ec1427 src/Tools/jEdit/src/rendering.scala --- a/src/Tools/jEdit/src/rendering.scala Sat Dec 15 21:34:32 2012 +0100 +++ b/src/Tools/jEdit/src/rendering.scala Sat Dec 15 22:19:14 2012 +0100 @@ -346,6 +346,9 @@ Library.separate(Pretty.FBreak, all_tips.toList) } + val tooltip_margin: Int = options.int("jedit_tooltip_margin") + val tooltip_bounds: Double = (options.real("jedit_tooltip_bounds") max 0.2) min 0.8 + def gutter_message(range: Text.Range): Option[Icon] = { diff -r 89c0d2f13cca -r e4dc37ec1427 src/Tools/jEdit/src/rich_text_area.scala --- a/src/Tools/jEdit/src/rich_text_area.scala Sat Dec 15 21:34:32 2012 +0100 +++ b/src/Tools/jEdit/src/rich_text_area.scala Sat Dec 15 22:19:14 2012 +0100 @@ -10,7 +10,7 @@ import isabelle._ -import java.awt.{Graphics2D, Shape, Window, Color, Point} +import java.awt.{Graphics2D, Shape, Window, Color, Point, Toolkit} import java.awt.event.{MouseMotionAdapter, MouseAdapter, MouseEvent, FocusAdapter, FocusEvent, WindowEvent, WindowAdapter} import java.awt.font.TextAttribute @@ -168,7 +168,7 @@ private val mouse_motion_listener = new MouseMotionAdapter { override def mouseMoved(e: MouseEvent) { robust_body(()) { - control = if (OperatingSystem.isMacOS()) e.isMetaDown else e.isControlDown + control = (e.getModifiers & Toolkit.getDefaultToolkit.getMenuShortcutKeyMask) != 0 if ((control || hovering) && !buffer.isLoading) { JEdit_Lib.buffer_lock(buffer) {