wenzelm@50202
|
1 |
/* Title: Tools/jEdit/src/rendering.scala
|
wenzelm@39178
|
2 |
Author: Makarius
|
wenzelm@39178
|
3 |
|
wenzelm@50199
|
4 |
Isabelle-specific implementation of quasi-abstract rendering and
|
wenzelm@50199
|
5 |
markup interpretation.
|
wenzelm@39178
|
6 |
*/
|
wenzelm@39178
|
7 |
|
wenzelm@39178
|
8 |
package isabelle.jedit
|
wenzelm@39178
|
9 |
|
wenzelm@39178
|
10 |
|
wenzelm@39178
|
11 |
import isabelle._
|
wenzelm@39178
|
12 |
|
wenzelm@39178
|
13 |
import java.awt.Color
|
wenzelm@46227
|
14 |
import javax.swing.Icon
|
wenzelm@39178
|
15 |
|
wenzelm@43414
|
16 |
import org.gjt.sp.jedit.syntax.{Token => JEditToken}
|
wenzelm@50206
|
17 |
import org.gjt.sp.jedit.{jEdit, GUIUtilities}
|
wenzelm@39178
|
18 |
|
wenzelm@45460
|
19 |
import scala.collection.immutable.SortedMap
|
wenzelm@45460
|
20 |
|
wenzelm@39178
|
21 |
|
wenzelm@50199
|
22 |
object Rendering
|
wenzelm@39178
|
23 |
{
|
wenzelm@50199
|
24 |
def apply(snapshot: Document.Snapshot, options: Options): Rendering =
|
wenzelm@50199
|
25 |
new Rendering(snapshot, options)
|
wenzelm@39178
|
26 |
|
wenzelm@49356
|
27 |
|
wenzelm@49406
|
28 |
/* message priorities */
|
wenzelm@43388
|
29 |
|
wenzelm@46227
|
30 |
private val writeln_pri = 1
|
wenzelm@49474
|
31 |
private val tracing_pri = 2
|
wenzelm@49474
|
32 |
private val warning_pri = 3
|
wenzelm@49474
|
33 |
private val legacy_pri = 4
|
wenzelm@49474
|
34 |
private val error_pri = 5
|
wenzelm@49474
|
35 |
|
wenzelm@49474
|
36 |
private val message_pri = Map(
|
wenzelm@50201
|
37 |
Markup.WRITELN -> writeln_pri, Markup.WRITELN_MESSAGE -> writeln_pri,
|
wenzelm@50201
|
38 |
Markup.TRACING -> tracing_pri, Markup.TRACING_MESSAGE -> tracing_pri,
|
wenzelm@50201
|
39 |
Markup.WARNING -> warning_pri, Markup.WARNING_MESSAGE -> warning_pri,
|
wenzelm@50201
|
40 |
Markup.ERROR -> error_pri, Markup.ERROR_MESSAGE -> error_pri)
|
wenzelm@39178
|
41 |
|
wenzelm@49406
|
42 |
|
wenzelm@49406
|
43 |
/* icons */
|
wenzelm@49406
|
44 |
|
wenzelm@49406
|
45 |
private def load_icon(name: String): Icon =
|
wenzelm@49406
|
46 |
{
|
wenzelm@49406
|
47 |
val icon = GUIUtilities.loadIcon(name)
|
wenzelm@50200
|
48 |
if (icon.getIconWidth < 0 || icon.getIconHeight < 0) error("Bad icon: " + name)
|
wenzelm@50200
|
49 |
else icon
|
wenzelm@49406
|
50 |
}
|
wenzelm@49406
|
51 |
|
wenzelm@46234
|
52 |
private val gutter_icons = Map(
|
wenzelm@49406
|
53 |
warning_pri -> load_icon("16x16/status/dialog-information.png"),
|
wenzelm@49406
|
54 |
legacy_pri -> load_icon("16x16/status/dialog-warning.png"),
|
wenzelm@49406
|
55 |
error_pri -> load_icon("16x16/status/dialog-error.png"))
|
wenzelm@46234
|
56 |
|
wenzelm@49727
|
57 |
val tooltip_close_icon = load_icon("16x16/actions/document-close.png")
|
wenzelm@49726
|
58 |
val tooltip_detach_icon = load_icon("16x16/actions/window-new.png")
|
wenzelm@49725
|
59 |
|
wenzelm@46234
|
60 |
|
wenzelm@50206
|
61 |
/* jEdit font */
|
wenzelm@50206
|
62 |
|
wenzelm@50206
|
63 |
def font_family(): String = jEdit.getProperty("view.font")
|
wenzelm@50206
|
64 |
|
wenzelm@50206
|
65 |
def font_size(scale: String): Float =
|
wenzelm@50206
|
66 |
(jEdit.getIntegerProperty("view.fontsize", 16) * PIDE.options.real(scale)).toFloat
|
wenzelm@50206
|
67 |
|
wenzelm@50206
|
68 |
|
wenzelm@39179
|
69 |
/* token markup -- text styles */
|
wenzelm@39179
|
70 |
|
wenzelm@39179
|
71 |
private val command_style: Map[String, Byte] =
|
wenzelm@39179
|
72 |
{
|
wenzelm@43414
|
73 |
import JEditToken._
|
wenzelm@39179
|
74 |
Map[String, Byte](
|
wenzelm@39179
|
75 |
Keyword.THY_END -> KEYWORD2,
|
wenzelm@39179
|
76 |
Keyword.THY_SCRIPT -> LABEL,
|
wenzelm@50643
|
77 |
Keyword.PRF_SCRIPT -> DIGIT,
|
wenzelm@39179
|
78 |
Keyword.PRF_ASM -> KEYWORD3,
|
wenzelm@39179
|
79 |
Keyword.PRF_ASM_GOAL -> KEYWORD3
|
wenzelm@39179
|
80 |
).withDefaultValue(KEYWORD1)
|
wenzelm@39179
|
81 |
}
|
wenzelm@39179
|
82 |
|
wenzelm@43414
|
83 |
private val token_style: Map[Token.Kind.Value, Byte] =
|
wenzelm@39179
|
84 |
{
|
wenzelm@43414
|
85 |
import JEditToken._
|
wenzelm@43414
|
86 |
Map[Token.Kind.Value, Byte](
|
wenzelm@43414
|
87 |
Token.Kind.KEYWORD -> KEYWORD2,
|
wenzelm@43414
|
88 |
Token.Kind.IDENT -> NULL,
|
wenzelm@43414
|
89 |
Token.Kind.LONG_IDENT -> NULL,
|
wenzelm@43414
|
90 |
Token.Kind.SYM_IDENT -> NULL,
|
wenzelm@43414
|
91 |
Token.Kind.VAR -> NULL,
|
wenzelm@43414
|
92 |
Token.Kind.TYPE_IDENT -> NULL,
|
wenzelm@43414
|
93 |
Token.Kind.TYPE_VAR -> NULL,
|
wenzelm@43414
|
94 |
Token.Kind.NAT -> NULL,
|
wenzelm@43414
|
95 |
Token.Kind.FLOAT -> NULL,
|
wenzelm@43431
|
96 |
Token.Kind.STRING -> LITERAL1,
|
wenzelm@43431
|
97 |
Token.Kind.ALT_STRING -> LITERAL2,
|
wenzelm@43414
|
98 |
Token.Kind.VERBATIM -> COMMENT3,
|
wenzelm@43414
|
99 |
Token.Kind.SPACE -> NULL,
|
wenzelm@43414
|
100 |
Token.Kind.COMMENT -> COMMENT1,
|
wenzelm@48754
|
101 |
Token.Kind.ERROR -> INVALID
|
wenzelm@39179
|
102 |
).withDefaultValue(NULL)
|
wenzelm@39179
|
103 |
}
|
wenzelm@39179
|
104 |
|
wenzelm@43414
|
105 |
def token_markup(syntax: Outer_Syntax, token: Token): Byte =
|
wenzelm@43430
|
106 |
if (token.is_command) command_style(syntax.keyword_kind(token.content).getOrElse(""))
|
wenzelm@43430
|
107 |
else if (token.is_operator) JEditToken.OPERATOR
|
wenzelm@43414
|
108 |
else token_style(token.kind)
|
wenzelm@39178
|
109 |
}
|
wenzelm@49356
|
110 |
|
wenzelm@49356
|
111 |
|
wenzelm@50199
|
112 |
class Rendering private(val snapshot: Document.Snapshot, val options: Options)
|
wenzelm@49356
|
113 |
{
|
wenzelm@49356
|
114 |
/* colors */
|
wenzelm@49356
|
115 |
|
wenzelm@49356
|
116 |
def color_value(s: String): Color = Color_Value(options.string(s))
|
wenzelm@49356
|
117 |
|
wenzelm@49356
|
118 |
val outdated_color = color_value("outdated_color")
|
wenzelm@49356
|
119 |
val unprocessed_color = color_value("unprocessed_color")
|
wenzelm@49356
|
120 |
val unprocessed1_color = color_value("unprocessed1_color")
|
wenzelm@49356
|
121 |
val running_color = color_value("running_color")
|
wenzelm@49356
|
122 |
val running1_color = color_value("running1_color")
|
wenzelm@49356
|
123 |
val light_color = color_value("light_color")
|
wenzelm@49706
|
124 |
val tooltip_color = color_value("tooltip_color")
|
wenzelm@49356
|
125 |
val writeln_color = color_value("writeln_color")
|
wenzelm@49356
|
126 |
val warning_color = color_value("warning_color")
|
wenzelm@49356
|
127 |
val error_color = color_value("error_color")
|
wenzelm@49356
|
128 |
val error1_color = color_value("error1_color")
|
wenzelm@49418
|
129 |
val writeln_message_color = color_value("writeln_message_color")
|
wenzelm@49418
|
130 |
val tracing_message_color = color_value("tracing_message_color")
|
wenzelm@49418
|
131 |
val warning_message_color = color_value("warning_message_color")
|
wenzelm@49418
|
132 |
val error_message_color = color_value("error_message_color")
|
wenzelm@49356
|
133 |
val bad_color = color_value("bad_color")
|
wenzelm@49358
|
134 |
val intensify_color = color_value("intensify_color")
|
wenzelm@49356
|
135 |
val quoted_color = color_value("quoted_color")
|
wenzelm@49358
|
136 |
val highlight_color = color_value("highlight_color")
|
wenzelm@49356
|
137 |
val hyperlink_color = color_value("hyperlink_color")
|
wenzelm@50450
|
138 |
val active_color = color_value("active_color")
|
wenzelm@50450
|
139 |
val active_hover_color = color_value("active_hover_color")
|
wenzelm@50498
|
140 |
val active_result_color = color_value("active_result_color")
|
wenzelm@49356
|
141 |
val keyword1_color = color_value("keyword1_color")
|
wenzelm@49356
|
142 |
val keyword2_color = color_value("keyword2_color")
|
wenzelm@49356
|
143 |
|
wenzelm@49356
|
144 |
val tfree_color = color_value("tfree_color")
|
wenzelm@49356
|
145 |
val tvar_color = color_value("tvar_color")
|
wenzelm@49356
|
146 |
val free_color = color_value("free_color")
|
wenzelm@49356
|
147 |
val skolem_color = color_value("skolem_color")
|
wenzelm@49356
|
148 |
val bound_color = color_value("bound_color")
|
wenzelm@49356
|
149 |
val var_color = color_value("var_color")
|
wenzelm@49356
|
150 |
val inner_numeral_color = color_value("inner_numeral_color")
|
wenzelm@49356
|
151 |
val inner_quoted_color = color_value("inner_quoted_color")
|
wenzelm@49356
|
152 |
val inner_comment_color = color_value("inner_comment_color")
|
wenzelm@49356
|
153 |
val antiquotation_color = color_value("antiquotation_color")
|
wenzelm@49356
|
154 |
val dynamic_color = color_value("dynamic_color")
|
wenzelm@49356
|
155 |
|
wenzelm@49356
|
156 |
|
wenzelm@49356
|
157 |
/* command overview */
|
wenzelm@49356
|
158 |
|
wenzelm@49969
|
159 |
val overview_limit = options.int("jedit_text_overview_limit")
|
wenzelm@49969
|
160 |
|
wenzelm@50547
|
161 |
private val overview_include = Protocol.command_status_markup + Markup.WARNING + Markup.ERROR
|
wenzelm@50547
|
162 |
|
wenzelm@49356
|
163 |
def overview_color(range: Text.Range): Option[Color] =
|
wenzelm@49356
|
164 |
{
|
wenzelm@49356
|
165 |
if (snapshot.is_outdated) None
|
wenzelm@49356
|
166 |
else {
|
wenzelm@49356
|
167 |
val results =
|
wenzelm@49356
|
168 |
snapshot.cumulate_markup[(Protocol.Status, Int)](
|
wenzelm@50547
|
169 |
range, (Protocol.Status.init, 0), Some(overview_include), _ =>
|
wenzelm@49356
|
170 |
{
|
wenzelm@50547
|
171 |
case ((status, pri), Text.Info(_, XML.Elem(markup, _)))
|
wenzelm@50547
|
172 |
if overview_include(markup.name) =>
|
wenzelm@50201
|
173 |
if (markup.name == Markup.WARNING || markup.name == Markup.ERROR)
|
wenzelm@50199
|
174 |
(status, pri max Rendering.message_pri(markup.name))
|
wenzelm@49356
|
175 |
else (Protocol.command_status(status, markup), pri)
|
wenzelm@49356
|
176 |
})
|
wenzelm@49356
|
177 |
if (results.isEmpty) None
|
wenzelm@49356
|
178 |
else {
|
wenzelm@49356
|
179 |
val (status, pri) =
|
wenzelm@49356
|
180 |
((Protocol.Status.init, 0) /: results) {
|
wenzelm@49356
|
181 |
case ((s1, p1), Text.Info(_, (s2, p2))) => (s1 + s2, p1 max p2) }
|
wenzelm@49356
|
182 |
|
wenzelm@50199
|
183 |
if (pri == Rendering.warning_pri) Some(warning_color)
|
wenzelm@50199
|
184 |
else if (pri == Rendering.error_pri) Some(error_color)
|
wenzelm@49356
|
185 |
else if (status.is_unprocessed) Some(unprocessed_color)
|
wenzelm@49356
|
186 |
else if (status.is_running) Some(running_color)
|
wenzelm@49356
|
187 |
else if (status.is_failed) Some(error_color)
|
wenzelm@49356
|
188 |
else None
|
wenzelm@49356
|
189 |
}
|
wenzelm@49356
|
190 |
}
|
wenzelm@49356
|
191 |
}
|
wenzelm@49356
|
192 |
|
wenzelm@49356
|
193 |
|
wenzelm@49356
|
194 |
/* markup selectors */
|
wenzelm@49356
|
195 |
|
wenzelm@49358
|
196 |
private val highlight_include =
|
wenzelm@50201
|
197 |
Set(Markup.SORT, Markup.TYP, Markup.TERM, Markup.PROP, Markup.ML_TYPING, Markup.TOKEN_RANGE,
|
wenzelm@50201
|
198 |
Markup.ENTITY, Markup.PATH, Markup.SORTING, Markup.TYPING, Markup.FREE, Markup.SKOLEM,
|
wenzelm@50201
|
199 |
Markup.BOUND, Markup.VAR, Markup.TFREE, Markup.TVAR, Markup.ML_SOURCE, Markup.DOCUMENT_SOURCE)
|
wenzelm@49356
|
200 |
|
wenzelm@49358
|
201 |
def highlight(range: Text.Range): Option[Text.Info[Color]] =
|
wenzelm@49356
|
202 |
{
|
wenzelm@50500
|
203 |
snapshot.select_markup(range, Some(highlight_include), _ =>
|
wenzelm@49356
|
204 |
{
|
wenzelm@49358
|
205 |
case Text.Info(info_range, XML.Elem(Markup(name, _), _)) if highlight_include(name) =>
|
wenzelm@49358
|
206 |
Text.Info(snapshot.convert(info_range), highlight_color)
|
wenzelm@49356
|
207 |
}) match { case Text.Info(_, info) #:: _ => Some(info) case _ => None }
|
wenzelm@49356
|
208 |
}
|
wenzelm@49356
|
209 |
|
wenzelm@49356
|
210 |
|
wenzelm@50201
|
211 |
private val hyperlink_include = Set(Markup.ENTITY, Markup.PATH)
|
wenzelm@49356
|
212 |
|
wenzelm@49356
|
213 |
def hyperlink(range: Text.Range): Option[Text.Info[Hyperlink]] =
|
wenzelm@49356
|
214 |
{
|
wenzelm@50500
|
215 |
snapshot.cumulate_markup[List[Text.Info[Hyperlink]]](range, Nil, Some(hyperlink_include), _ =>
|
wenzelm@49356
|
216 |
{
|
wenzelm@50201
|
217 |
case (links, Text.Info(info_range, XML.Elem(Markup.Path(name), _)))
|
wenzelm@49356
|
218 |
if Path.is_ok(name) =>
|
wenzelm@50205
|
219 |
val jedit_file = PIDE.thy_load.append(snapshot.node_name.dir, Path.explode(name))
|
wenzelm@49356
|
220 |
Text.Info(snapshot.convert(info_range), Hyperlink(jedit_file, 0, 0)) :: links
|
wenzelm@49356
|
221 |
|
wenzelm@50201
|
222 |
case (links, Text.Info(info_range, XML.Elem(Markup(Markup.ENTITY, props), _)))
|
wenzelm@49815
|
223 |
if !props.exists(
|
wenzelm@50201
|
224 |
{ case (Markup.KIND, Markup.ML_OPEN) => true
|
wenzelm@50201
|
225 |
case (Markup.KIND, Markup.ML_STRUCT) => true
|
wenzelm@49815
|
226 |
case _ => false }) =>
|
wenzelm@49356
|
227 |
|
wenzelm@49356
|
228 |
props match {
|
wenzelm@49419
|
229 |
case Position.Def_Line_File(line, name) if Path.is_ok(name) =>
|
wenzelm@49356
|
230 |
Isabelle_System.source_file(Path.explode(name)) match {
|
wenzelm@49356
|
231 |
case Some(path) =>
|
wenzelm@49356
|
232 |
val jedit_file = Isabelle_System.platform_path(path)
|
wenzelm@49356
|
233 |
Text.Info(snapshot.convert(info_range), Hyperlink(jedit_file, line, 0)) :: links
|
wenzelm@49356
|
234 |
case None => links
|
wenzelm@49356
|
235 |
}
|
wenzelm@49356
|
236 |
|
wenzelm@49419
|
237 |
case Position.Def_Id_Offset(id, offset) =>
|
wenzelm@49356
|
238 |
snapshot.state.find_command(snapshot.version, id) match {
|
wenzelm@49356
|
239 |
case Some((node, command)) =>
|
wenzelm@49356
|
240 |
val sources =
|
wenzelm@49356
|
241 |
node.commands.iterator.takeWhile(_ != command).map(_.source) ++
|
wenzelm@49356
|
242 |
Iterator.single(command.source(Text.Range(0, command.decode(offset))))
|
wenzelm@49356
|
243 |
val (line, column) = ((1, 1) /: sources)(Symbol.advance_line_column)
|
wenzelm@49356
|
244 |
Text.Info(snapshot.convert(info_range),
|
wenzelm@49356
|
245 |
Hyperlink(command.node_name.node, line, column)) :: links
|
wenzelm@49356
|
246 |
case None => links
|
wenzelm@49356
|
247 |
}
|
wenzelm@49356
|
248 |
|
wenzelm@49356
|
249 |
case _ => links
|
wenzelm@49356
|
250 |
}
|
wenzelm@49356
|
251 |
}) match { case Text.Info(_, info :: _) #:: _ => Some(info) case _ => None }
|
wenzelm@49356
|
252 |
}
|
wenzelm@49356
|
253 |
|
wenzelm@49356
|
254 |
|
wenzelm@50500
|
255 |
private val active_include = Set(Markup.GRAPHVIEW, Markup.SENDBACK, Markup.DIALOG)
|
wenzelm@50450
|
256 |
|
wenzelm@50450
|
257 |
def active(range: Text.Range): Option[Text.Info[XML.Elem]] =
|
wenzelm@50500
|
258 |
snapshot.select_markup(range, Some(active_include), command_state =>
|
wenzelm@50501
|
259 |
{
|
wenzelm@50501
|
260 |
case Text.Info(info_range, elem @ Protocol.Dialog(_, serial, _))
|
wenzelm@50507
|
261 |
if !command_state.results.defined(serial) =>
|
wenzelm@50501
|
262 |
Text.Info(snapshot.convert(info_range), elem)
|
wenzelm@50501
|
263 |
case Text.Info(info_range, elem @ XML.Elem(Markup(name, _), _))
|
wenzelm@50501
|
264 |
if name == Markup.GRAPHVIEW || name == Markup.SENDBACK =>
|
wenzelm@50501
|
265 |
Text.Info(snapshot.convert(info_range), elem)
|
wenzelm@50163
|
266 |
}) match { case Text.Info(_, info) #:: _ => Some(info) case _ => None }
|
wenzelm@49492
|
267 |
|
wenzelm@49492
|
268 |
|
wenzelm@50502
|
269 |
def command_results(range: Text.Range): Command.Results =
|
wenzelm@50502
|
270 |
{
|
wenzelm@50502
|
271 |
val results =
|
wenzelm@50502
|
272 |
snapshot.select_markup[Command.Results](range, None, command_state =>
|
wenzelm@50502
|
273 |
{ case _ => command_state.results }).map(_.info)
|
wenzelm@50507
|
274 |
(Command.Results.empty /: results)(_ ++ _)
|
wenzelm@50502
|
275 |
}
|
wenzelm@50502
|
276 |
|
wenzelm@49700
|
277 |
def tooltip_message(range: Text.Range): XML.Body =
|
wenzelm@49356
|
278 |
{
|
wenzelm@49356
|
279 |
val msgs =
|
wenzelm@50507
|
280 |
Command.Results.merge(
|
wenzelm@50507
|
281 |
snapshot.cumulate_markup[Command.Results](range, Command.Results.empty,
|
wenzelm@50507
|
282 |
Some(Set(Markup.WRITELN, Markup.WARNING, Markup.ERROR, Markup.BAD)), _ =>
|
wenzelm@50507
|
283 |
{
|
wenzelm@50507
|
284 |
case (msgs, Text.Info(_, XML.Elem(Markup(name, props @ Markup.Serial(serial)), body)))
|
wenzelm@50507
|
285 |
if name == Markup.WRITELN ||
|
wenzelm@50507
|
286 |
name == Markup.WARNING ||
|
wenzelm@50507
|
287 |
name == Markup.ERROR =>
|
wenzelm@50507
|
288 |
msgs + (serial -> XML.Elem(Markup(Markup.message(name), props), body))
|
wenzelm@50507
|
289 |
case (msgs, Text.Info(_, msg @ XML.Elem(Markup(Markup.BAD, _), body)))
|
wenzelm@50507
|
290 |
if !body.isEmpty => msgs + (Document.new_id() -> msg)
|
wenzelm@50507
|
291 |
}).map(_.info))
|
wenzelm@50507
|
292 |
Pretty.separate(msgs.entries.map(_._2).toList)
|
wenzelm@49356
|
293 |
}
|
wenzelm@49356
|
294 |
|
wenzelm@49356
|
295 |
|
wenzelm@49356
|
296 |
private val tooltips: Map[String, String] =
|
wenzelm@49356
|
297 |
Map(
|
wenzelm@50201
|
298 |
Markup.SORT -> "sort",
|
wenzelm@50201
|
299 |
Markup.TYP -> "type",
|
wenzelm@50201
|
300 |
Markup.TERM -> "term",
|
wenzelm@50201
|
301 |
Markup.PROP -> "proposition",
|
wenzelm@50201
|
302 |
Markup.TOKEN_RANGE -> "inner syntax token",
|
wenzelm@50201
|
303 |
Markup.FREE -> "free variable",
|
wenzelm@50201
|
304 |
Markup.SKOLEM -> "skolem variable",
|
wenzelm@50201
|
305 |
Markup.BOUND -> "bound variable",
|
wenzelm@50201
|
306 |
Markup.VAR -> "schematic variable",
|
wenzelm@50201
|
307 |
Markup.TFREE -> "free type variable",
|
wenzelm@50201
|
308 |
Markup.TVAR -> "schematic type variable",
|
wenzelm@50201
|
309 |
Markup.ML_SOURCE -> "ML source",
|
wenzelm@50201
|
310 |
Markup.DOCUMENT_SOURCE -> "document source")
|
wenzelm@49356
|
311 |
|
wenzelm@49356
|
312 |
private val tooltip_elements =
|
wenzelm@50201
|
313 |
Set(Markup.ENTITY, Markup.SORTING, Markup.TYPING, Markup.ML_TYPING, Markup.PATH) ++
|
wenzelm@50201
|
314 |
tooltips.keys
|
wenzelm@49356
|
315 |
|
wenzelm@49700
|
316 |
private def pretty_typing(kind: String, body: XML.Body): XML.Tree =
|
wenzelm@49700
|
317 |
Pretty.block(XML.Text(kind) :: Pretty.Break(1) :: body)
|
wenzelm@49356
|
318 |
|
wenzelm@49700
|
319 |
def tooltip(range: Text.Range): XML.Body =
|
wenzelm@49356
|
320 |
{
|
wenzelm@49700
|
321 |
def add(prev: Text.Info[List[(Boolean, XML.Tree)]], r: Text.Range, p: (Boolean, XML.Tree)) =
|
wenzelm@49356
|
322 |
if (prev.range == r) Text.Info(r, p :: prev.info) else Text.Info(r, List(p))
|
wenzelm@49356
|
323 |
|
wenzelm@49356
|
324 |
val tips =
|
wenzelm@49700
|
325 |
snapshot.cumulate_markup[Text.Info[(List[(Boolean, XML.Tree)])]](
|
wenzelm@50500
|
326 |
range, Text.Info(range, Nil), Some(tooltip_elements), _ =>
|
wenzelm@49356
|
327 |
{
|
wenzelm@50201
|
328 |
case (prev, Text.Info(r, XML.Elem(Markup.Entity(kind, name), _))) =>
|
wenzelm@49356
|
329 |
val kind1 = space_explode('_', kind).mkString(" ")
|
wenzelm@49700
|
330 |
add(prev, r, (true, XML.Text(kind1 + " " + quote(name))))
|
wenzelm@50201
|
331 |
case (prev, Text.Info(r, XML.Elem(Markup.Path(name), _)))
|
wenzelm@49356
|
332 |
if Path.is_ok(name) =>
|
wenzelm@50205
|
333 |
val jedit_file = PIDE.thy_load.append(snapshot.node_name.dir, Path.explode(name))
|
wenzelm@49700
|
334 |
add(prev, r, (true, XML.Text("file " + quote(jedit_file))))
|
wenzelm@49674
|
335 |
case (prev, Text.Info(r, XML.Elem(Markup(name, _), body)))
|
wenzelm@50201
|
336 |
if name == Markup.SORTING || name == Markup.TYPING =>
|
wenzelm@49700
|
337 |
add(prev, r, (true, pretty_typing("::", body)))
|
wenzelm@50201
|
338 |
case (prev, Text.Info(r, XML.Elem(Markup(Markup.ML_TYPING, _), body))) =>
|
wenzelm@49700
|
339 |
add(prev, r, (false, pretty_typing("ML:", body)))
|
wenzelm@49356
|
340 |
case (prev, Text.Info(r, XML.Elem(Markup(name, _), _)))
|
wenzelm@49700
|
341 |
if tooltips.isDefinedAt(name) => add(prev, r, (true, XML.Text(tooltips(name))))
|
wenzelm@49356
|
342 |
}).toList.flatMap(_.info.info)
|
wenzelm@49356
|
343 |
|
wenzelm@49356
|
344 |
val all_tips =
|
wenzelm@49356
|
345 |
(tips.filter(_._1) ++ tips.filter(!_._1).lastOption.toList).map(_._2)
|
wenzelm@49700
|
346 |
Library.separate(Pretty.FBreak, all_tips.toList)
|
wenzelm@49356
|
347 |
}
|
wenzelm@49356
|
348 |
|
wenzelm@50554
|
349 |
val tooltip_margin: Int = options.int("jedit_tooltip_margin")
|
wenzelm@50554
|
350 |
val tooltip_bounds: Double = (options.real("jedit_tooltip_bounds") max 0.2) min 0.8
|
wenzelm@50554
|
351 |
|
wenzelm@49356
|
352 |
|
wenzelm@49356
|
353 |
def gutter_message(range: Text.Range): Option[Icon] =
|
wenzelm@49356
|
354 |
{
|
wenzelm@49356
|
355 |
val results =
|
wenzelm@50500
|
356 |
snapshot.cumulate_markup[Int](range, 0, Some(Set(Markup.WARNING, Markup.ERROR)), _ =>
|
wenzelm@49356
|
357 |
{
|
wenzelm@50201
|
358 |
case (pri, Text.Info(_, XML.Elem(Markup(Markup.WARNING, _), body))) =>
|
wenzelm@49356
|
359 |
body match {
|
wenzelm@50201
|
360 |
case List(XML.Elem(Markup(Markup.LEGACY, _), _)) =>
|
wenzelm@50199
|
361 |
pri max Rendering.legacy_pri
|
wenzelm@50199
|
362 |
case _ => pri max Rendering.warning_pri
|
wenzelm@49356
|
363 |
}
|
wenzelm@50201
|
364 |
case (pri, Text.Info(_, XML.Elem(Markup(Markup.ERROR, _), _))) =>
|
wenzelm@50199
|
365 |
pri max Rendering.error_pri
|
wenzelm@49356
|
366 |
})
|
wenzelm@49356
|
367 |
val pri = (0 /: results) { case (p1, Text.Info(_, p2)) => p1 max p2 }
|
wenzelm@50199
|
368 |
Rendering.gutter_icons.get(pri)
|
wenzelm@49356
|
369 |
}
|
wenzelm@49356
|
370 |
|
wenzelm@49356
|
371 |
|
wenzelm@49356
|
372 |
private val squiggly_colors = Map(
|
wenzelm@50199
|
373 |
Rendering.writeln_pri -> writeln_color,
|
wenzelm@50199
|
374 |
Rendering.warning_pri -> warning_color,
|
wenzelm@50199
|
375 |
Rendering.error_pri -> error_color)
|
wenzelm@49356
|
376 |
|
wenzelm@49356
|
377 |
def squiggly_underline(range: Text.Range): Stream[Text.Info[Color]] =
|
wenzelm@49356
|
378 |
{
|
wenzelm@49356
|
379 |
val results =
|
wenzelm@49356
|
380 |
snapshot.cumulate_markup[Int](range, 0,
|
wenzelm@50500
|
381 |
Some(Set(Markup.WRITELN, Markup.WARNING, Markup.ERROR)), _ =>
|
wenzelm@49356
|
382 |
{
|
wenzelm@49474
|
383 |
case (pri, Text.Info(_, XML.Elem(Markup(name, _), _)))
|
wenzelm@50201
|
384 |
if name == Markup.WRITELN ||
|
wenzelm@50201
|
385 |
name == Markup.WARNING ||
|
wenzelm@50201
|
386 |
name == Markup.ERROR => pri max Rendering.message_pri(name)
|
wenzelm@49356
|
387 |
})
|
wenzelm@49356
|
388 |
for {
|
wenzelm@49356
|
389 |
Text.Info(r, pri) <- results
|
wenzelm@49356
|
390 |
color <- squiggly_colors.get(pri)
|
wenzelm@49356
|
391 |
} yield Text.Info(r, color)
|
wenzelm@49356
|
392 |
}
|
wenzelm@49356
|
393 |
|
wenzelm@49474
|
394 |
|
wenzelm@49493
|
395 |
private val messages_include =
|
wenzelm@50201
|
396 |
Set(Markup.WRITELN_MESSAGE, Markup.TRACING_MESSAGE, Markup.WARNING_MESSAGE, Markup.ERROR_MESSAGE)
|
wenzelm@49493
|
397 |
|
wenzelm@49474
|
398 |
private val message_colors = Map(
|
wenzelm@50199
|
399 |
Rendering.writeln_pri -> writeln_message_color,
|
wenzelm@50199
|
400 |
Rendering.tracing_pri -> tracing_message_color,
|
wenzelm@50199
|
401 |
Rendering.warning_pri -> warning_message_color,
|
wenzelm@50199
|
402 |
Rendering.error_pri -> error_message_color)
|
wenzelm@49474
|
403 |
|
wenzelm@49473
|
404 |
def line_background(range: Text.Range): Option[(Color, Boolean)] =
|
wenzelm@49473
|
405 |
{
|
wenzelm@49474
|
406 |
val results =
|
wenzelm@50500
|
407 |
snapshot.cumulate_markup[Int](range, 0, Some(messages_include), _ =>
|
wenzelm@49473
|
408 |
{
|
wenzelm@49474
|
409 |
case (pri, Text.Info(_, XML.Elem(Markup(name, _), _)))
|
wenzelm@50201
|
410 |
if name == Markup.WRITELN_MESSAGE ||
|
wenzelm@50201
|
411 |
name == Markup.TRACING_MESSAGE ||
|
wenzelm@50201
|
412 |
name == Markup.WARNING_MESSAGE ||
|
wenzelm@50201
|
413 |
name == Markup.ERROR_MESSAGE => pri max Rendering.message_pri(name)
|
wenzelm@49474
|
414 |
})
|
wenzelm@49474
|
415 |
val pri = (0 /: results) { case (p1, Text.Info(_, p2)) => p1 max p2 }
|
wenzelm@49474
|
416 |
|
wenzelm@49474
|
417 |
val is_separator = pri > 0 &&
|
wenzelm@50500
|
418 |
snapshot.cumulate_markup[Boolean](range, false, Some(Set(Markup.SEPARATOR)), _ =>
|
wenzelm@49473
|
419 |
{
|
wenzelm@50201
|
420 |
case (_, Text.Info(_, XML.Elem(Markup(Markup.SEPARATOR, _), _))) => true
|
wenzelm@49473
|
421 |
}).exists(_.info)
|
wenzelm@49473
|
422 |
|
wenzelm@49474
|
423 |
message_colors.get(pri).map((_, is_separator))
|
wenzelm@49473
|
424 |
}
|
wenzelm@49356
|
425 |
|
wenzelm@49492
|
426 |
|
wenzelm@50500
|
427 |
def output_messages(st: Command.State): List[XML.Tree] =
|
wenzelm@50500
|
428 |
{
|
wenzelm@50507
|
429 |
st.results.entries.map(_._2).filterNot(Protocol.is_result(_)).toList
|
wenzelm@50500
|
430 |
}
|
wenzelm@50500
|
431 |
|
wenzelm@50500
|
432 |
|
wenzelm@49492
|
433 |
private val background1_include =
|
wenzelm@50201
|
434 |
Protocol.command_status_markup + Markup.WRITELN_MESSAGE + Markup.TRACING_MESSAGE +
|
wenzelm@50500
|
435 |
Markup.WARNING_MESSAGE + Markup.ERROR_MESSAGE + Markup.BAD + Markup.INTENSIFY ++
|
wenzelm@50500
|
436 |
active_include
|
wenzelm@49492
|
437 |
|
wenzelm@49356
|
438 |
def background1(range: Text.Range): Stream[Text.Info[Color]] =
|
wenzelm@49356
|
439 |
{
|
wenzelm@49356
|
440 |
if (snapshot.is_outdated) Stream(Text.Info(range, outdated_color))
|
wenzelm@49356
|
441 |
else
|
wenzelm@49356
|
442 |
for {
|
wenzelm@49356
|
443 |
Text.Info(r, result) <-
|
wenzelm@50500
|
444 |
snapshot.cumulate_markup[(Option[Protocol.Status], Option[Color])](
|
wenzelm@50500
|
445 |
range, (Some(Protocol.Status.init), None), Some(background1_include), command_state =>
|
wenzelm@49356
|
446 |
{
|
wenzelm@50500
|
447 |
case (((Some(status), color), Text.Info(_, XML.Elem(markup, _))))
|
wenzelm@49356
|
448 |
if (Protocol.command_status_markup(markup.name)) =>
|
wenzelm@50500
|
449 |
(Some(Protocol.command_status(status, markup)), color)
|
wenzelm@50500
|
450 |
case (_, Text.Info(_, XML.Elem(Markup(Markup.BAD, _), _))) =>
|
wenzelm@50500
|
451 |
(None, Some(bad_color))
|
wenzelm@50500
|
452 |
case (_, Text.Info(_, XML.Elem(Markup(Markup.INTENSIFY, _), _))) =>
|
wenzelm@50500
|
453 |
(None, Some(intensify_color))
|
wenzelm@50501
|
454 |
case (acc, Text.Info(_, elem @ XML.Elem(Markup(Markup.DIALOG, _), _))) =>
|
wenzelm@50501
|
455 |
// FIXME pattern match problem in scala-2.9.2 (!??)
|
wenzelm@50501
|
456 |
elem match {
|
wenzelm@50501
|
457 |
case Protocol.Dialog(_, serial, result) =>
|
wenzelm@50501
|
458 |
command_state.results.get(serial) match {
|
wenzelm@50501
|
459 |
case Some(Protocol.Dialog_Result(res)) if res == result =>
|
wenzelm@50501
|
460 |
(None, Some(active_result_color))
|
wenzelm@50501
|
461 |
case _ =>
|
wenzelm@50501
|
462 |
(None, Some(active_color))
|
wenzelm@50501
|
463 |
}
|
wenzelm@50501
|
464 |
case _ => acc
|
wenzelm@50500
|
465 |
}
|
wenzelm@50500
|
466 |
case (_, Text.Info(_, XML.Elem(markup, _))) if active_include(markup.name) =>
|
wenzelm@50500
|
467 |
(None, Some(active_color))
|
wenzelm@49356
|
468 |
})
|
wenzelm@49356
|
469 |
color <-
|
wenzelm@49356
|
470 |
(result match {
|
wenzelm@50500
|
471 |
case (Some(status), opt_color) =>
|
wenzelm@49356
|
472 |
if (status.is_unprocessed) Some(unprocessed1_color)
|
wenzelm@49356
|
473 |
else if (status.is_running) Some(running1_color)
|
wenzelm@49356
|
474 |
else opt_color
|
wenzelm@50500
|
475 |
case (_, opt_color) => opt_color
|
wenzelm@49356
|
476 |
})
|
wenzelm@49356
|
477 |
} yield Text.Info(r, color)
|
wenzelm@49356
|
478 |
}
|
wenzelm@49356
|
479 |
|
wenzelm@49356
|
480 |
|
wenzelm@49356
|
481 |
def background2(range: Text.Range): Stream[Text.Info[Color]] =
|
wenzelm@50500
|
482 |
snapshot.select_markup(range, Some(Set(Markup.TOKEN_RANGE)), _ =>
|
wenzelm@49356
|
483 |
{
|
wenzelm@50201
|
484 |
case Text.Info(_, XML.Elem(Markup(Markup.TOKEN_RANGE, _), _)) => light_color
|
wenzelm@49356
|
485 |
})
|
wenzelm@49356
|
486 |
|
wenzelm@49356
|
487 |
|
wenzelm@49356
|
488 |
def foreground(range: Text.Range): Stream[Text.Info[Color]] =
|
wenzelm@50500
|
489 |
snapshot.select_markup(range, Some(Set(Markup.STRING, Markup.ALTSTRING, Markup.VERBATIM)), _ =>
|
wenzelm@49356
|
490 |
{
|
wenzelm@50201
|
491 |
case Text.Info(_, XML.Elem(Markup(Markup.STRING, _), _)) => quoted_color
|
wenzelm@50201
|
492 |
case Text.Info(_, XML.Elem(Markup(Markup.ALTSTRING, _), _)) => quoted_color
|
wenzelm@50201
|
493 |
case Text.Info(_, XML.Elem(Markup(Markup.VERBATIM, _), _)) => quoted_color
|
wenzelm@49356
|
494 |
})
|
wenzelm@49356
|
495 |
|
wenzelm@49356
|
496 |
|
wenzelm@49356
|
497 |
private val text_colors: Map[String, Color] = Map(
|
wenzelm@50201
|
498 |
Markup.KEYWORD1 -> keyword1_color,
|
wenzelm@50201
|
499 |
Markup.KEYWORD2 -> keyword2_color,
|
wenzelm@50201
|
500 |
Markup.STRING -> Color.BLACK,
|
wenzelm@50201
|
501 |
Markup.ALTSTRING -> Color.BLACK,
|
wenzelm@50201
|
502 |
Markup.VERBATIM -> Color.BLACK,
|
wenzelm@50201
|
503 |
Markup.LITERAL -> keyword1_color,
|
wenzelm@50201
|
504 |
Markup.DELIMITER -> Color.BLACK,
|
wenzelm@50201
|
505 |
Markup.TFREE -> tfree_color,
|
wenzelm@50201
|
506 |
Markup.TVAR -> tvar_color,
|
wenzelm@50201
|
507 |
Markup.FREE -> free_color,
|
wenzelm@50201
|
508 |
Markup.SKOLEM -> skolem_color,
|
wenzelm@50201
|
509 |
Markup.BOUND -> bound_color,
|
wenzelm@50201
|
510 |
Markup.VAR -> var_color,
|
wenzelm@50201
|
511 |
Markup.INNER_STRING -> inner_quoted_color,
|
wenzelm@50201
|
512 |
Markup.INNER_COMMENT -> inner_comment_color,
|
wenzelm@50201
|
513 |
Markup.DYNAMIC_FACT -> dynamic_color,
|
wenzelm@50201
|
514 |
Markup.ML_KEYWORD -> keyword1_color,
|
wenzelm@50201
|
515 |
Markup.ML_DELIMITER -> Color.BLACK,
|
wenzelm@50201
|
516 |
Markup.ML_NUMERAL -> inner_numeral_color,
|
wenzelm@50201
|
517 |
Markup.ML_CHAR -> inner_quoted_color,
|
wenzelm@50201
|
518 |
Markup.ML_STRING -> inner_quoted_color,
|
wenzelm@50201
|
519 |
Markup.ML_COMMENT -> inner_comment_color,
|
wenzelm@50201
|
520 |
Markup.ANTIQ -> antiquotation_color)
|
wenzelm@49356
|
521 |
|
wenzelm@49356
|
522 |
private val text_color_elements = Set.empty[String] ++ text_colors.keys
|
wenzelm@49356
|
523 |
|
wenzelm@49356
|
524 |
def text_color(range: Text.Range, color: Color)
|
wenzelm@49356
|
525 |
: Stream[Text.Info[Color]] =
|
wenzelm@49356
|
526 |
{
|
wenzelm@50196
|
527 |
if (color == Token_Markup.hidden_color) Stream(Text.Info(range, color))
|
wenzelm@50196
|
528 |
else
|
wenzelm@50500
|
529 |
snapshot.cumulate_markup(range, color, Some(text_color_elements), _ =>
|
wenzelm@50196
|
530 |
{
|
wenzelm@50196
|
531 |
case (_, Text.Info(_, XML.Elem(Markup(m, _), _)))
|
wenzelm@50196
|
532 |
if text_colors.isDefinedAt(m) => text_colors(m)
|
wenzelm@50196
|
533 |
})
|
wenzelm@49356
|
534 |
}
|
wenzelm@50542
|
535 |
|
wenzelm@50542
|
536 |
|
wenzelm@50542
|
537 |
/* nested text structure -- folds */
|
wenzelm@50542
|
538 |
|
wenzelm@50545
|
539 |
private val fold_depth_include = Set(Markup.TEXT_FOLD, Markup.GOAL, Markup.SUBGOAL)
|
wenzelm@50542
|
540 |
|
wenzelm@50542
|
541 |
def fold_depth(range: Text.Range): Stream[Text.Info[Int]] =
|
wenzelm@50542
|
542 |
snapshot.cumulate_markup[Int](range, 0, Some(fold_depth_include), _ =>
|
wenzelm@50542
|
543 |
{
|
wenzelm@50542
|
544 |
case (depth, Text.Info(_, XML.Elem(Markup(name, _), _)))
|
wenzelm@50542
|
545 |
if fold_depth_include(name) => depth + 1
|
wenzelm@50542
|
546 |
})
|
wenzelm@49356
|
547 |
}
|