| author | wenzelm | 
| Wed, 21 Feb 2024 19:54:17 +0100 | |
| changeset 79683 | ade429ddb1fc | 
| parent 78243 | 0e221a8128e4 | 
| child 82142 | 508a673c87ac | 
| permissions | -rw-r--r-- | 
| 38425 | 1 | /* Title: Pure/PIDE/text.scala | 
| 34276 | 2 | Author: Fabian Immler, TU Munich | 
| 3 | Author: Makarius | |
| 4 | ||
| 38425 | 5 | Basic operations on plain text. | 
| 34276 | 6 | */ | 
| 7 | ||
| 8 | package isabelle | |
| 9 | ||
| 10 | ||
| 44379 | 11 | import scala.collection.mutable | 
| 12 | import scala.util.Sorting | |
| 13 | ||
| 14 | ||
| 75393 | 15 | object Text {
 | 
| 38477 | 16 | /* offset */ | 
| 38426 | 17 | |
| 18 | type Offset = Int | |
| 19 | ||
| 38477 | 20 | |
| 38565 
32b924a832c4
further clarification/unification of Position.Range and Text.Range concerning singularities: start offset is always included;
 wenzelm parents: 
38564diff
changeset | 21 | /* range -- with total quasi-ordering */ | 
| 38477 | 22 | |
| 75393 | 23 |   object Range {
 | 
| 38568 | 24 | def apply(start: Offset): Range = Range(start, start) | 
| 76235 | 25 | def length(text: CharSequence): Range = Range(0, text.length) | 
| 44379 | 26 | |
| 76235 | 27 | val zero: Range = apply(0) | 
| 78243 | 28 | val full: Range = apply(0, Int.MaxValue / 2) | 
| 56172 | 29 | val offside: Range = apply(-1) | 
| 30 | ||
| 75393 | 31 |     object Ordering extends scala.math.Ordering[Range] {
 | 
| 65154 | 32 | def compare(r1: Range, r2: Range): Int = r1 compare r2 | 
| 44379 | 33 | } | 
| 38568 | 34 | } | 
| 35 | ||
| 75393 | 36 |   sealed case class Range(start: Offset, stop: Offset) {
 | 
| 38565 
32b924a832c4
further clarification/unification of Position.Range and Text.Range concerning singularities: start offset is always included;
 wenzelm parents: 
38564diff
changeset | 37 |     // denotation: {start} Un {i. start < i & i < stop}
 | 
| 43425 | 38 | if (start > stop) | 
| 64546 
134ae7da2ccf
clarified output: avoid confusion with line:column notation;
 wenzelm parents: 
64370diff
changeset | 39 |       error("Bad range: [" + start.toString + ".." + stop.toString + "]")
 | 
| 38477 | 40 | |
| 64546 
134ae7da2ccf
clarified output: avoid confusion with line:column notation;
 wenzelm parents: 
64370diff
changeset | 41 | override def toString: String = "[" + start.toString + ".." + stop.toString + "]" | 
| 38563 | 42 | |
| 65155 
25bccf5bf33e
clarified native Text.Offset versus Text.Length index Int;
 wenzelm parents: 
65154diff
changeset | 43 | def length: Offset = stop - start | 
| 47542 
26d0a76fef0a
more robust Sendback handling: JVM/jEdit paranoia for case matching, treat Pretty body not just XML.Text, replace proper_range only (without trailing whitespace);
 wenzelm parents: 
46712diff
changeset | 44 | |
| 38427 | 45 | def map(f: Offset => Offset): Range = Range(f(start), f(stop)) | 
| 56308 | 46 | def +(i: Offset): Range = if (i == 0) this else map(_ + i) | 
| 47 | def -(i: Offset): Range = if (i == 0) this else map(_ - i) | |
| 38662 | 48 | |
| 38725 | 49 | def is_singularity: Boolean = start == stop | 
| 56590 
d01d183e84ea
clarified treatment of markup ranges wrt. revert/convert: inflate_singularity allows to retrieve information like language_context more reliably during editing;
 wenzelm parents: 
56473diff
changeset | 50 | def inflate_singularity: Range = if (is_singularity) Range(start, start + 1) else this | 
| 38662 | 51 | |
| 58749 
83b0f633190e
some structure matching, based on line token iterators;
 wenzelm parents: 
57912diff
changeset | 52 | def touches(i: Offset): Boolean = start <= i && i <= stop | 
| 
83b0f633190e
some structure matching, based on line token iterators;
 wenzelm parents: 
57912diff
changeset | 53 | |
| 38565 
32b924a832c4
further clarification/unification of Position.Range and Text.Range concerning singularities: start offset is always included;
 wenzelm parents: 
38564diff
changeset | 54 | def contains(i: Offset): Boolean = start == i || start < i && i < stop | 
| 
32b924a832c4
further clarification/unification of Position.Range and Text.Range concerning singularities: start offset is always included;
 wenzelm parents: 
38564diff
changeset | 55 | def contains(that: Range): Boolean = this.contains(that.start) && that.stop <= this.stop | 
| 
32b924a832c4
further clarification/unification of Position.Range and Text.Range concerning singularities: start offset is always included;
 wenzelm parents: 
38564diff
changeset | 56 | def overlaps(that: Range): Boolean = this.contains(that.start) || that.contains(this.start) | 
| 
32b924a832c4
further clarification/unification of Position.Range and Text.Range concerning singularities: start offset is always included;
 wenzelm parents: 
38564diff
changeset | 57 | def compare(that: Range): Int = if (overlaps(that)) 0 else this.start compare that.start | 
| 38485 | 58 | |
| 45240 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 59 | def apart(that: Range): Boolean = | 
| 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 60 | (this.start max that.start) > (this.stop min that.stop) | 
| 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 61 | |
| 38564 | 62 | def restrict(that: Range): Range = | 
| 38485 | 63 | Range(this.start max that.start, this.stop min that.stop) | 
| 43428 
b41dea5772c6
more robust treatment of partial range restriction;
 wenzelm parents: 
43425diff
changeset | 64 | |
| 
b41dea5772c6
more robust treatment of partial range restriction;
 wenzelm parents: 
43425diff
changeset | 65 | def try_restrict(that: Range): Option[Range] = | 
| 45240 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 66 | if (this apart that) None | 
| 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 67 | else Some(restrict(that)) | 
| 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 68 | |
| 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 69 | def try_join(that: Range): Option[Range] = | 
| 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 70 | if (this apart that) None | 
| 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 71 | else Some(Range(this.start min that.start, this.stop max that.stop)) | 
| 65522 | 72 | |
| 73 | def substring(text: String): String = text.substring(start, stop) | |
| 66114 | 74 | |
| 75 | def try_substring(text: String): Option[String] = | |
| 76 |       try { Some(substring(text)) }
 | |
| 77 |       catch { case _: IndexOutOfBoundsException => None }
 | |
| 38427 | 78 | } | 
| 38426 | 79 | |
| 80 | ||
| 44379 | 81 | /* perspective */ | 
| 82 | ||
| 75393 | 83 |   object Perspective {
 | 
| 44474 | 84 | val empty: Perspective = Perspective(Nil) | 
| 44379 | 85 | |
| 64678 | 86 | def full: Perspective = Perspective(List(Range.full)) | 
| 46576 
ae9286f64574
approximate Perspective.full within the bounds of the JVM;
 wenzelm parents: 
46207diff
changeset | 87 | |
| 75393 | 88 |     def apply(ranges: List[Range]): Perspective = {
 | 
| 65154 | 89 | val result = new mutable.ListBuffer[Range] | 
| 90 | var last: Option[Range] = None | |
| 73340 | 91 |       def ship(next: Option[Range]): Unit = { result ++= last; last = next }
 | 
| 45240 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 92 | |
| 75393 | 93 |       for (range <- ranges.sortBy(_.start)) {
 | 
| 44473 | 94 |         last match {
 | 
| 45240 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 95 | case None => ship(Some(range)) | 
| 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 96 | case Some(last_range) => | 
| 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 97 |             last_range.try_join(range) match {
 | 
| 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 98 | case None => ship(Some(range)) | 
| 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 99 | case joined => last = joined | 
| 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 100 | } | 
| 44473 | 101 | } | 
| 44379 | 102 | } | 
| 45240 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 103 | ship(None) | 
| 44473 | 104 | new Perspective(result.toList) | 
| 44379 | 105 | } | 
| 44473 | 106 | } | 
| 107 | ||
| 46712 | 108 | final class Perspective private( | 
| 75393 | 109 | val ranges: List[Range] // visible text partitioning in canonical order | 
| 110 |   ) {
 | |
| 44473 | 111 | def is_empty: Boolean = ranges.isEmpty | 
| 112 | def range: Range = | |
| 76235 | 113 | if (is_empty) Range.zero | 
| 44473 | 114 | else Range(ranges.head.start, ranges.last.stop) | 
| 45631 
6bdf8b926f50
recovered structural equality from 9d97bd3c086a, otherwise update_perspective might be issued over and over again, canceling a pending/slow execution;
 wenzelm parents: 
45470diff
changeset | 115 | |
| 
6bdf8b926f50
recovered structural equality from 9d97bd3c086a, otherwise update_perspective might be issued over and over again, canceling a pending/slow execution;
 wenzelm parents: 
45470diff
changeset | 116 | override def hashCode: Int = ranges.hashCode | 
| 
6bdf8b926f50
recovered structural equality from 9d97bd3c086a, otherwise update_perspective might be issued over and over again, canceling a pending/slow execution;
 wenzelm parents: 
45470diff
changeset | 117 | override def equals(that: Any): Boolean = | 
| 
6bdf8b926f50
recovered structural equality from 9d97bd3c086a, otherwise update_perspective might be issued over and over again, canceling a pending/slow execution;
 wenzelm parents: 
45470diff
changeset | 118 |       that match {
 | 
| 
6bdf8b926f50
recovered structural equality from 9d97bd3c086a, otherwise update_perspective might be issued over and over again, canceling a pending/slow execution;
 wenzelm parents: 
45470diff
changeset | 119 | case other: Perspective => ranges == other.ranges | 
| 
6bdf8b926f50
recovered structural equality from 9d97bd3c086a, otherwise update_perspective might be issued over and over again, canceling a pending/slow execution;
 wenzelm parents: 
45470diff
changeset | 120 | case _ => false | 
| 
6bdf8b926f50
recovered structural equality from 9d97bd3c086a, otherwise update_perspective might be issued over and over again, canceling a pending/slow execution;
 wenzelm parents: 
45470diff
changeset | 121 | } | 
| 57912 | 122 | override def toString: String = ranges.toString | 
| 44379 | 123 | } | 
| 124 | ||
| 125 | ||
| 38577 | 126 | /* information associated with text range */ | 
| 127 | ||
| 75393 | 128 |   sealed case class Info[A](range: Range, info: A) {
 | 
| 65154 | 129 | def restrict(r: Range): Info[A] = Info(range.restrict(r), info) | 
| 130 | def try_restrict(r: Range): Option[Info[A]] = range.try_restrict(r).map(Info(_, info)) | |
| 65132 | 131 | |
| 132 | def map[B](f: A => B): Info[B] = Info(range, f(info)) | |
| 38577 | 133 | } | 
| 134 | ||
| 45470 | 135 | type Markup = Info[XML.Elem] | 
| 45455 | 136 | |
| 38577 | 137 | |
| 38426 | 138 | /* editing */ | 
| 34286 | 139 | |
| 75393 | 140 |   object Edit {
 | 
| 38426 | 141 | def insert(start: Offset, text: String): Edit = new Edit(true, start, text) | 
| 142 | def remove(start: Offset, text: String): Edit = new Edit(false, start, text) | |
| 65156 | 143 | def inserts(start: Offset, text: String): List[Edit] = | 
| 144 | if (text == "") Nil else List(insert(start, text)) | |
| 145 | def removes(start: Offset, text: String): List[Edit] = | |
| 146 | if (text == "") Nil else List(remove(start, text)) | |
| 64816 | 147 | def replace(start: Offset, old_text: String, new_text: String): List[Edit] = | 
| 148 | if (old_text == new_text) Nil | |
| 65156 | 149 | else removes(start, old_text) ::: inserts(start, new_text) | 
| 38425 | 150 | } | 
| 34286 | 151 | |
| 75393 | 152 |   final class Edit private(val is_insert: Boolean, val start: Offset, val text: String) {
 | 
| 57912 | 153 | override def toString: String = | 
| 38425 | 154 |       (if (is_insert) "Insert(" else "Remove(") + (start, text).toString + ")"
 | 
| 34286 | 155 | |
| 156 | ||
| 38425 | 157 | /* transform offsets */ | 
| 34286 | 158 | |
| 38426 | 159 | private def transform(do_insert: Boolean, i: Offset): Offset = | 
| 160 | if (i < start) i | |
| 43425 | 161 | else if (do_insert) i + text.length | 
| 38426 | 162 | else (i - text.length) max start | 
| 34286 | 163 | |
| 43425 | 164 | def convert(i: Offset): Offset = transform(is_insert, i) | 
| 165 | def revert(i: Offset): Offset = transform(!is_insert, i) | |
| 38425 | 166 | |
| 34286 | 167 | |
| 38425 | 168 | /* edit strings */ | 
| 169 | ||
| 38426 | 170 | private def insert(i: Offset, string: String): String = | 
| 171 | string.substring(0, i) + text + string.substring(i) | |
| 34276 | 172 | |
| 65155 
25bccf5bf33e
clarified native Text.Offset versus Text.Length index Int;
 wenzelm parents: 
65154diff
changeset | 173 | private def remove(i: Offset, count: Offset, string: String): String = | 
| 38426 | 174 | string.substring(0, i) + string.substring(i + count) | 
| 38425 | 175 | |
| 65155 
25bccf5bf33e
clarified native Text.Offset versus Text.Length index Int;
 wenzelm parents: 
65154diff
changeset | 176 | def can_edit(string: String, shift: Offset): Boolean = | 
| 38425 | 177 | shift <= start && start < shift + string.length | 
| 178 | ||
| 65155 
25bccf5bf33e
clarified native Text.Offset versus Text.Length index Int;
 wenzelm parents: 
65154diff
changeset | 179 | def edit(string: String, shift: Offset): (Option[Edit], String) = | 
| 38425 | 180 | if (!can_edit(string, shift)) (Some(this), string) | 
| 181 | else if (is_insert) (None, insert(start - shift, string)) | |
| 182 |       else {
 | |
| 38426 | 183 | val i = start - shift | 
| 184 | val count = text.length min (string.length - i) | |
| 38425 | 185 | val rest = | 
| 186 | if (count == text.length) None | |
| 187 | else Some(Edit.remove(start, text.substring(count))) | |
| 38426 | 188 | (rest, remove(i, count, string)) | 
| 38425 | 189 | } | 
| 190 | } | |
| 34276 | 191 | } |