--- a/src/HOL/Tools/Nitpick/nitpick.ML Fri Sep 14 12:09:27 2012 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick.ML Fri Sep 14 21:23:06 2012 +0200
@@ -245,7 +245,7 @@
if mode = Auto_Try then
Unsynchronized.change state_ref o Proof.goal_message o K
o Pretty.chunks o cons (Pretty.str "") o single
- o Pretty.mark Isabelle_Markup.hilite
+ o Pretty.mark Isabelle_Markup.intensify
else
print o Pretty.string_of
val pprint_a = pprint Output.urgent_message
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_run.ML Fri Sep 14 12:09:27 2012 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_run.ML Fri Sep 14 21:23:06 2012 +0200
@@ -137,7 +137,7 @@
|> outcome_code = someN
? Proof.goal_message (fn () =>
[Pretty.str "",
- Pretty.mark Isabelle_Markup.hilite (Pretty.str (message ()))]
+ Pretty.mark Isabelle_Markup.intensify (Pretty.str (message ()))]
|> Pretty.chunks))
end
else if blocking then
--- a/src/HOL/Tools/try0.ML Fri Sep 14 12:09:27 2012 +0200
+++ b/src/HOL/Tools/try0.ML Fri Sep 14 21:23:06 2012 +0200
@@ -146,7 +146,7 @@
(true, (s, st |> (if mode = Auto_Try then
Proof.goal_message
(fn () => Pretty.chunks [Pretty.str "",
- Pretty.markup Isabelle_Markup.hilite
+ Pretty.markup Isabelle_Markup.intensify
[Pretty.str message]])
else
tap (fn _ => Output.urgent_message message))))
--- a/src/Pure/General/name_space.ML Fri Sep 14 12:09:27 2012 +0200
+++ b/src/Pure/General/name_space.ML Fri Sep 14 21:23:06 2012 +0200
@@ -133,7 +133,7 @@
fun markup (Name_Space {kind, entries, ...}) name =
(case Symtab.lookup entries name of
- NONE => Isabelle_Markup.hilite
+ NONE => Isabelle_Markup.intensify
| SOME (_, entry) => entry_markup false kind (name, entry));
fun is_concealed space name = #concealed (the_entry space name);
--- a/src/Pure/PIDE/command.scala Fri Sep 14 12:09:27 2012 +0200
+++ b/src/Pure/PIDE/command.scala Fri Sep 14 21:23:06 2012 +0200
@@ -102,6 +102,8 @@
def unparsed(source: String): Command =
Command(Document.no_id, Document.Node.Name.empty, List(Token(Token.Kind.UNPARSED, source)))
+ val empty = Command(Document.no_id, Document.Node.Name.empty, Nil)
+
/* perspective */
--- a/src/Pure/PIDE/isabelle_markup.ML Fri Sep 14 12:09:27 2012 +0200
+++ b/src/Pure/PIDE/isabelle_markup.ML Fri Sep 14 21:23:06 2012 +0200
@@ -84,7 +84,7 @@
val stateN: string val state: Markup.T
val subgoalN: string val subgoal: Markup.T
val sendbackN: string val sendback: Markup.T
- val hiliteN: string val hilite: Markup.T
+ val intensifyN: string val intensify: Markup.T
val taskN: string
val acceptedN: string val accepted: Markup.T
val forkedN: string val forked: Markup.T
@@ -264,7 +264,7 @@
val (stateN, state) = markup_elem "state";
val (subgoalN, subgoal) = markup_elem "subgoal";
val (sendbackN, sendback) = markup_elem "sendback";
-val (hiliteN, hilite) = markup_elem "hilite";
+val (intensifyN, intensify) = markup_elem "intensify";
(* command status *)
--- a/src/Pure/PIDE/isabelle_markup.scala Fri Sep 14 12:09:27 2012 +0200
+++ b/src/Pure/PIDE/isabelle_markup.scala Fri Sep 14 21:23:06 2012 +0200
@@ -185,7 +185,7 @@
val STATE = "state"
val SUBGOAL = "subgoal"
val SENDBACK = "sendback"
- val HILITE = "hilite"
+ val INTENSIFY = "intensify"
/* command status */
--- a/src/Pure/ProofGeneral/proof_general_emacs.ML Fri Sep 14 12:09:27 2012 +0200
+++ b/src/Pure/ProofGeneral/proof_general_emacs.ML Fri Sep 14 21:23:06 2012 +0200
@@ -41,7 +41,7 @@
if null ts then Markup.no_output
else if name = Isabelle_Markup.stateN then (special "O" ^ "\n", "\n" ^ special "P")
else if name = Isabelle_Markup.sendbackN then (special "W", special "X")
- else if name = Isabelle_Markup.hiliteN then (special "0", special "1")
+ else if name = Isabelle_Markup.intensifyN then (special "0", special "1")
else if name = Isabelle_Markup.tfreeN then (special "C", special "A")
else if name = Isabelle_Markup.tvarN then (special "D", special "A")
else if name = Isabelle_Markup.freeN then (special "E", special "A")
--- a/src/Pure/Syntax/syntax_phases.ML Fri Sep 14 12:09:27 2012 +0200
+++ b/src/Pure/Syntax/syntax_phases.ML Fri Sep 14 21:23:06 2012 +0200
@@ -586,7 +586,7 @@
val m =
if Variable.is_fixed ctxt x orelse Syntax.is_pretty_global ctxt
then Isabelle_Markup.fixed x
- else Isabelle_Markup.hilite;
+ else Isabelle_Markup.intensify;
in
if can Name.dest_skolem x
then ([m, Isabelle_Markup.skolem], Variable.revert_fixed ctxt x)
--- a/src/Tools/jEdit/etc/isabelle-jedit.css Fri Sep 14 12:09:27 2012 +0200
+++ b/src/Tools/jEdit/etc/isabelle-jedit.css Fri Sep 14 21:23:06 2012 +0200
@@ -9,7 +9,7 @@
.report { display: none; }
-.hilite { background-color: #FFCC66; }
+.intensify { background-color: #FFCC66; }
.keyword { font-weight: bold; color: #009966; }
.operator { font-weight: bold; }
--- a/src/Tools/jEdit/etc/options Fri Sep 14 12:09:27 2012 +0200
+++ b/src/Tools/jEdit/etc/options Fri Sep 14 21:23:06 2012 +0200
@@ -21,32 +21,33 @@
section "Rendering of Document Content"
-option color_outdated : string = "EEE3E3FF"
-option color_unprocessed : string = "FFA0A0FF"
-option color_unprocessed1 : string = "FFA0A032"
-option color_running : string = "610061FF"
-option color_running1 : string = "61006164"
-option color_light : string = "F0F0F0FF"
-option color_writeln : string = "C0C0C0FF"
-option color_warning : string = "FF8C00FF"
-option color_error : string = "B22222FF"
-option color_error1 : string = "B2222232"
-option color_bad : string = "FF6A6A64"
-option color_hilite : string = "FFCC6664"
-option color_quoted : string = "8B8B8B19"
-option color_subexp : string = "50505032"
-option color_hyperlink : string = "000000FF"
-option color_keyword1 : string = "006699FF"
-option color_keyword2 : string = "009966FF"
+option outdated_color : string = "EEE3E3FF"
+option unprocessed_color : string = "FFA0A0FF"
+option unprocessed1_color : string = "FFA0A032"
+option running_color : string = "610061FF"
+option running1_color : string = "61006164"
+option light_color : string = "F0F0F0FF"
+option writeln_color : string = "C0C0C0FF"
+option warning_color : string = "FF8C00FF"
+option error_color : string = "B22222FF"
+option error1_color : string = "B2222232"
+option bad_color : string = "FF6A6A64"
+option intensify_color : string = "FFCC6664"
+option quoted_color : string = "8B8B8B19"
+option highlight_color : string = "50505032"
+option hyperlink_color : string = "000000FF"
+option keyword1_color : string = "006699FF"
+option keyword2_color : string = "009966FF"
-option color_variable_free : string = "0000FFFF"
-option color_variable_type : string = "A020F0FF"
-option color_variable_skolem : string = "D2691EFF"
-option color_variable_bound : string = "008000FF"
-option color_variable_schematic : string = "00009BFF"
-option color_inner_quoted : string = "D2691EFF"
-option color_inner_comment : string = "8B0000FF"
-option color_inner_numeral : string = "FF0000FF"
-option color_antiquotation : string = "0000FFFF"
-option color_dynamic : string = "7BA428FF"
+option tfree_color : string = "A020F0FF"
+option tvar_color : string = "A020F0FF"
+option free_color : string = "0000FFFF"
+option skolem_color : string = "D2691EFF"
+option bound_color : string = "008000FF"
+option var_color : string = "00009BFF"
+option inner_numeral_color : string = "FF0000FF"
+option inner_quoted_color : string = "D2691EFF"
+option inner_comment_color : string = "8B0000FF"
+option antiquotation_color : string = "0000FFFF"
+option dynamic_color : string = "7BA428FF"
--- a/src/Tools/jEdit/src/Isabelle.props Fri Sep 14 12:09:27 2012 +0200
+++ b/src/Tools/jEdit/src/Isabelle.props Fri Sep 14 21:23:06 2012 +0200
@@ -20,9 +20,11 @@
plugin.isabelle.jedit.Plugin.depend.4=plugin sidekick.SideKickPlugin 0.8
#options
-plugin.isabelle.jedit.Plugin.option-pane=isabelle
-options.isabelle.label=Isabelle
-options.isabelle.code=new isabelle.jedit.Isabelle_Options();
+plugin.isabelle.jedit.Plugin.option-group=isabelle-general isabelle-rendering
+options.isabelle-general.label=General
+options.isabelle-general.code=new isabelle.jedit.Isabelle_Options1();
+options.isabelle-rendering.label=Rendering
+options.isabelle-rendering.code=new isabelle.jedit.Isabelle_Options2();
#actions
isabelle.check-buffer.label=Commence full proof checking of current buffer
--- a/src/Tools/jEdit/src/document_model.scala Fri Sep 14 12:09:27 2012 +0200
+++ b/src/Tools/jEdit/src/document_model.scala Fri Sep 14 21:23:06 2012 +0200
@@ -92,6 +92,24 @@
}
+
+ /* point range */
+
+ def point_range(offset: Text.Offset): Text.Range =
+ Isabelle.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) }
+ }
+
+
/* pending text edits */
private object pending_edits // owned by Swing thread
--- a/src/Tools/jEdit/src/document_view.scala Fri Sep 14 12:09:27 2012 +0200
+++ b/src/Tools/jEdit/src/document_view.scala Fri Sep 14 21:23:06 2012 +0200
@@ -19,7 +19,6 @@
import java.awt.{Color, Graphics2D, Point}
import java.awt.event.{MouseMotionAdapter, MouseAdapter, MouseEvent,
FocusAdapter, FocusEvent, WindowEvent, WindowAdapter}
-import javax.swing.{Popup, PopupFactory, SwingUtilities, BorderFactory}
import javax.swing.event.{CaretListener, CaretEvent}
import org.gjt.sp.util.Log
@@ -143,77 +142,56 @@
}
- /* HTML popups */
+ /* active areas within the text */
- private var html_popup: Option[Popup] = None
+ private class Active_Area[A](
+ rendering: Isabelle_Rendering => Text.Range => Option[Text.Info[A]])
+ {
+ private var the_info: Option[Text.Info[A]] = None
- private def exit_popup() { html_popup.map(_.hide) }
+ def info: Option[Text.Info[A]] = the_info
- private val html_panel =
- new HTML_Panel(Isabelle.font_family(), scala.math.round(Isabelle.font_size()))
- html_panel.setBorder(BorderFactory.createLineBorder(Color.black))
+ def update(new_info: Option[Text.Info[A]])
+ {
+ val old_info = the_info
+ if (new_info != old_info) {
+ for { opt <- List(old_info, new_info); Text.Info(range, _) <- opt }
+ invalidate_range(range)
+ the_info = new_info
+ }
+ }
- private def html_panel_resize()
- {
- Swing_Thread.now {
- html_panel.resize(Isabelle.font_family(), scala.math.round(Isabelle.font_size()))
- }
+ def update_rendering(r: Isabelle_Rendering, range: Text.Range)
+ { update(rendering(r)(range)) }
+
+ def reset { update(None) }
}
- private def init_popup(snapshot: Document.Snapshot, x: Int, y: Int)
- {
- exit_popup()
-/* FIXME broken
- val offset = text_area.xyToOffset(x, y)
- val p = new Point(x, y); SwingUtilities.convertPointToScreen(p, text_area.getPainter)
-
- // FIXME snapshot.cumulate
- snapshot.select_markup(Text.Range(offset, offset + 1))(Isabelle_Rendering.popup) match {
- case Text.Info(_, Some(msg)) #:: _ =>
- val popup = PopupFactory.getSharedInstance().getPopup(text_area, html_panel, p.x, p.y + 60)
- html_panel.render_sync(List(msg))
- Thread.sleep(10) // FIXME !?
- popup.show
- html_popup = Some(popup)
- case _ =>
- }
-*/
- }
-
-
- /* subexpression highlighting and hyperlinks */
-
- @volatile private var _highlight_range: Option[Text.Info[Color]] = None
- def highlight_range(): Option[Text.Info[Color]] = _highlight_range
-
- @volatile private var _hyperlink_range: Option[Text.Info[Hyperlink]] = None
- def hyperlink_range(): Option[Text.Info[Hyperlink]] = _hyperlink_range
+ // owned by Swing thread
private var control: Boolean = false
- private def exit_control()
- {
- exit_popup()
- _highlight_range = None
- _hyperlink_range = None
- }
+ private val highlight_area = new Active_Area[Color]((r: Isabelle_Rendering) => r.highlight _)
+ def highlight_info(): Option[Text.Info[Color]] = highlight_area.info
+
+ private val hyperlink_area = new Active_Area[Hyperlink]((r: Isabelle_Rendering) => r.hyperlink _)
+ def hyperlink_info(): Option[Text.Info[Hyperlink]] = hyperlink_area.info
+
+ private val active_areas = List(highlight_area, hyperlink_area)
+ private def active_reset(): Unit = active_areas.foreach(_.reset)
private val focus_listener = new FocusAdapter {
- override def focusLost(e: FocusEvent) {
- // FIXME exit_control !?
- _highlight_range = None
- _hyperlink_range = None
- }
+ override def focusLost(e: FocusEvent) { active_reset() }
}
private val window_listener = new WindowAdapter {
- override def windowIconified(e: WindowEvent) { exit_control() }
- override def windowDeactivated(e: WindowEvent) { exit_control() }
+ override def windowIconified(e: WindowEvent) { active_reset() }
+ override def windowDeactivated(e: WindowEvent) { active_reset() }
}
private val mouse_listener = new MouseAdapter {
override def mouseClicked(e: MouseEvent) {
- hyperlink_range match {
+ hyperlink_area.info match {
case Some(Text.Info(range, link)) => link.follow(text_area.getView)
case None =>
}
@@ -222,36 +200,15 @@
private val mouse_motion_listener = new MouseMotionAdapter {
override def mouseMoved(e: MouseEvent) {
- Swing_Thread.assert()
-
control = if (OperatingSystem.isMacOS()) e.isMetaDown else e.isControlDown
- val x = e.getX()
- val y = e.getY()
-
- if (!model.buffer.isLoaded) exit_control()
- else
+ if (control && model.buffer.isLoaded) {
Isabelle.buffer_lock(model.buffer) {
- val snapshot = model.snapshot()
-
- if (control) init_popup(snapshot, x, y)
-
- def update_range[A](
- rendering: (Document.Snapshot, Text.Range) => Option[Text.Info[A]],
- info: Option[Text.Info[A]]): Option[Text.Info[A]] =
- {
- for (Text.Info(range, _) <- info) invalidate_range(range)
- val new_info =
- if (control) {
- val offset = text_area.xyToOffset(x, y)
- rendering(snapshot, Text.Range(offset, offset + 1))
- }
- else None
- for (Text.Info(range, _) <- info) invalidate_range(range)
- new_info
- }
- _highlight_range = update_range(Isabelle_Rendering.subexp, _highlight_range)
- _hyperlink_range = update_range(Isabelle_Rendering.hyperlink, _hyperlink_range)
+ val rendering = Isabelle_Rendering(model.snapshot(), Isabelle.options.value)
+ val mouse_range = model.point_range(text_area.xyToOffset(e.getX(), e.getY()))
+ active_areas.foreach(_.update_rendering(rendering, mouse_range))
}
+ }
+ else active_reset()
}
}
@@ -265,12 +222,12 @@
override def getToolTipText(x: Int, y: Int): String =
{
robust_body(null: String) {
- val snapshot = model.snapshot()
+ val rendering = Isabelle_Rendering(model.snapshot(), Isabelle.options.value)
val offset = text_area.xyToOffset(x, y)
val range = Text.Range(offset, offset + 1)
val tip =
- if (control) Isabelle_Rendering.tooltip(snapshot, range)
- else Isabelle_Rendering.tooltip_message(snapshot, range)
+ if (control) rendering.tooltip(range)
+ else rendering.tooltip_message(range)
tip.map(Isabelle.tooltip(_)) getOrElse null
}
}
@@ -292,13 +249,14 @@
if (gutter.isSelectionAreaEnabled && !gutter.isExpanded && width >= 12 && line_height >= 12) {
Isabelle.buffer_lock(model.buffer) {
- val snapshot = model.snapshot()
+ val rendering = Isabelle_Rendering(model.snapshot(), Isabelle.options.value)
+
for (i <- 0 until physical_lines.length) {
if (physical_lines(i) != -1) {
val line_range = proper_line_range(start(i), end(i))
// gutter icons
- Isabelle_Rendering.gutter_message(snapshot, line_range) match {
+ rendering.gutter_message(line_range) match {
case Some(icon) =>
val x0 = (FOLD_MARKER_SIZE + width - border_width - icon.getIconWidth) max 10
val y0 = y + i * line_height + (((line_height - icon.getIconHeight) / 2) max 0)
@@ -314,32 +272,8 @@
}
- /* caret range */
-
- def caret_range(): Text.Range =
- Isabelle.buffer_lock(model.buffer) {
- def text(i: Text.Offset): Char = model.buffer.getText(i, 1).charAt(0)
- val caret = text_area.getCaretPosition
- try {
- val c = text(caret)
- if (Character.isHighSurrogate(c) && Character.isLowSurrogate(text(caret + 1)))
- Text.Range(caret, caret + 2)
- else if (Character.isLowSurrogate(c) && Character.isHighSurrogate(text(caret - 1)))
- Text.Range(caret - 1, caret + 1)
- else Text.Range(caret, caret + 1)
- }
- catch { case _: ArrayIndexOutOfBoundsException => Text.Range(caret, caret + 1) }
- }
-
-
/* caret handling */
- def selected_command(): Option[Command] =
- {
- Swing_Thread.require()
- model.snapshot().node.command_at(text_area.getCaretPosition).map(_._1)
- }
-
private val delay_caret_update =
Swing_Thread.delay_last(Time.seconds(Isabelle.options.real("editor_input_delay"))) {
session.caret_focus.event(Session.Caret_Focus)
@@ -406,8 +340,6 @@
}
}
- case Session.Global_Settings => html_panel_resize()
-
case bad => System.err.println("command_change_actor: ignoring bad message " + bad)
}
}
@@ -431,7 +363,6 @@
text_area.addLeftOfScrollBar(overview)
session.raw_edits += main_actor
session.commands_changed += main_actor
- session.global_settings += main_actor
}
private def deactivate()
@@ -439,7 +370,6 @@
val painter = text_area.getPainter
session.raw_edits -= main_actor
session.commands_changed -= main_actor
- session.global_settings -= main_actor
text_area.removeFocusListener(focus_listener)
text_area.getView.removeWindowListener(window_listener)
painter.removeMouseMotionListener(mouse_motion_listener)
@@ -450,6 +380,5 @@
text_area_painter.deactivate()
painter.removeExtension(tooltip_painter)
painter.removeExtension(update_perspective)
- exit_popup()
}
}
--- a/src/Tools/jEdit/src/isabelle_options.scala Fri Sep 14 12:09:27 2012 +0200
+++ b/src/Tools/jEdit/src/isabelle_options.scala Fri Sep 14 21:23:06 2012 +0200
@@ -12,25 +12,9 @@
import org.gjt.sp.jedit.{jEdit, AbstractOptionPane}
-class Isabelle_Options extends AbstractOptionPane("isabelle")
+abstract class Isabelle_Options(name: String) extends AbstractOptionPane(name)
{
- // FIXME avoid hard-wired stuff
- private val relevant_options =
- Set("jedit_logic", "jedit_auto_start", "jedit_font_scale", "jedit_tooltip_font_size",
- "jedit_tooltip_margin", "jedit_tooltip_dismiss_delay", "editor_load_delay",
- "editor_input_delay", "editor_output_delay", "editor_update_delay")
-
- relevant_options.foreach(Isabelle.options.value.check_name _)
-
- private val predefined =
- Isabelle_Logic.logic_selector(false) ::
- (for {
- (name, opt) <- Isabelle.options.value.options.toList
- // FIXME avoid hard-wired stuff
- if (name.startsWith("color_") && opt.section == "Rendering of Document Content")
- } yield Isabelle.options.make_color_component(opt))
-
- private val components = Isabelle.options.make_components(predefined, relevant_options)
+ protected val components: List[(String, List[Option_Component])]
override def _init()
{
@@ -51,3 +35,34 @@
for ((_, cs) <- components) cs.foreach(_.save())
}
}
+
+
+class Isabelle_Options1 extends Isabelle_Options("isabelle-general")
+{
+ // FIXME avoid hard-wired stuff
+ private val relevant_options =
+ Set("jedit_logic", "jedit_auto_start", "jedit_font_scale", "jedit_tooltip_font_size",
+ "jedit_tooltip_margin", "jedit_tooltip_dismiss_delay", "editor_load_delay",
+ "editor_input_delay", "editor_output_delay", "editor_update_delay")
+
+ relevant_options.foreach(Isabelle.options.value.check_name _)
+
+ protected val components =
+ Isabelle.options.make_components(List(Isabelle_Logic.logic_selector(false)), relevant_options)
+}
+
+
+class Isabelle_Options2 extends Isabelle_Options("isabelle-rendering")
+{
+ // FIXME avoid hard-wired stuff
+ private val predefined =
+ (for {
+ (name, opt) <- Isabelle.options.value.options.toList
+ if (name.endsWith("_color") && opt.section == "Rendering of Document Content")
+ } yield Isabelle.options.make_color_component(opt))
+
+ assert(!predefined.isEmpty)
+
+ protected val components = Isabelle.options.make_components(predefined, _ => false)
+}
+
--- a/src/Tools/jEdit/src/isabelle_rendering.scala Fri Sep 14 12:09:27 2012 +0200
+++ b/src/Tools/jEdit/src/isabelle_rendering.scala Fri Sep 14 21:23:06 2012 +0200
@@ -20,326 +20,22 @@
object Isabelle_Rendering
{
- /* physical rendering */
+ def apply(snapshot: Document.Snapshot, options: Options): Isabelle_Rendering =
+ new Isabelle_Rendering(snapshot, options)
- def color_value(s: String): Color = Color_Value(Isabelle.options.value.string(s))
+
+ /* physical rendering */
private val writeln_pri = 1
private val warning_pri = 2
private val legacy_pri = 3
private val error_pri = 4
-
- /* command overview */
-
- def overview_color(snapshot: Document.Snapshot, range: Text.Range): Option[Color] =
- {
- if (snapshot.is_outdated) None
- else {
- val results =
- snapshot.cumulate_markup[(Protocol.Status, Int)](
- range, (Protocol.Status.init, 0),
- Some(Protocol.command_status_markup + Isabelle_Markup.WARNING + Isabelle_Markup.ERROR),
- {
- case ((status, pri), Text.Info(_, XML.Elem(markup, _))) =>
- if (markup.name == Isabelle_Markup.WARNING) (status, pri max warning_pri)
- else if (markup.name == Isabelle_Markup.ERROR) (status, pri max error_pri)
- else (Protocol.command_status(status, markup), pri)
- })
- if (results.isEmpty) None
- else {
- val (status, pri) =
- ((Protocol.Status.init, 0) /: results) {
- case ((s1, p1), Text.Info(_, (s2, p2))) => (s1 + s2, p1 max p2) }
-
- if (pri == warning_pri) Some(color_value("color_warning"))
- else if (pri == error_pri) Some(color_value("color_error"))
- else if (status.is_unprocessed) Some(color_value("color_unprocessed"))
- else if (status.is_running) Some(color_value("color_running"))
- else if (status.is_failed) Some(color_value("color_error"))
- else None
- }
- }
- }
-
-
- /* markup selectors */
-
- private val subexp_include =
- Set(Isabelle_Markup.SORT, Isabelle_Markup.TYP, Isabelle_Markup.TERM, Isabelle_Markup.PROP,
- Isabelle_Markup.ML_TYPING, Isabelle_Markup.TOKEN_RANGE, Isabelle_Markup.ENTITY,
- Isabelle_Markup.PATH, Isabelle_Markup.TYPING, Isabelle_Markup.FREE, Isabelle_Markup.SKOLEM,
- Isabelle_Markup.BOUND, Isabelle_Markup.VAR, Isabelle_Markup.TFREE, Isabelle_Markup.TVAR,
- Isabelle_Markup.ML_SOURCE, Isabelle_Markup.DOCUMENT_SOURCE)
-
- def subexp(snapshot: Document.Snapshot, range: Text.Range): Option[Text.Info[Color]] =
- {
- snapshot.select_markup(range, Some(subexp_include),
- {
- case Text.Info(info_range, XML.Elem(Markup(name, _), _)) if subexp_include(name) =>
- Text.Info(snapshot.convert(info_range), color_value("color_subexp"))
- }) match { case Text.Info(_, info) #:: _ => Some(info) case _ => None }
- }
-
-
- private val hyperlink_include = Set(Isabelle_Markup.ENTITY, Isabelle_Markup.PATH)
-
- def hyperlink(snapshot: Document.Snapshot, range: Text.Range): Option[Text.Info[Hyperlink]] =
- {
- snapshot.cumulate_markup[List[Text.Info[Hyperlink]]](range, Nil, Some(hyperlink_include),
- {
- case (links, Text.Info(info_range, XML.Elem(Isabelle_Markup.Path(name), _)))
- if Path.is_ok(name) =>
- val jedit_file = Isabelle.thy_load.append(snapshot.node_name.dir, Path.explode(name))
- Text.Info(snapshot.convert(info_range), Hyperlink(jedit_file, 0, 0)) :: links
-
- case (links, Text.Info(info_range, XML.Elem(Markup(Isabelle_Markup.ENTITY, props), _)))
- if (props.find(
- { case (Markup.KIND, Isabelle_Markup.ML_OPEN) => true
- case (Markup.KIND, Isabelle_Markup.ML_STRUCT) => true
- case _ => false }).isEmpty) =>
-
- props match {
- case Position.Line_File(line, name) if Path.is_ok(name) =>
- Isabelle_System.source_file(Path.explode(name)) match {
- case Some(path) =>
- val jedit_file = Isabelle_System.platform_path(path)
- Text.Info(snapshot.convert(info_range), Hyperlink(jedit_file, line, 0)) :: links
- case None => links
- }
-
- case Position.Id_Offset(id, offset) if !snapshot.is_outdated =>
- snapshot.state.find_command(snapshot.version, id) match {
- case Some((node, command)) =>
- val sources =
- node.commands.iterator.takeWhile(_ != command).map(_.source) ++
- Iterator.single(command.source(Text.Range(0, command.decode(offset))))
- val (line, column) = ((1, 1) /: sources)(Symbol.advance_line_column)
- Text.Info(snapshot.convert(info_range),
- Hyperlink(command.node_name.node, line, column)) :: links
- case None => links
- }
-
- case _ => links
- }
- }) match { case Text.Info(_, info :: _) #:: _ => Some(info) case _ => None }
- }
-
-
- private def tooltip_text(msg: XML.Tree): String =
- Pretty.string_of(List(msg), margin = Isabelle.options.int("jedit_tooltip_margin"))
-
- def tooltip_message(snapshot: Document.Snapshot, range: Text.Range): Option[String] =
- {
- val msgs =
- snapshot.cumulate_markup[SortedMap[Long, String]](range, SortedMap.empty,
- Some(Set(Isabelle_Markup.WRITELN, Isabelle_Markup.WARNING, Isabelle_Markup.ERROR,
- Isabelle_Markup.BAD)),
- {
- case (msgs, Text.Info(_, msg @
- XML.Elem(Markup(markup, Isabelle_Markup.Serial(serial)), _)))
- if markup == Isabelle_Markup.WRITELN ||
- markup == Isabelle_Markup.WARNING ||
- markup == Isabelle_Markup.ERROR =>
- msgs + (serial -> tooltip_text(msg))
- case (msgs, Text.Info(_, msg @ XML.Elem(Markup(Isabelle_Markup.BAD, _), body)))
- if !body.isEmpty => msgs + (Document.new_id() -> tooltip_text(msg))
- }).toList.flatMap(_.info)
- if (msgs.isEmpty) None else Some(cat_lines(msgs.iterator.map(_._2)))
- }
-
-
- private val tooltips: Map[String, String] =
- Map(
- Isabelle_Markup.SORT -> "sort",
- Isabelle_Markup.TYP -> "type",
- Isabelle_Markup.TERM -> "term",
- Isabelle_Markup.PROP -> "proposition",
- Isabelle_Markup.TOKEN_RANGE -> "inner syntax token",
- Isabelle_Markup.FREE -> "free variable",
- Isabelle_Markup.SKOLEM -> "skolem variable",
- Isabelle_Markup.BOUND -> "bound variable",
- Isabelle_Markup.VAR -> "schematic variable",
- Isabelle_Markup.TFREE -> "free type variable",
- Isabelle_Markup.TVAR -> "schematic type variable",
- Isabelle_Markup.ML_SOURCE -> "ML source",
- Isabelle_Markup.DOCUMENT_SOURCE -> "document source")
-
- private val tooltip_elements =
- Set(Isabelle_Markup.ENTITY, Isabelle_Markup.TYPING, Isabelle_Markup.ML_TYPING,
- Isabelle_Markup.PATH) ++ tooltips.keys
-
- private def string_of_typing(kind: String, body: XML.Body): String =
- Pretty.string_of(List(Pretty.block(XML.Text(kind) :: Pretty.Break(1) :: body)),
- margin = Isabelle.options.int("jedit_tooltip_margin"))
-
- def tooltip(snapshot: Document.Snapshot, range: Text.Range): Option[String] =
- {
- def add(prev: Text.Info[List[(Boolean, String)]], r: Text.Range, p: (Boolean, String)) =
- if (prev.range == r) Text.Info(r, p :: prev.info) else Text.Info(r, List(p))
-
- val tips =
- snapshot.cumulate_markup[Text.Info[(List[(Boolean, String)])]](
- range, Text.Info(range, Nil), Some(tooltip_elements),
- {
- case (prev, Text.Info(r, XML.Elem(Isabelle_Markup.Entity(kind, name), _))) =>
- val kind1 = space_explode('_', kind).mkString(" ")
- add(prev, r, (true, kind1 + " " + quote(name)))
- case (prev, Text.Info(r, XML.Elem(Isabelle_Markup.Path(name), _)))
- if Path.is_ok(name) =>
- val jedit_file = Isabelle.thy_load.append(snapshot.node_name.dir, Path.explode(name))
- add(prev, r, (true, "file " + quote(jedit_file)))
- case (prev, Text.Info(r, XML.Elem(Markup(Isabelle_Markup.TYPING, _), body))) =>
- add(prev, r, (true, string_of_typing("::", body)))
- case (prev, Text.Info(r, XML.Elem(Markup(Isabelle_Markup.ML_TYPING, _), body))) =>
- add(prev, r, (false, string_of_typing("ML:", body)))
- case (prev, Text.Info(r, XML.Elem(Markup(name, _), _)))
- if tooltips.isDefinedAt(name) => add(prev, r, (true, tooltips(name)))
- }).toList.flatMap(_.info.info)
-
- val all_tips =
- (tips.filter(_._1) ++ tips.filter(!_._1).lastOption.toList).map(_._2)
- if (all_tips.isEmpty) None else Some(cat_lines(all_tips))
- }
-
-
private val gutter_icons = Map(
warning_pri -> Isabelle.load_icon("16x16/status/dialog-information.png"),
legacy_pri -> Isabelle.load_icon("16x16/status/dialog-warning.png"),
error_pri -> Isabelle.load_icon("16x16/status/dialog-error.png"))
- def gutter_message(snapshot: Document.Snapshot, range: Text.Range): Option[Icon] =
- {
- val results =
- snapshot.cumulate_markup[Int](range, 0,
- Some(Set(Isabelle_Markup.WARNING, Isabelle_Markup.ERROR)),
- {
- case (pri, Text.Info(_, XML.Elem(Markup(Isabelle_Markup.WARNING, _), body))) =>
- body match {
- case List(XML.Elem(Markup(Isabelle_Markup.LEGACY, _), _)) => pri max legacy_pri
- case _ => pri max warning_pri
- }
- case (pri, Text.Info(_, XML.Elem(Markup(Isabelle_Markup.ERROR, _), _))) =>
- pri max error_pri
- })
- val pri = (0 /: results) { case (p1, Text.Info(_, p2)) => p1 max p2 }
- gutter_icons.get(pri)
- }
-
-
- def squiggly_underline(snapshot: Document.Snapshot, range: Text.Range): Stream[Text.Info[Color]] =
- {
- val squiggly_colors = Map(
- writeln_pri -> color_value("color_writeln"),
- warning_pri -> color_value("color_warning"),
- error_pri -> color_value("color_error"))
-
- val results =
- snapshot.cumulate_markup[Int](range, 0,
- Some(Set(Isabelle_Markup.WRITELN, Isabelle_Markup.WARNING, Isabelle_Markup.ERROR)),
- {
- case (pri, Text.Info(_, XML.Elem(Markup(name, _), _))) =>
- name match {
- case Isabelle_Markup.WRITELN => pri max writeln_pri
- case Isabelle_Markup.WARNING => pri max warning_pri
- case Isabelle_Markup.ERROR => pri max error_pri
- case _ => pri
- }
- })
- for {
- Text.Info(r, pri) <- results
- color <- squiggly_colors.get(pri)
- } yield Text.Info(r, color)
- }
-
-
- def background1(snapshot: Document.Snapshot, range: Text.Range): Stream[Text.Info[Color]] =
- {
- if (snapshot.is_outdated) Stream(Text.Info(range, color_value("color_outdated")))
- else
- for {
- Text.Info(r, result) <-
- snapshot.cumulate_markup[(Option[Protocol.Status], Option[Color])](
- range, (Some(Protocol.Status.init), None),
- Some(Protocol.command_status_markup + Isabelle_Markup.BAD + Isabelle_Markup.HILITE),
- {
- case (((Some(status), color), Text.Info(_, XML.Elem(markup, _))))
- if (Protocol.command_status_markup(markup.name)) =>
- (Some(Protocol.command_status(status, markup)), color)
- case (_, Text.Info(_, XML.Elem(Markup(Isabelle_Markup.BAD, _), _))) =>
- (None, Some(color_value("color_bad")))
- case (_, Text.Info(_, XML.Elem(Markup(Isabelle_Markup.HILITE, _), _))) =>
- (None, Some(color_value("color_hilite")))
- })
- color <-
- (result match {
- case (Some(status), opt_color) =>
- if (status.is_unprocessed) Some(color_value("color_unprocessed1"))
- else if (status.is_running) Some(color_value("color_running1"))
- else opt_color
- case (_, opt_color) => opt_color
- })
- } yield Text.Info(r, color)
- }
-
-
- def background2(snapshot: Document.Snapshot, range: Text.Range): Stream[Text.Info[Color]] =
- snapshot.select_markup(range,
- Some(Set(Isabelle_Markup.TOKEN_RANGE)),
- {
- case Text.Info(_, XML.Elem(Markup(Isabelle_Markup.TOKEN_RANGE, _), _)) =>
- color_value("color_light")
- })
-
-
- def foreground(snapshot: Document.Snapshot, range: Text.Range): Stream[Text.Info[Color]] =
- snapshot.select_markup(range,
- Some(Set(Isabelle_Markup.STRING, Isabelle_Markup.ALTSTRING, Isabelle_Markup.VERBATIM)),
- {
- case Text.Info(_, XML.Elem(Markup(Isabelle_Markup.STRING, _), _)) =>
- color_value("color_quoted")
- case Text.Info(_, XML.Elem(Markup(Isabelle_Markup.ALTSTRING, _), _)) =>
- color_value("color_quoted")
- case Text.Info(_, XML.Elem(Markup(Isabelle_Markup.VERBATIM, _), _)) =>
- color_value("color_quoted")
- })
-
-
- def text_color(snapshot: Document.Snapshot, range: Text.Range, color: Color)
- : Stream[Text.Info[Color]] =
- {
- val text_colors: Map[String, Color] = Map(
- Isabelle_Markup.STRING -> Color.BLACK,
- Isabelle_Markup.ALTSTRING -> Color.BLACK,
- Isabelle_Markup.VERBATIM -> Color.BLACK,
- Isabelle_Markup.LITERAL -> color_value("color_keyword1"),
- Isabelle_Markup.DELIMITER -> Color.BLACK,
- Isabelle_Markup.TFREE -> color_value("color_variable_type"),
- Isabelle_Markup.TVAR -> color_value("color_variable_type"),
- Isabelle_Markup.FREE -> color_value("color_variable_free"),
- Isabelle_Markup.SKOLEM -> color_value("color_variable_skolem"),
- Isabelle_Markup.BOUND -> color_value("color_variable_bound"),
- Isabelle_Markup.VAR -> color_value("color_variable_schematic"),
- Isabelle_Markup.INNER_STRING -> color_value("color_inner_quoted"),
- Isabelle_Markup.INNER_COMMENT -> color_value("color_inner_comment"),
- Isabelle_Markup.DYNAMIC_FACT -> color_value("color_dynamic"),
- Isabelle_Markup.ML_KEYWORD -> color_value("color_keyword1"),
- Isabelle_Markup.ML_DELIMITER -> Color.BLACK,
- Isabelle_Markup.ML_NUMERAL -> color_value("color_inner_numeral"),
- Isabelle_Markup.ML_CHAR -> color_value("color_inner_quoted"),
- Isabelle_Markup.ML_STRING -> color_value("color_inner_quoted"),
- Isabelle_Markup.ML_COMMENT -> color_value("color_inner_comment"),
- Isabelle_Markup.ANTIQ -> color_value("color_antiquotation"))
-
- val text_color_elements = Set.empty[String] ++ text_colors.keys
-
- snapshot.cumulate_markup(range, color, Some(text_color_elements),
- {
- case (_, Text.Info(_, XML.Elem(Markup(m, _), _)))
- if text_colors.isDefinedAt(m) => text_colors(m)
- })
- }
-
/* token markup -- text styles */
@@ -382,3 +78,346 @@
else if (token.is_operator) JEditToken.OPERATOR
else token_style(token.kind)
}
+
+
+class Isabelle_Rendering private(val snapshot: Document.Snapshot, val options: Options)
+{
+ /* colors */
+
+ def color_value(s: String): Color = Color_Value(options.string(s))
+
+ val outdated_color = color_value("outdated_color")
+ val unprocessed_color = color_value("unprocessed_color")
+ val unprocessed1_color = color_value("unprocessed1_color")
+ val running_color = color_value("running_color")
+ val running1_color = color_value("running1_color")
+ val light_color = color_value("light_color")
+ val writeln_color = color_value("writeln_color")
+ val warning_color = color_value("warning_color")
+ val error_color = color_value("error_color")
+ val error1_color = color_value("error1_color")
+ val bad_color = color_value("bad_color")
+ val intensify_color = color_value("intensify_color")
+ val quoted_color = color_value("quoted_color")
+ val highlight_color = color_value("highlight_color")
+ val hyperlink_color = color_value("hyperlink_color")
+ val keyword1_color = color_value("keyword1_color")
+ val keyword2_color = color_value("keyword2_color")
+
+ val tfree_color = color_value("tfree_color")
+ val tvar_color = color_value("tvar_color")
+ val free_color = color_value("free_color")
+ val skolem_color = color_value("skolem_color")
+ val bound_color = color_value("bound_color")
+ val var_color = color_value("var_color")
+ val inner_numeral_color = color_value("inner_numeral_color")
+ val inner_quoted_color = color_value("inner_quoted_color")
+ val inner_comment_color = color_value("inner_comment_color")
+ val antiquotation_color = color_value("antiquotation_color")
+ val dynamic_color = color_value("dynamic_color")
+
+
+ /* command overview */
+
+ def overview_color(range: Text.Range): Option[Color] =
+ {
+ if (snapshot.is_outdated) None
+ else {
+ val results =
+ snapshot.cumulate_markup[(Protocol.Status, Int)](
+ range, (Protocol.Status.init, 0),
+ Some(Protocol.command_status_markup + Isabelle_Markup.WARNING + Isabelle_Markup.ERROR),
+ {
+ case ((status, pri), Text.Info(_, XML.Elem(markup, _))) =>
+ if (markup.name == Isabelle_Markup.WARNING)
+ (status, pri max Isabelle_Rendering.warning_pri)
+ else if (markup.name == Isabelle_Markup.ERROR)
+ (status, pri max Isabelle_Rendering.error_pri)
+ else (Protocol.command_status(status, markup), pri)
+ })
+ if (results.isEmpty) None
+ else {
+ val (status, pri) =
+ ((Protocol.Status.init, 0) /: results) {
+ case ((s1, p1), Text.Info(_, (s2, p2))) => (s1 + s2, p1 max p2) }
+
+ if (pri == Isabelle_Rendering.warning_pri) Some(warning_color)
+ else if (pri == Isabelle_Rendering.error_pri) Some(error_color)
+ else if (status.is_unprocessed) Some(unprocessed_color)
+ else if (status.is_running) Some(running_color)
+ else if (status.is_failed) Some(error_color)
+ else None
+ }
+ }
+ }
+
+
+ /* markup selectors */
+
+ private val highlight_include =
+ Set(Isabelle_Markup.SORT, Isabelle_Markup.TYP, Isabelle_Markup.TERM, Isabelle_Markup.PROP,
+ Isabelle_Markup.ML_TYPING, Isabelle_Markup.TOKEN_RANGE, Isabelle_Markup.ENTITY,
+ Isabelle_Markup.PATH, Isabelle_Markup.TYPING, Isabelle_Markup.FREE, Isabelle_Markup.SKOLEM,
+ Isabelle_Markup.BOUND, Isabelle_Markup.VAR, Isabelle_Markup.TFREE, Isabelle_Markup.TVAR,
+ Isabelle_Markup.ML_SOURCE, Isabelle_Markup.DOCUMENT_SOURCE)
+
+ def highlight(range: Text.Range): Option[Text.Info[Color]] =
+ {
+ snapshot.select_markup(range, Some(highlight_include),
+ {
+ case Text.Info(info_range, XML.Elem(Markup(name, _), _)) if highlight_include(name) =>
+ Text.Info(snapshot.convert(info_range), highlight_color)
+ }) match { case Text.Info(_, info) #:: _ => Some(info) case _ => None }
+ }
+
+
+ private val hyperlink_include = Set(Isabelle_Markup.ENTITY, Isabelle_Markup.PATH)
+
+ def hyperlink(range: Text.Range): Option[Text.Info[Hyperlink]] =
+ {
+ snapshot.cumulate_markup[List[Text.Info[Hyperlink]]](range, Nil, Some(hyperlink_include),
+ {
+ case (links, Text.Info(info_range, XML.Elem(Isabelle_Markup.Path(name), _)))
+ if Path.is_ok(name) =>
+ val jedit_file = Isabelle.thy_load.append(snapshot.node_name.dir, Path.explode(name))
+ Text.Info(snapshot.convert(info_range), Hyperlink(jedit_file, 0, 0)) :: links
+
+ case (links, Text.Info(info_range, XML.Elem(Markup(Isabelle_Markup.ENTITY, props), _)))
+ if (props.find(
+ { case (Markup.KIND, Isabelle_Markup.ML_OPEN) => true
+ case (Markup.KIND, Isabelle_Markup.ML_STRUCT) => true
+ case _ => false }).isEmpty) =>
+
+ props match {
+ case Position.Line_File(line, name) if Path.is_ok(name) =>
+ Isabelle_System.source_file(Path.explode(name)) match {
+ case Some(path) =>
+ val jedit_file = Isabelle_System.platform_path(path)
+ Text.Info(snapshot.convert(info_range), Hyperlink(jedit_file, line, 0)) :: links
+ case None => links
+ }
+
+ case Position.Id_Offset(id, offset) if !snapshot.is_outdated =>
+ snapshot.state.find_command(snapshot.version, id) match {
+ case Some((node, command)) =>
+ val sources =
+ node.commands.iterator.takeWhile(_ != command).map(_.source) ++
+ Iterator.single(command.source(Text.Range(0, command.decode(offset))))
+ val (line, column) = ((1, 1) /: sources)(Symbol.advance_line_column)
+ Text.Info(snapshot.convert(info_range),
+ Hyperlink(command.node_name.node, line, column)) :: links
+ case None => links
+ }
+
+ case _ => links
+ }
+ }) match { case Text.Info(_, info :: _) #:: _ => Some(info) case _ => None }
+ }
+
+
+ private def tooltip_text(msg: XML.Tree): String =
+ Pretty.string_of(List(msg), margin = options.int("jedit_tooltip_margin"))
+
+ def tooltip_message(range: Text.Range): Option[String] =
+ {
+ val msgs =
+ snapshot.cumulate_markup[SortedMap[Long, String]](range, SortedMap.empty,
+ Some(Set(Isabelle_Markup.WRITELN, Isabelle_Markup.WARNING, Isabelle_Markup.ERROR,
+ Isabelle_Markup.BAD)),
+ {
+ case (msgs, Text.Info(_, msg @
+ XML.Elem(Markup(markup, Isabelle_Markup.Serial(serial)), _)))
+ if markup == Isabelle_Markup.WRITELN ||
+ markup == Isabelle_Markup.WARNING ||
+ markup == Isabelle_Markup.ERROR =>
+ msgs + (serial -> tooltip_text(msg))
+ case (msgs, Text.Info(_, msg @ XML.Elem(Markup(Isabelle_Markup.BAD, _), body)))
+ if !body.isEmpty => msgs + (Document.new_id() -> tooltip_text(msg))
+ }).toList.flatMap(_.info)
+ if (msgs.isEmpty) None else Some(cat_lines(msgs.iterator.map(_._2)))
+ }
+
+
+ private val tooltips: Map[String, String] =
+ Map(
+ Isabelle_Markup.SORT -> "sort",
+ Isabelle_Markup.TYP -> "type",
+ Isabelle_Markup.TERM -> "term",
+ Isabelle_Markup.PROP -> "proposition",
+ Isabelle_Markup.TOKEN_RANGE -> "inner syntax token",
+ Isabelle_Markup.FREE -> "free variable",
+ Isabelle_Markup.SKOLEM -> "skolem variable",
+ Isabelle_Markup.BOUND -> "bound variable",
+ Isabelle_Markup.VAR -> "schematic variable",
+ Isabelle_Markup.TFREE -> "free type variable",
+ Isabelle_Markup.TVAR -> "schematic type variable",
+ Isabelle_Markup.ML_SOURCE -> "ML source",
+ Isabelle_Markup.DOCUMENT_SOURCE -> "document source")
+
+ private val tooltip_elements =
+ Set(Isabelle_Markup.ENTITY, Isabelle_Markup.TYPING, Isabelle_Markup.ML_TYPING,
+ Isabelle_Markup.PATH) ++ tooltips.keys
+
+ private def string_of_typing(kind: String, body: XML.Body): String =
+ Pretty.string_of(List(Pretty.block(XML.Text(kind) :: Pretty.Break(1) :: body)),
+ margin = options.int("jedit_tooltip_margin"))
+
+ def tooltip(range: Text.Range): Option[String] =
+ {
+ def add(prev: Text.Info[List[(Boolean, String)]], r: Text.Range, p: (Boolean, String)) =
+ if (prev.range == r) Text.Info(r, p :: prev.info) else Text.Info(r, List(p))
+
+ val tips =
+ snapshot.cumulate_markup[Text.Info[(List[(Boolean, String)])]](
+ range, Text.Info(range, Nil), Some(tooltip_elements),
+ {
+ case (prev, Text.Info(r, XML.Elem(Isabelle_Markup.Entity(kind, name), _))) =>
+ val kind1 = space_explode('_', kind).mkString(" ")
+ add(prev, r, (true, kind1 + " " + quote(name)))
+ case (prev, Text.Info(r, XML.Elem(Isabelle_Markup.Path(name), _)))
+ if Path.is_ok(name) =>
+ val jedit_file = Isabelle.thy_load.append(snapshot.node_name.dir, Path.explode(name))
+ add(prev, r, (true, "file " + quote(jedit_file)))
+ case (prev, Text.Info(r, XML.Elem(Markup(Isabelle_Markup.TYPING, _), body))) =>
+ add(prev, r, (true, string_of_typing("::", body)))
+ case (prev, Text.Info(r, XML.Elem(Markup(Isabelle_Markup.ML_TYPING, _), body))) =>
+ add(prev, r, (false, string_of_typing("ML:", body)))
+ case (prev, Text.Info(r, XML.Elem(Markup(name, _), _)))
+ if tooltips.isDefinedAt(name) => add(prev, r, (true, tooltips(name)))
+ }).toList.flatMap(_.info.info)
+
+ val all_tips =
+ (tips.filter(_._1) ++ tips.filter(!_._1).lastOption.toList).map(_._2)
+ if (all_tips.isEmpty) None else Some(cat_lines(all_tips))
+ }
+
+
+ def gutter_message(range: Text.Range): Option[Icon] =
+ {
+ val results =
+ snapshot.cumulate_markup[Int](range, 0,
+ Some(Set(Isabelle_Markup.WARNING, Isabelle_Markup.ERROR)),
+ {
+ case (pri, Text.Info(_, XML.Elem(Markup(Isabelle_Markup.WARNING, _), body))) =>
+ body match {
+ case List(XML.Elem(Markup(Isabelle_Markup.LEGACY, _), _)) =>
+ pri max Isabelle_Rendering.legacy_pri
+ case _ => pri max Isabelle_Rendering.warning_pri
+ }
+ case (pri, Text.Info(_, XML.Elem(Markup(Isabelle_Markup.ERROR, _), _))) =>
+ pri max Isabelle_Rendering.error_pri
+ })
+ val pri = (0 /: results) { case (p1, Text.Info(_, p2)) => p1 max p2 }
+ Isabelle_Rendering.gutter_icons.get(pri)
+ }
+
+
+ private val squiggly_colors = Map(
+ Isabelle_Rendering.writeln_pri -> writeln_color,
+ Isabelle_Rendering.warning_pri -> warning_color,
+ Isabelle_Rendering.error_pri -> error_color)
+
+ def squiggly_underline(range: Text.Range): Stream[Text.Info[Color]] =
+ {
+ val results =
+ snapshot.cumulate_markup[Int](range, 0,
+ Some(Set(Isabelle_Markup.WRITELN, Isabelle_Markup.WARNING, Isabelle_Markup.ERROR)),
+ {
+ case (pri, Text.Info(_, XML.Elem(Markup(name, _), _))) =>
+ name match {
+ case Isabelle_Markup.WRITELN => pri max Isabelle_Rendering.writeln_pri
+ case Isabelle_Markup.WARNING => pri max Isabelle_Rendering.warning_pri
+ case Isabelle_Markup.ERROR => pri max Isabelle_Rendering.error_pri
+ case _ => pri
+ }
+ })
+ for {
+ Text.Info(r, pri) <- results
+ color <- squiggly_colors.get(pri)
+ } yield Text.Info(r, color)
+ }
+
+
+ def background1(range: Text.Range): Stream[Text.Info[Color]] =
+ {
+ if (snapshot.is_outdated) Stream(Text.Info(range, outdated_color))
+ else
+ for {
+ Text.Info(r, result) <-
+ snapshot.cumulate_markup[(Option[Protocol.Status], Option[Color])](
+ range, (Some(Protocol.Status.init), None),
+ Some(Protocol.command_status_markup + Isabelle_Markup.BAD + Isabelle_Markup.INTENSIFY),
+ {
+ case (((Some(status), color), Text.Info(_, XML.Elem(markup, _))))
+ if (Protocol.command_status_markup(markup.name)) =>
+ (Some(Protocol.command_status(status, markup)), color)
+ case (_, Text.Info(_, XML.Elem(Markup(Isabelle_Markup.BAD, _), _))) =>
+ (None, Some(bad_color))
+ case (_, Text.Info(_, XML.Elem(Markup(Isabelle_Markup.INTENSIFY, _), _))) =>
+ (None, Some(intensify_color))
+ })
+ color <-
+ (result match {
+ case (Some(status), opt_color) =>
+ if (status.is_unprocessed) Some(unprocessed1_color)
+ else if (status.is_running) Some(running1_color)
+ else opt_color
+ case (_, opt_color) => opt_color
+ })
+ } yield Text.Info(r, color)
+ }
+
+
+ def background2(range: Text.Range): Stream[Text.Info[Color]] =
+ snapshot.select_markup(range,
+ Some(Set(Isabelle_Markup.TOKEN_RANGE)),
+ {
+ case Text.Info(_, XML.Elem(Markup(Isabelle_Markup.TOKEN_RANGE, _), _)) => light_color
+ })
+
+
+ def foreground(range: Text.Range): Stream[Text.Info[Color]] =
+ snapshot.select_markup(range,
+ Some(Set(Isabelle_Markup.STRING, Isabelle_Markup.ALTSTRING, Isabelle_Markup.VERBATIM)),
+ {
+ case Text.Info(_, XML.Elem(Markup(Isabelle_Markup.STRING, _), _)) => quoted_color
+ case Text.Info(_, XML.Elem(Markup(Isabelle_Markup.ALTSTRING, _), _)) => quoted_color
+ case Text.Info(_, XML.Elem(Markup(Isabelle_Markup.VERBATIM, _), _)) => quoted_color
+ })
+
+
+ private val text_colors: Map[String, Color] = Map(
+ Isabelle_Markup.STRING -> Color.BLACK,
+ Isabelle_Markup.ALTSTRING -> Color.BLACK,
+ Isabelle_Markup.VERBATIM -> Color.BLACK,
+ Isabelle_Markup.LITERAL -> keyword1_color,
+ Isabelle_Markup.DELIMITER -> Color.BLACK,
+ Isabelle_Markup.TFREE -> tfree_color,
+ Isabelle_Markup.TVAR -> tvar_color,
+ Isabelle_Markup.FREE -> free_color,
+ Isabelle_Markup.SKOLEM -> skolem_color,
+ Isabelle_Markup.BOUND -> bound_color,
+ Isabelle_Markup.VAR -> var_color,
+ Isabelle_Markup.INNER_STRING -> inner_quoted_color,
+ Isabelle_Markup.INNER_COMMENT -> inner_comment_color,
+ Isabelle_Markup.DYNAMIC_FACT -> dynamic_color,
+ Isabelle_Markup.ML_KEYWORD -> keyword1_color,
+ Isabelle_Markup.ML_DELIMITER -> Color.BLACK,
+ Isabelle_Markup.ML_NUMERAL -> inner_numeral_color,
+ Isabelle_Markup.ML_CHAR -> inner_quoted_color,
+ Isabelle_Markup.ML_STRING -> inner_quoted_color,
+ Isabelle_Markup.ML_COMMENT -> inner_comment_color,
+ Isabelle_Markup.ANTIQ -> antiquotation_color)
+
+ private val text_color_elements = Set.empty[String] ++ text_colors.keys
+
+ def text_color(range: Text.Range, color: Color)
+ : Stream[Text.Info[Color]] =
+ {
+ snapshot.cumulate_markup(range, color, Some(text_color_elements),
+ {
+ case (_, Text.Info(_, XML.Elem(Markup(m, _), _)))
+ if text_colors.isDefinedAt(m) => text_colors(m)
+ })
+ }
+}
--- a/src/Tools/jEdit/src/jedit_options.scala Fri Sep 14 12:09:27 2012 +0200
+++ b/src/Tools/jEdit/src/jedit_options.scala Fri Sep 14 21:23:06 2012 +0200
@@ -9,6 +9,7 @@
import isabelle._
+import java.awt.Color
import javax.swing.{InputVerifier, JComponent, UIManager}
import javax.swing.text.JTextComponent
@@ -26,6 +27,8 @@
class JEdit_Options extends Options_Variable
{
+ def color_value(s: String): Color = Color_Value(string(s))
+
def make_color_component(opt: Options.Opt): Option_Component =
{
Swing_Thread.require()
--- a/src/Tools/jEdit/src/output_dockable.scala Fri Sep 14 12:09:27 2012 +0200
+++ b/src/Tools/jEdit/src/output_dockable.scala Fri Sep 14 21:23:06 2012 +0200
@@ -25,6 +25,18 @@
{
Swing_Thread.require()
+
+ /* component state -- owned by Swing thread */
+
+ private var zoom_factor = 100
+ private var show_tracing = false
+ private var do_update = true
+ private var current_state = Command.empty.empty_state
+ private var current_body: XML.Body = Nil
+
+
+ /* HTML panel */
+
private val html_panel =
new HTML_Panel(Isabelle.font_family(), scala.math.round(Isabelle.font_size()))
{
@@ -36,23 +48,20 @@
Document_View(view.getTextArea) match {
case Some(doc_view) =>
doc_view.robust_body() {
- current_command match {
- case Some(cmd) =>
- val model = doc_view.model
- val buffer = model.buffer
- val snapshot = model.snapshot()
- snapshot.node.command_start(cmd) match {
- case Some(start) if !snapshot.is_outdated =>
- val text = Pretty.string_of(sendback)
- try {
- buffer.beginCompoundEdit()
- buffer.remove(start, cmd.proper_range.length)
- buffer.insert(start, text)
- }
- finally { buffer.endCompoundEdit() }
- case _ =>
+ val cmd = current_state.command
+ val model = doc_view.model
+ val buffer = model.buffer
+ val snapshot = model.snapshot()
+ snapshot.node.command_start(cmd) match {
+ case Some(start) if !snapshot.is_outdated =>
+ val text = Pretty.string_of(sendback)
+ try {
+ buffer.beginCompoundEdit()
+ buffer.remove(start, cmd.proper_range.length)
+ buffer.insert(start, text)
}
- case None =>
+ finally { buffer.endCompoundEdit() }
+ case _ =>
}
}
case None =>
@@ -63,55 +72,45 @@
set_content(html_panel)
- /* component state -- owned by Swing thread */
-
- private var zoom_factor = 100
- private var show_tracing = false
- private var follow_caret = true
- private var current_command: Option[Command] = None
-
-
private def handle_resize()
{
- Swing_Thread.now {
- html_panel.resize(Isabelle.font_family(),
- scala.math.round(Isabelle.font_size() * zoom_factor / 100))
- }
+ Swing_Thread.require()
+
+ html_panel.resize(Isabelle.font_family(),
+ scala.math.round(Isabelle.font_size() * zoom_factor / 100))
}
- private def handle_perspective(): Boolean =
- Swing_Thread.now {
- Document_View(view.getTextArea) match {
- case Some(doc_view) =>
- val cmd = doc_view.selected_command()
- if (current_command == cmd) false
- else { current_command = cmd; true }
- case None => false
- }
- }
-
- private def handle_update(restriction: Option[Set[Command]] = None)
+ private def handle_update(follow: Boolean, restriction: Option[Set[Command]])
{
- Swing_Thread.now {
- if (follow_caret) handle_perspective()
- Document_View(view.getTextArea) match {
- case Some(doc_view) =>
- current_command match {
- case Some(cmd) if !restriction.isDefined || restriction.get.contains(cmd) =>
- val snapshot = doc_view.model.snapshot()
- val filtered_results =
- snapshot.state.command_state(snapshot.version, cmd).results.iterator
- .map(_._2).filter(
- { // FIXME not scalable
- case XML.Elem(Markup(Isabelle_Markup.TRACING, _), _) => show_tracing
- case _ => true
- }).toList
- html_panel.render(filtered_results)
- case _ => html_panel.render(Nil)
- }
- case None => html_panel.render(Nil)
+ Swing_Thread.require()
+
+ val new_state =
+ if (follow) {
+ Document_View(view.getTextArea) match {
+ case Some(doc_view) =>
+ val snapshot = doc_view.model.snapshot()
+ snapshot.node.command_at(doc_view.text_area.getCaretPosition).map(_._1) match {
+ case Some(cmd) => snapshot.state.command_state(snapshot.version, cmd)
+ case None => Command.empty.empty_state
+ }
+ case None => Command.empty.empty_state
+ }
}
- }
+ else current_state
+
+ val new_body =
+ if (!restriction.isDefined || restriction.get.contains(new_state.command))
+ new_state.results.iterator.map(_._2).filter(
+ { // FIXME not scalable
+ case XML.Elem(Markup(Isabelle_Markup.TRACING, _), _) => show_tracing
+ case _ => true
+ }).toList
+ else current_body
+
+ if (new_body != current_body) html_panel.render(new_body)
+
+ current_state = new_state
+ current_body = new_body
}
@@ -120,9 +119,12 @@
private val main_actor = actor {
loop {
react {
- case Session.Global_Settings => handle_resize()
- case changed: Session.Commands_Changed => handle_update(Some(changed.commands))
- case Session.Caret_Focus => if (follow_caret && handle_perspective()) handle_update()
+ case Session.Global_Settings =>
+ Swing_Thread.later { handle_resize() }
+ case changed: Session.Commands_Changed =>
+ Swing_Thread.later { handle_update(do_update, Some(changed.commands)) }
+ case Session.Caret_Focus =>
+ Swing_Thread.later { handle_update(do_update, None) }
case bad => System.err.println("Output_Dockable: ignoring bad message " + bad)
}
}
@@ -130,14 +132,18 @@
override def init()
{
+ Swing_Thread.require()
+
Isabelle.session.global_settings += main_actor
Isabelle.session.commands_changed += main_actor
Isabelle.session.caret_focus += main_actor
- handle_update()
+ handle_update(true, None)
}
override def exit()
{
+ Swing_Thread.require()
+
Isabelle.session.global_settings -= main_actor
Isabelle.session.commands_changed -= main_actor
Isabelle.session.caret_focus -= main_actor
@@ -162,19 +168,21 @@
zoom.tooltip = "Zoom factor for basic font size"
private val tracing = new CheckBox("Tracing") {
- reactions += { case ButtonClicked(_) => show_tracing = this.selected; handle_update() }
+ reactions += {
+ case ButtonClicked(_) => show_tracing = this.selected; handle_update(do_update, None) }
}
tracing.selected = show_tracing
tracing.tooltip = "Indicate output of tracing messages"
private val auto_update = new CheckBox("Auto update") {
- reactions += { case ButtonClicked(_) => follow_caret = this.selected; handle_update() }
+ reactions += {
+ case ButtonClicked(_) => do_update = this.selected; handle_update(do_update, None) }
}
- auto_update.selected = follow_caret
+ auto_update.selected = do_update
auto_update.tooltip = "Indicate automatic update following cursor movement"
private val update = new Button("Update") {
- reactions += { case ButtonClicked(_) => handle_perspective(); handle_update() }
+ reactions += { case ButtonClicked(_) => handle_update(true, None) }
}
update.tooltip = "Update display according to the command at cursor position"
--- a/src/Tools/jEdit/src/session_dockable.scala Fri Sep 14 12:09:27 2012 +0200
+++ b/src/Tools/jEdit/src/session_dockable.scala Fri Sep 14 21:23:06 2012 +0200
@@ -89,10 +89,10 @@
var end = size.width - insets.right
for {
(n, color) <- List(
- (st.unprocessed, Isabelle_Rendering.color_value("color_unprocessed1")),
- (st.running, Isabelle_Rendering.color_value("color_running")),
- (st.warned, Isabelle_Rendering.color_value("color_warning")),
- (st.failed, Isabelle_Rendering.color_value("color_error"))) }
+ (st.unprocessed, Isabelle.options.color_value("unprocessed1_color")),
+ (st.running, Isabelle.options.color_value("running_color")),
+ (st.warned, Isabelle.options.color_value("warning_color")),
+ (st.failed, Isabelle.options.color_value("error_color"))) }
{
gfx.setColor(color)
val v = (n * w / st.total) max (if (n > 0) 2 else 0)
--- a/src/Tools/jEdit/src/text_area_painter.scala Fri Sep 14 12:09:27 2012 +0200
+++ b/src/Tools/jEdit/src/text_area_painter.scala Fri Sep 14 21:23:06 2012 +0200
@@ -75,7 +75,7 @@
/* common painter state */
- @volatile private var painter_snapshot: Document.Snapshot = null
+ @volatile private var painter_rendering: Isabelle_Rendering = null
@volatile private var painter_clip: Shape = null
private val set_state = new TextAreaExtension
@@ -84,7 +84,7 @@
first_line: Int, last_line: Int, physical_lines: Array[Int],
start: Array[Int], end: Array[Int], y: Int, line_height: Int)
{
- painter_snapshot = model.snapshot()
+ painter_rendering = Isabelle_Rendering(model.snapshot(), Isabelle.options.value)
painter_clip = gfx.getClip
}
}
@@ -95,14 +95,14 @@
first_line: Int, last_line: Int, physical_lines: Array[Int],
start: Array[Int], end: Array[Int], y: Int, line_height: Int)
{
- painter_snapshot = null
+ painter_rendering = null
painter_clip = null
}
}
- private def robust_snapshot(body: Document.Snapshot => Unit)
+ private def robust_rendering(body: Isabelle_Rendering => Unit)
{
- doc_view.robust_body(()) { body(painter_snapshot) }
+ doc_view.robust_body(()) { body(painter_rendering) }
}
@@ -114,7 +114,7 @@
first_line: Int, last_line: Int, physical_lines: Array[Int],
start: Array[Int], end: Array[Int], y: Int, line_height: Int)
{
- robust_snapshot { snapshot =>
+ robust_rendering { rendering =>
val ascent = text_area.getPainter.getFontMetrics.getAscent
for (i <- 0 until physical_lines.length) {
@@ -123,7 +123,7 @@
// background color (1)
for {
- Text.Info(range, color) <- Isabelle_Rendering.background1(snapshot, line_range)
+ Text.Info(range, color) <- rendering.background1(line_range)
r <- gfx_range(range)
} {
gfx.setColor(color)
@@ -132,7 +132,7 @@
// background color (2)
for {
- Text.Info(range, color) <- Isabelle_Rendering.background2(snapshot, line_range)
+ Text.Info(range, color) <- rendering.background2(line_range)
r <- gfx_range(range)
} {
gfx.setColor(color)
@@ -141,7 +141,7 @@
// squiggly underline
for {
- Text.Info(range, color) <- Isabelle_Rendering.squiggly_underline(snapshot, line_range)
+ Text.Info(range, color) <- rendering.squiggly_underline(line_range)
r <- gfx_range(range)
} {
gfx.setColor(color)
@@ -161,7 +161,7 @@
/* text */
- private def paint_chunk_list(snapshot: Document.Snapshot,
+ private def paint_chunk_list(rendering: Isabelle_Rendering,
gfx: Graphics2D, line_start: Text.Offset, head: Chunk, x: Float, y: Float): Float =
{
val clip_rect = gfx.getClipBounds
@@ -185,12 +185,12 @@
else chunk_font.getStringBounds(s, font_context).getWidth.toFloat
val caret_range =
- if (text_area.isCaretVisible) doc_view.caret_range()
+ if (text_area.isCaretVisible) model.point_range(text_area.getCaretPosition)
else Text.Range(-1)
val markup =
for {
- r1 <- Isabelle_Rendering.text_color(snapshot, chunk_range, chunk_color)
+ r1 <- rendering.text_color(chunk_range, chunk_color)
r2 <- r1.try_restrict(chunk_range)
} yield r2
@@ -246,7 +246,7 @@
first_line: Int, last_line: Int, physical_lines: Array[Int],
start: Array[Int], end: Array[Int], y: Int, line_height: Int)
{
- robust_snapshot { snapshot =>
+ robust_rendering { rendering =>
val clip = gfx.getClip
val x0 = text_area.getHorizontalOffset
val fm = text_area.getPainter.getFontMetrics
@@ -260,7 +260,7 @@
if (chunks != null) {
val line_start = text_area.getBuffer.getLineStartOffset(line)
gfx.clipRect(x0, y + line_height * i, Integer.MAX_VALUE, line_height)
- val w = paint_chunk_list(snapshot, gfx, line_start, chunks, x0, y0).toInt
+ val w = paint_chunk_list(rendering, gfx, line_start, chunks, x0, y0).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)
@@ -282,23 +282,23 @@
first_line: Int, last_line: Int, physical_lines: Array[Int],
start: Array[Int], end: Array[Int], y: Int, line_height: Int)
{
- robust_snapshot { snapshot =>
+ robust_rendering { rendering =>
for (i <- 0 until physical_lines.length) {
if (physical_lines(i) != -1) {
val line_range = doc_view.proper_line_range(start(i), end(i))
// foreground color
for {
- Text.Info(range, color) <- Isabelle_Rendering.foreground(snapshot, line_range)
+ Text.Info(range, color) <- rendering.foreground(line_range)
r <- gfx_range(range)
} {
gfx.setColor(color)
gfx.fillRect(r.x, y + i * line_height, r.length, line_height)
}
- // highlighted range -- potentially from other snapshot
+ // highlight range -- potentially from other snapshot
for {
- info <- doc_view.highlight_range()
+ info <- doc_view.highlight_info()
Text.Info(range, color) <- info.try_restrict(line_range)
r <- gfx_range(range)
} {
@@ -308,11 +308,11 @@
// hyperlink range -- potentially from other snapshot
for {
- info <- doc_view.hyperlink_range()
+ info <- doc_view.hyperlink_info()
Text.Info(range, _) <- info.try_restrict(line_range)
r <- gfx_range(range)
} {
- gfx.setColor(Isabelle_Rendering.color_value("color_hyperlink"))
+ gfx.setColor(rendering.hyperlink_color)
gfx.drawRect(r.x, y + i * line_height, r.length - 1, line_height - 1)
}
}
@@ -329,7 +329,7 @@
override def paintValidLine(gfx: Graphics2D,
screen_line: Int, physical_line: Int, start: Int, end: Int, y: Int)
{
- robust_snapshot { _ =>
+ robust_rendering { _ =>
if (before) gfx.clipRect(0, 0, 0, 0)
else gfx.setClip(painter_clip)
}
@@ -346,7 +346,7 @@
override def paintValidLine(gfx: Graphics2D,
screen_line: Int, physical_line: Int, start: Int, end: Int, y: Int)
{
- robust_snapshot { _ =>
+ robust_rendering { _ =>
if (text_area.isCaretVisible) {
val caret = text_area.getCaretPosition
if (start <= caret && caret == end - 1) {
--- a/src/Tools/jEdit/src/text_overview.scala Fri Sep 14 12:09:27 2012 +0200
+++ b/src/Tools/jEdit/src/text_overview.scala Fri Sep 14 21:23:06 2012 +0200
@@ -54,6 +54,7 @@
private var cached_colors: List[(Color, Int, Int)] = Nil
private var last_snapshot = Document.State.init.snapshot()
+ private var last_options = Isabelle.options.value
private var last_line_count = 0
private var last_char_count = 0
private var last_L = 0
@@ -69,7 +70,7 @@
val snapshot = doc_view.model.snapshot()
if (snapshot.is_outdated) {
- gfx.setColor(Isabelle_Rendering.color_value("color_outdated"))
+ gfx.setColor(Isabelle.options.color_value("outdated_color"))
gfx.asInstanceOf[Graphics2D].fill(gfx.getClipBounds)
}
else {
@@ -82,9 +83,14 @@
val L = lines()
val H = getHeight()
+ val options = Isabelle.options.value
+
if (!(line_count == last_line_count && char_count == last_char_count &&
- L == last_L && H == last_H && (snapshot eq_markup last_snapshot)))
+ L == last_L && H == last_H && (snapshot eq_markup last_snapshot) &&
+ (options eq last_options)))
{
+ val rendering = Isabelle_Rendering(snapshot, options)
+
@tailrec def loop(l: Int, h: Int, p: Int, q: Int, colors: List[(Color, Int, Int)])
: List[(Color, Int, Int)] =
{
@@ -102,7 +108,7 @@
val range = Text.Range(start, end)
val colors1 =
- (Isabelle_Rendering.overview_color(snapshot, range), colors) match {
+ (rendering.overview_color(range), 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
@@ -115,6 +121,7 @@
cached_colors = loop(0, 0, 0, 0, Nil)
last_snapshot = snapshot
+ last_options = options
last_line_count = line_count
last_char_count = char_count
last_L = L
--- a/src/Tools/quickcheck.ML Fri Sep 14 12:09:27 2012 +0200
+++ b/src/Tools/quickcheck.ML Fri Sep 14 21:23:06 2012 +0200
@@ -529,7 +529,7 @@
state
|> (if auto then
Proof.goal_message (K (Pretty.chunks [Pretty.str "",
- Pretty.mark Isabelle_Markup.hilite msg]))
+ Pretty.mark Isabelle_Markup.intensify msg]))
else
tap (fn _ => Output.urgent_message (Pretty.string_of msg))))
else
--- a/src/Tools/solve_direct.ML Fri Sep 14 12:09:27 2012 +0200
+++ b/src/Tools/solve_direct.ML Fri Sep 14 21:23:06 2012 +0200
@@ -89,7 +89,7 @@
Proof.goal_message
(fn () =>
Pretty.chunks
- [Pretty.str "", Pretty.markup Isabelle_Markup.hilite (message results)])
+ [Pretty.str "", Pretty.markup Isabelle_Markup.intensify (message results)])
else
tap (fn _ =>
Output.urgent_message (Pretty.string_of (Pretty.chunks (message results))))))