| author | wenzelm | 
| Tue, 24 Jun 2025 21:05:48 +0200 | |
| changeset 82747 | 00828818a607 | 
| parent 82142 | 508a673c87ac | 
| permissions | -rw-r--r-- | 
| 63422 | 1 | /* Title: Tools/jEdit/src/text_structure.scala | 
| 58748 | 2 | Author: Makarius | 
| 3 | ||
| 63422 | 4 | Text structure based on Isabelle/Isar outer syntax. | 
| 58748 | 5 | */ | 
| 6 | ||
| 7 | package isabelle.jedit | |
| 8 | ||
| 9 | ||
| 10 | import isabelle._ | |
| 11 | ||
| 73909 | 12 | import java.util.{List => JList}
 | 
| 13 | ||
| 63422 | 14 | import org.gjt.sp.jedit.indent.{IndentRule, IndentAction}
 | 
| 58803 
7a0f675eb671
proper selectMatch, e.g. relevant for S-click on gutter;
 wenzelm parents: 
58800diff
changeset | 15 | import org.gjt.sp.jedit.textarea.{TextArea, StructureMatcher, Selection}
 | 
| 63422 | 16 | import org.gjt.sp.jedit.buffer.JEditBuffer | 
| 58748 | 17 | |
| 18 | ||
| 75393 | 19 | object Text_Structure {
 | 
| 63425 | 20 | /* token navigator */ | 
| 21 | ||
| 75393 | 22 |   class Navigator(syntax: Outer_Syntax, buffer: JEditBuffer, comments: Boolean) {
 | 
| 71601 | 23 |     val limit: Int = PIDE.options.value.int("jedit_structure_limit") max 0
 | 
| 63425 | 24 | |
| 75393 | 25 |     def iterator(line: Int, lim: Int = limit): Iterator[Text.Info[Token]] = {
 | 
| 63445 
5761bb8592dc
observe comments in indentation, but not in fold structure;
 wenzelm parents: 
63442diff
changeset | 26 | val it = Token_Markup.line_token_iterator(syntax, buffer, line, line + lim) | 
| 68730 | 27 | if (comments) it.filterNot(_.info.is_space) else it.filter(_.info.is_proper) | 
| 63445 
5761bb8592dc
observe comments in indentation, but not in fold structure;
 wenzelm parents: 
63442diff
changeset | 28 | } | 
| 63425 | 29 | |
| 75393 | 30 |     def reverse_iterator(line: Int, lim: Int = limit): Iterator[Text.Info[Token]] = {
 | 
| 63445 
5761bb8592dc
observe comments in indentation, but not in fold structure;
 wenzelm parents: 
63442diff
changeset | 31 | val it = Token_Markup.line_token_reverse_iterator(syntax, buffer, line, line - lim) | 
| 68730 | 32 | if (comments) it.filterNot(_.info.is_space) else it.filter(_.info.is_proper) | 
| 63445 
5761bb8592dc
observe comments in indentation, but not in fold structure;
 wenzelm parents: 
63442diff
changeset | 33 | } | 
| 63425 | 34 | } | 
| 35 | ||
| 36 | ||
| 63422 | 37 | /* indentation */ | 
| 38 | ||
| 75393 | 39 |   object Indent_Rule extends IndentRule {
 | 
| 63428 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 40 | private val keyword_open = Keyword.theory_goal ++ Keyword.proof_open | 
| 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 41 | private val keyword_close = Keyword.proof_close | 
| 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 42 | |
| 75393 | 43 | def apply( | 
| 44 | buffer: JEditBuffer, | |
| 45 | current_line: Int, | |
| 46 | prev_line0: Int, | |
| 47 | prev_prev_line0: Int, | |
| 48 | actions: JList[IndentAction] | |
| 49 |     ): Unit = {
 | |
| 63425 | 50 |       Isabelle.buffer_syntax(buffer) match {
 | 
| 66183 | 51 | case Some(syntax) => | 
| 63425 | 52 | val keywords = syntax.keywords | 
| 66183 | 53 | val nav = new Navigator(syntax, buffer, true) | 
| 63422 | 54 | |
| 63474 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 55 | val indent_size = buffer.getIndentSize | 
| 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 56 | |
| 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 57 | |
| 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 58 | def line_indent(line: Int): Int = | 
| 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 59 | if (line < 0 || line >= buffer.getLineCount) 0 | 
| 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 60 | else buffer.getCurrentIndentForLine(line, null) | 
| 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 61 | |
| 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 62 | def line_head(line: Int): Option[Text.Info[Token]] = | 
| 73345 | 63 | nav.iterator(line, 1).nextOption() | 
| 63428 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 64 | |
| 63434 | 65 | def head_is_quasi_command(line: Int): Boolean = | 
| 63474 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 66 |             line_head(line) match {
 | 
| 63434 | 67 | case None => false | 
| 63474 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 68 | case Some(Text.Info(_, tok)) => keywords.is_quasi_command(tok) | 
| 63434 | 69 | } | 
| 70 | ||
| 64518 | 71 | val prev_line: Int = | 
| 72 | Range.inclusive(current_line - 1, 0, -1).find(line => | |
| 66176 | 73 | Token_Markup.Line_Context.before(buffer, line).get_context == Scan.Finished && | 
| 66178 | 74 | (!Token_Markup.Line_Context.after(buffer, line).structure.improper || | 
| 75 | Token_Markup.Line_Context.after(buffer, line).structure.blank)) getOrElse -1 | |
| 64518 | 76 | |
| 63477 
f5c81436b930
clarified indentation: 'begin' is treated like a separate command without indent;
 wenzelm parents: 
63474diff
changeset | 77 | def prev_line_command: Option[Token] = | 
| 63428 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 78 | nav.reverse_iterator(prev_line, 1). | 
| 63477 
f5c81436b930
clarified indentation: 'begin' is treated like a separate command without indent;
 wenzelm parents: 
63474diff
changeset | 79 |               collectFirst({ case Text.Info(_, tok) if tok.is_begin_or_command => tok })
 | 
| 
f5c81436b930
clarified indentation: 'begin' is treated like a separate command without indent;
 wenzelm parents: 
63474diff
changeset | 80 | |
| 
f5c81436b930
clarified indentation: 'begin' is treated like a separate command without indent;
 wenzelm parents: 
63474diff
changeset | 81 | def prev_line_span: Iterator[Token] = | 
| 
f5c81436b930
clarified indentation: 'begin' is treated like a separate command without indent;
 wenzelm parents: 
63474diff
changeset | 82 | nav.reverse_iterator(prev_line, 1).map(_.info).takeWhile(tok => !tok.is_begin_or_command) | 
| 63428 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 83 | |
| 63434 | 84 | def prev_span: Iterator[Token] = | 
| 63477 
f5c81436b930
clarified indentation: 'begin' is treated like a separate command without indent;
 wenzelm parents: 
63474diff
changeset | 85 | nav.reverse_iterator(prev_line).map(_.info).takeWhile(tok => !tok.is_begin_or_command) | 
| 63450 | 86 | |
| 63434 | 87 | |
| 75393 | 88 |           val script_indent: Text.Info[Token] => Int = {
 | 
| 64621 | 89 | val opt_rendering: Option[JEdit_Rendering] = | 
| 63474 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 90 |               if (PIDE.options.value.bool("jedit_indent_script"))
 | 
| 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 91 |                 GUI_Thread.now {
 | 
| 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 92 |                   (for {
 | 
| 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 93 | text_area <- JEdit_Lib.jedit_text_areas(buffer) | 
| 76765 
c654103e9c9d
more robust Document.Pending_Edits: cover all nodes simulataneously, and thus support proper Snapshot.switch;
 wenzelm parents: 
75423diff
changeset | 94 | rendering <- Document_View.get_rendering(text_area) | 
| 
c654103e9c9d
more robust Document.Pending_Edits: cover all nodes simulataneously, and thus support proper Snapshot.switch;
 wenzelm parents: 
75423diff
changeset | 95 | } yield rendering).nextOption() | 
| 63474 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 96 | } | 
| 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 97 | else None | 
| 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 98 |             val limit = PIDE.options.value.int("jedit_indent_script_limit")
 | 
| 63481 | 99 | (info: Text.Info[Token]) => | 
| 63474 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 100 |               opt_rendering match {
 | 
| 63481 | 101 | case Some(rendering) if keywords.is_command(info.info, Keyword.prf_script) => | 
| 102 | (rendering.indentation(info.range) min limit) max 0 | |
| 103 | case _ => 0 | |
| 63474 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 104 | } | 
| 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 105 | } | 
| 63428 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 106 | |
| 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 107 | def indent_indent(tok: Token): Int = | 
| 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 108 | if (keywords.is_command(tok, keyword_open)) indent_size | 
| 75423 | 109 |             else if (keywords.is_command(tok, keyword_close)) { - indent_size }
 | 
| 63428 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 110 | else 0 | 
| 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 111 | |
| 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 112 | def indent_offset(tok: Token): Int = | 
| 63477 
f5c81436b930
clarified indentation: 'begin' is treated like a separate command without indent;
 wenzelm parents: 
63474diff
changeset | 113 | if (keywords.is_command(tok, Keyword.proof_enclose)) indent_size | 
| 63428 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 114 | else 0 | 
| 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 115 | |
| 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 116 | def indent_structure: Int = | 
| 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 117 | nav.reverse_iterator(current_line - 1).scanLeft((0, false))( | 
| 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 118 |               { case ((ind, _), Text.Info(range, tok)) =>
 | 
| 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 119 | val ind1 = ind + indent_indent(tok) | 
| 63479 | 120 |                   if (tok.is_begin_or_command && !keywords.is_command(tok, Keyword.prf_script)) {
 | 
| 63428 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 121 | val line = buffer.getLineOfOffset(range.start) | 
| 63474 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 122 |                     line_head(line) match {
 | 
| 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 123 | case Some(info) if info.info == tok => | 
| 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 124 | (ind1 + indent_offset(tok) + line_indent(line), true) | 
| 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 125 | case _ => (ind1, false) | 
| 
f66e3c3b0fb1
semantic indentation for unstructured proof scripts;
 wenzelm parents: 
63450diff
changeset | 126 | } | 
| 63428 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 127 | } | 
| 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 128 | else (ind1, false) | 
| 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 129 |               }).collectFirst({ case (i, true) => i }).getOrElse(0)
 | 
| 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 130 | |
| 63480 | 131 | def indent_brackets: Int = | 
| 73359 | 132 |             prev_line_span.foldLeft(0) {
 | 
| 133 | case (i, tok) => | |
| 134 | if (tok.is_open_bracket) i + indent_size | |
| 135 | else if (tok.is_close_bracket) i - indent_size | |
| 136 | else i | |
| 137 | } | |
| 63480 | 138 | |
| 139 | def indent_extra: Int = | |
| 71601 | 140 | if (prev_span.exists(keywords.is_quasi_command)) indent_size | 
| 63480 | 141 | else 0 | 
| 142 | ||
| 63428 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 143 | val indent = | 
| 66179 
148d61626014
indent = 0 for blank lines: produce less whitespace by default;
 wenzelm parents: 
66178diff
changeset | 144 | if (Token_Markup.Line_Context.before(buffer, current_line).get_context != Scan.Finished) | 
| 
148d61626014
indent = 0 for blank lines: produce less whitespace by default;
 wenzelm parents: 
66178diff
changeset | 145 | line_indent(current_line) | 
| 
148d61626014
indent = 0 for blank lines: produce less whitespace by default;
 wenzelm parents: 
66178diff
changeset | 146 | else if (Token_Markup.Line_Context.after(buffer, current_line).structure.blank) 0 | 
| 
148d61626014
indent = 0 for blank lines: produce less whitespace by default;
 wenzelm parents: 
66178diff
changeset | 147 |             else {
 | 
| 64518 | 148 |               line_head(current_line) match {
 | 
| 73105 
578a33042aa6
clarified: command keyword position is sufficient (amending 693a39f2cddc);
 wenzelm parents: 
71601diff
changeset | 149 | case Some(info) => | 
| 
578a33042aa6
clarified: command keyword position is sufficient (amending 693a39f2cddc);
 wenzelm parents: 
71601diff
changeset | 150 | val tok = info.info | 
| 64518 | 151 | if (tok.is_begin || | 
| 152 | keywords.is_before_command(tok) || | |
| 153 | keywords.is_command(tok, Keyword.theory)) 0 | |
| 154 | else if (keywords.is_command(tok, Keyword.proof_enclose)) | |
| 155 | indent_structure + script_indent(info) - indent_offset(tok) | |
| 156 | else if (keywords.is_command(tok, Keyword.proof)) | |
| 157 | (indent_structure + script_indent(info) - indent_offset(tok)) max indent_size | |
| 158 | else if (tok.is_command) indent_structure - indent_offset(tok) | |
| 159 |                   else {
 | |
| 160 |                     prev_line_command match {
 | |
| 161 | case None => | |
| 162 | val extra = | |
| 163 |                           (keywords.is_quasi_command(tok), head_is_quasi_command(prev_line)) match {
 | |
| 164 | case (true, true) | (false, false) => 0 | |
| 165 | case (true, false) => - indent_extra | |
| 166 | case (false, true) => indent_extra | |
| 167 | } | |
| 168 | line_indent(prev_line) + indent_brackets + extra - indent_offset(tok) | |
| 169 | case Some(prev_tok) => | |
| 170 | indent_structure + indent_brackets + indent_size - indent_offset(tok) - | |
| 171 | indent_offset(prev_tok) - indent_indent(prev_tok) | |
| 172 | } | |
| 64536 
e61de633a3ed
more uniform indentation of new line, even if it is empty (relevant for non-proof commands, e.g. 'definition', 'context');
 wenzelm parents: 
64518diff
changeset | 173 | } | 
| 
e61de633a3ed
more uniform indentation of new line, even if it is empty (relevant for non-proof commands, e.g. 'definition', 'context');
 wenzelm parents: 
64518diff
changeset | 174 | case None => | 
| 
e61de633a3ed
more uniform indentation of new line, even if it is empty (relevant for non-proof commands, e.g. 'definition', 'context');
 wenzelm parents: 
64518diff
changeset | 175 |                   prev_line_command match {
 | 
| 
e61de633a3ed
more uniform indentation of new line, even if it is empty (relevant for non-proof commands, e.g. 'definition', 'context');
 wenzelm parents: 
64518diff
changeset | 176 | case None => | 
| 
e61de633a3ed
more uniform indentation of new line, even if it is empty (relevant for non-proof commands, e.g. 'definition', 'context');
 wenzelm parents: 
64518diff
changeset | 177 | val extra = if (head_is_quasi_command(prev_line)) indent_extra else 0 | 
| 
e61de633a3ed
more uniform indentation of new line, even if it is empty (relevant for non-proof commands, e.g. 'definition', 'context');
 wenzelm parents: 
64518diff
changeset | 178 | line_indent(prev_line) + indent_brackets + extra | 
| 
e61de633a3ed
more uniform indentation of new line, even if it is empty (relevant for non-proof commands, e.g. 'definition', 'context');
 wenzelm parents: 
64518diff
changeset | 179 | case Some(prev_tok) => | 
| 
e61de633a3ed
more uniform indentation of new line, even if it is empty (relevant for non-proof commands, e.g. 'definition', 'context');
 wenzelm parents: 
64518diff
changeset | 180 | indent_structure + indent_brackets + indent_size - | 
| 
e61de633a3ed
more uniform indentation of new line, even if it is empty (relevant for non-proof commands, e.g. 'definition', 'context');
 wenzelm parents: 
64518diff
changeset | 181 | indent_offset(prev_tok) - indent_indent(prev_tok) | 
| 
e61de633a3ed
more uniform indentation of new line, even if it is empty (relevant for non-proof commands, e.g. 'definition', 'context');
 wenzelm parents: 
64518diff
changeset | 182 | } | 
| 64518 | 183 | } | 
| 63428 
005b490f0ce2
indentation in reminiscence to Proof General (see proof-indent.el);
 wenzelm parents: 
63427diff
changeset | 184 | } | 
| 63423 | 185 | |
| 63425 | 186 | actions.clear() | 
| 63439 | 187 | actions.add(new IndentAction.AlignOffset(indent max 0)) | 
| 66183 | 188 | case None => | 
| 63423 | 189 | } | 
| 63422 | 190 | } | 
| 191 | } | |
| 192 | ||
| 75393 | 193 | def line_content( | 
| 194 | buffer: JEditBuffer, | |
| 195 | keywords: Keyword.Keywords, | |
| 196 | range: Text.Range, | |
| 197 | ctxt: Scan.Line_Context | |
| 198 |   ): (List[Token], Scan.Line_Context) = {
 | |
| 67014 | 199 |     val text = JEdit_Lib.get_text(buffer, range).getOrElse("")
 | 
| 66175 | 200 | val (toks, ctxt1) = Token.explode_line(keywords, text, ctxt) | 
| 66173 | 201 | val toks1 = toks.filterNot(_.is_space) | 
| 66175 | 202 | (toks1, ctxt1) | 
| 66173 | 203 | } | 
| 204 | ||
| 75393 | 205 | def split_line_content( | 
| 206 | buffer: JEditBuffer, | |
| 207 | keywords: Keyword.Keywords, | |
| 208 | line: Int, | |
| 209 | caret: Int | |
| 210 |   ): (List[Token], List[Token]) = {
 | |
| 66173 | 211 | val line_range = JEdit_Lib.line_range(buffer, line) | 
| 66176 | 212 | val ctxt0 = Token_Markup.Line_Context.before(buffer, line).get_context | 
| 66175 | 213 | val (toks1, ctxt1) = line_content(buffer, keywords, Text.Range(line_range.start, caret), ctxt0) | 
| 214 | val (toks2, _) = line_content(buffer, keywords, Text.Range(caret, line_range.stop), ctxt1) | |
| 66173 | 215 | (toks1, toks2) | 
| 216 | } | |
| 217 | ||
| 63422 | 218 | |
| 219 | /* structure matching */ | |
| 220 | ||
| 75393 | 221 |   object Matcher extends StructureMatcher {
 | 
| 58803 
7a0f675eb671
proper selectMatch, e.g. relevant for S-click on gutter;
 wenzelm parents: 
58800diff
changeset | 222 | private def find_block( | 
| 58752 | 223 | open: Token => Boolean, | 
| 224 | close: Token => Boolean, | |
| 225 | reset: Token => Boolean, | |
| 58762 | 226 | restrict: Token => Boolean, | 
| 75393 | 227 | it: Iterator[Text.Info[Token]] | 
| 228 |     ): Option[(Text.Range, Text.Range)] = {
 | |
| 73344 | 229 | val range1 = it.next().range | 
| 58762 | 230 | it.takeWhile(info => !info.info.is_command || restrict(info.info)). | 
| 231 | scanLeft((range1, 1))( | |
| 232 |           { case ((r, d), Text.Info(range, tok)) =>
 | |
| 233 | if (open(tok)) (range, d + 1) | |
| 234 | else if (close(tok)) (range, d - 1) | |
| 235 | else if (reset(tok)) (range, 0) | |
| 236 | else (r, d) } | |
| 237 |         ).collectFirst({ case (range2, 0) => (range1, range2) })
 | |
| 58752 | 238 | } | 
| 239 | ||
| 75393 | 240 |     private def find_pair(text_area: TextArea): Option[(Text.Range, Text.Range)] = {
 | 
| 58748 | 241 | val buffer = text_area.getBuffer | 
| 242 | val caret_line = text_area.getCaretLine | |
| 58749 
83b0f633190e
some structure matching, based on line token iterators;
 wenzelm parents: 
58748diff
changeset | 243 | val caret = text_area.getCaretPosition | 
| 58748 | 244 | |
| 59074 | 245 |       Isabelle.buffer_syntax(text_area.getBuffer) match {
 | 
| 66183 | 246 | case Some(syntax) => | 
| 63424 | 247 | val keywords = syntax.keywords | 
| 58750 | 248 | |
| 66183 | 249 | val nav = new Navigator(syntax, buffer, false) | 
| 58750 | 250 | |
| 58752 | 251 | def caret_iterator(): Iterator[Text.Info[Token]] = | 
| 63425 | 252 | nav.iterator(caret_line).dropWhile(info => !info.range.touches(caret)) | 
| 58752 | 253 | |
| 63427 | 254 | def reverse_caret_iterator(): Iterator[Text.Info[Token]] = | 
| 255 | nav.reverse_iterator(caret_line).dropWhile(info => !info.range.touches(caret)) | |
| 58752 | 256 | |
| 63425 | 257 | nav.iterator(caret_line, 1).find(info => info.range.touches(caret)) | 
| 58756 
eb5d0c58564d
ignore improper tokens to avoid ambiguity of Range.touches (assuming that relevant tokens are separated properly);
 wenzelm parents: 
58755diff
changeset | 258 |           match {
 | 
| 63424 | 259 | case Some(Text.Info(range1, tok)) if keywords.is_command(tok, Keyword.theory_goal) => | 
| 58755 | 260 | find_block( | 
| 63424 | 261 | keywords.is_command(_, Keyword.proof_goal), | 
| 262 | keywords.is_command(_, Keyword.qed), | |
| 263 | keywords.is_command(_, Keyword.qed_global), | |
| 58762 | 264 | t => | 
| 63424 | 265 | keywords.is_command(t, Keyword.diag) || | 
| 266 | keywords.is_command(t, Keyword.proof), | |
| 58755 | 267 | caret_iterator()) | 
| 268 | ||
| 63424 | 269 | case Some(Text.Info(range1, tok)) if keywords.is_command(tok, Keyword.proof_goal) => | 
| 58755 | 270 | find_block( | 
| 63424 | 271 | keywords.is_command(_, Keyword.proof_goal), | 
| 272 | keywords.is_command(_, Keyword.qed), | |
| 58755 | 273 | _ => false, | 
| 58762 | 274 | t => | 
| 63424 | 275 | keywords.is_command(t, Keyword.diag) || | 
| 276 | keywords.is_command(t, Keyword.proof), | |
| 58755 | 277 | caret_iterator()) | 
| 278 | ||
| 63424 | 279 | case Some(Text.Info(range1, tok)) if keywords.is_command(tok, Keyword.qed_global) => | 
| 63427 | 280 | reverse_caret_iterator().find(info => keywords.is_command(info.info, Keyword.theory)) | 
| 58749 
83b0f633190e
some structure matching, based on line token iterators;
 wenzelm parents: 
58748diff
changeset | 281 |               match {
 | 
| 58750 | 282 | case Some(Text.Info(range2, tok)) | 
| 63424 | 283 | if keywords.is_command(tok, Keyword.theory_goal) => Some((range1, range2)) | 
| 58749 
83b0f633190e
some structure matching, based on line token iterators;
 wenzelm parents: 
58748diff
changeset | 284 | case _ => None | 
| 
83b0f633190e
some structure matching, based on line token iterators;
 wenzelm parents: 
58748diff
changeset | 285 | } | 
| 58752 | 286 | |
| 63424 | 287 | case Some(Text.Info(range1, tok)) if keywords.is_command(tok, Keyword.qed) => | 
| 58755 | 288 | find_block( | 
| 63424 | 289 | keywords.is_command(_, Keyword.qed), | 
| 58755 | 290 | t => | 
| 63424 | 291 | keywords.is_command(t, Keyword.proof_goal) || | 
| 292 | keywords.is_command(t, Keyword.theory_goal), | |
| 58755 | 293 | _ => false, | 
| 58762 | 294 | t => | 
| 63424 | 295 | keywords.is_command(t, Keyword.diag) || | 
| 296 | keywords.is_command(t, Keyword.proof) || | |
| 297 | keywords.is_command(t, Keyword.theory_goal), | |
| 63427 | 298 | reverse_caret_iterator()) | 
| 58755 | 299 | |
| 58752 | 300 | case Some(Text.Info(range1, tok)) if tok.is_begin => | 
| 58762 | 301 | find_block(_.is_begin, _.is_end, _ => false, _ => true, caret_iterator()) | 
| 58752 | 302 | |
| 303 | case Some(Text.Info(range1, tok)) if tok.is_end => | |
| 63427 | 304 | find_block(_.is_end, _.is_begin, _ => false, _ => true, reverse_caret_iterator()) | 
| 58763 | 305 |               match {
 | 
| 306 | case Some((_, range2)) => | |
| 63427 | 307 | reverse_caret_iterator(). | 
| 58800 
bfed1c26caed
explicit keyword category for commands that may start a block;
 wenzelm parents: 
58763diff
changeset | 308 | dropWhile(info => info.range != range2). | 
| 
bfed1c26caed
explicit keyword category for commands that may start a block;
 wenzelm parents: 
58763diff
changeset | 309 | dropWhile(info => info.range == range2). | 
| 
bfed1c26caed
explicit keyword category for commands that may start a block;
 wenzelm parents: 
58763diff
changeset | 310 | find(info => info.info.is_command || info.info.is_begin) | 
| 58763 | 311 |                   match {
 | 
| 58800 
bfed1c26caed
explicit keyword category for commands that may start a block;
 wenzelm parents: 
58763diff
changeset | 312 | case Some(Text.Info(range3, tok)) => | 
| 63424 | 313 | if (keywords.is_command(tok, Keyword.theory_block)) Some((range1, range3)) | 
| 58800 
bfed1c26caed
explicit keyword category for commands that may start a block;
 wenzelm parents: 
58763diff
changeset | 314 | else Some((range1, range2)) | 
| 58763 | 315 | case None => None | 
| 316 | } | |
| 317 | case None => None | |
| 318 | } | |
| 58752 | 319 | |
| 58749 
83b0f633190e
some structure matching, based on line token iterators;
 wenzelm parents: 
58748diff
changeset | 320 | case _ => None | 
| 58748 | 321 | } | 
| 66183 | 322 | case None => None | 
| 58748 | 323 | } | 
| 324 | } | |
| 325 | ||
| 58749 
83b0f633190e
some structure matching, based on line token iterators;
 wenzelm parents: 
58748diff
changeset | 326 | def getMatch(text_area: TextArea): StructureMatcher.Match = | 
| 
83b0f633190e
some structure matching, based on line token iterators;
 wenzelm parents: 
58748diff
changeset | 327 |       find_pair(text_area) match {
 | 
| 
83b0f633190e
some structure matching, based on line token iterators;
 wenzelm parents: 
58748diff
changeset | 328 | case Some((_, range)) => | 
| 
83b0f633190e
some structure matching, based on line token iterators;
 wenzelm parents: 
58748diff
changeset | 329 | val line = text_area.getBuffer.getLineOfOffset(range.start) | 
| 63422 | 330 | new StructureMatcher.Match(Matcher, line, range.start, line, range.stop) | 
| 58749 
83b0f633190e
some structure matching, based on line token iterators;
 wenzelm parents: 
58748diff
changeset | 331 | case None => null | 
| 
83b0f633190e
some structure matching, based on line token iterators;
 wenzelm parents: 
58748diff
changeset | 332 | } | 
| 
83b0f633190e
some structure matching, based on line token iterators;
 wenzelm parents: 
58748diff
changeset | 333 | |
| 75393 | 334 |     def selectMatch(text_area: TextArea): Unit = {
 | 
| 58803 
7a0f675eb671
proper selectMatch, e.g. relevant for S-click on gutter;
 wenzelm parents: 
58800diff
changeset | 335 | def get_span(offset: Text.Offset): Option[Text.Range] = | 
| 
7a0f675eb671
proper selectMatch, e.g. relevant for S-click on gutter;
 wenzelm parents: 
58800diff
changeset | 336 |         for {
 | 
| 59074 | 337 | syntax <- Isabelle.buffer_syntax(text_area.getBuffer) | 
| 58803 
7a0f675eb671
proper selectMatch, e.g. relevant for S-click on gutter;
 wenzelm parents: 
58800diff
changeset | 338 | span <- Token_Markup.command_span(syntax, text_area.getBuffer, offset) | 
| 
7a0f675eb671
proper selectMatch, e.g. relevant for S-click on gutter;
 wenzelm parents: 
58800diff
changeset | 339 | } yield span.range | 
| 
7a0f675eb671
proper selectMatch, e.g. relevant for S-click on gutter;
 wenzelm parents: 
58800diff
changeset | 340 | |
| 
7a0f675eb671
proper selectMatch, e.g. relevant for S-click on gutter;
 wenzelm parents: 
58800diff
changeset | 341 |       find_pair(text_area) match {
 | 
| 
7a0f675eb671
proper selectMatch, e.g. relevant for S-click on gutter;
 wenzelm parents: 
58800diff
changeset | 342 | case Some((r1, r2)) => | 
| 
7a0f675eb671
proper selectMatch, e.g. relevant for S-click on gutter;
 wenzelm parents: 
58800diff
changeset | 343 |           (get_span(r1.start), get_span(r2.start)) match {
 | 
| 
7a0f675eb671
proper selectMatch, e.g. relevant for S-click on gutter;
 wenzelm parents: 
58800diff
changeset | 344 | case (Some(range1), Some(range2)) => | 
| 
7a0f675eb671
proper selectMatch, e.g. relevant for S-click on gutter;
 wenzelm parents: 
58800diff
changeset | 345 | val start = range1.start min range2.start | 
| 
7a0f675eb671
proper selectMatch, e.g. relevant for S-click on gutter;
 wenzelm parents: 
58800diff
changeset | 346 | val stop = range1.stop max range2.stop | 
| 
7a0f675eb671
proper selectMatch, e.g. relevant for S-click on gutter;
 wenzelm parents: 
58800diff
changeset | 347 | |
| 
7a0f675eb671
proper selectMatch, e.g. relevant for S-click on gutter;
 wenzelm parents: 
58800diff
changeset | 348 | text_area.moveCaretPosition(stop, false) | 
| 
7a0f675eb671
proper selectMatch, e.g. relevant for S-click on gutter;
 wenzelm parents: 
58800diff
changeset | 349 | if (!text_area.isMultipleSelectionEnabled) text_area.selectNone | 
| 
7a0f675eb671
proper selectMatch, e.g. relevant for S-click on gutter;
 wenzelm parents: 
58800diff
changeset | 350 | text_area.addToSelection(new Selection.Range(start, stop)) | 
| 
7a0f675eb671
proper selectMatch, e.g. relevant for S-click on gutter;
 wenzelm parents: 
58800diff
changeset | 351 | case _ => | 
| 
7a0f675eb671
proper selectMatch, e.g. relevant for S-click on gutter;
 wenzelm parents: 
58800diff
changeset | 352 | } | 
| 
7a0f675eb671
proper selectMatch, e.g. relevant for S-click on gutter;
 wenzelm parents: 
58800diff
changeset | 353 | case None => | 
| 
7a0f675eb671
proper selectMatch, e.g. relevant for S-click on gutter;
 wenzelm parents: 
58800diff
changeset | 354 | } | 
| 58748 | 355 | } | 
| 356 | } | |
| 357 | } |