| author | wenzelm | 
| Wed, 21 Mar 2012 23:26:35 +0100 | |
| changeset 47069 | 451fc10a81f3 | 
| parent 46712 | 8650d9a95736 | 
| child 47542 | 26d0a76fef0a | 
| permissions | -rw-r--r-- | 
| 38425 | 1 | /* Title: Pure/PIDE/text.scala | 
| 45673 
cd41e3903fbf
separate compilation of PIDE vs. Pure sources, which enables independent Scala library;
 wenzelm parents: 
45667diff
changeset | 2 | Module: PIDE | 
| 34276 | 3 | Author: Fabian Immler, TU Munich | 
| 4 | Author: Makarius | |
| 5 | ||
| 38425 | 6 | Basic operations on plain text. | 
| 34276 | 7 | */ | 
| 8 | ||
| 9 | package isabelle | |
| 10 | ||
| 11 | ||
| 44379 | 12 | import scala.collection.mutable | 
| 13 | import scala.math.Ordering | |
| 14 | import scala.util.Sorting | |
| 15 | ||
| 16 | ||
| 38425 | 17 | object Text | 
| 34276 | 18 | {
 | 
| 38477 | 19 | /* offset */ | 
| 38426 | 20 | |
| 21 | type Offset = Int | |
| 22 | ||
| 38477 | 23 | |
| 38565 
32b924a832c4
further clarification/unification of Position.Range and Text.Range concerning singularities: start offset is always included;
 wenzelm parents: 
38564diff
changeset | 24 | /* range -- with total quasi-ordering */ | 
| 38477 | 25 | |
| 38568 | 26 | object Range | 
| 27 |   {
 | |
| 28 | def apply(start: Offset): Range = Range(start, start) | |
| 44379 | 29 | |
| 30 | object Ordering extends scala.math.Ordering[Text.Range] | |
| 31 |     {
 | |
| 32 | def compare(r1: Text.Range, r2: Text.Range): Int = r1 compare r2 | |
| 33 | } | |
| 38568 | 34 | } | 
| 35 | ||
| 38426 | 36 | sealed case class Range(val start: Offset, val stop: Offset) | 
| 38427 | 37 |   {
 | 
| 38565 
32b924a832c4
further clarification/unification of Position.Range and Text.Range concerning singularities: start offset is always included;
 wenzelm parents: 
38564diff
changeset | 38 |     // denotation: {start} Un {i. start < i & i < stop}
 | 
| 43425 | 39 | if (start > stop) | 
| 40 |       error("Bad range: [" + start.toString + ":" + stop.toString + "]")
 | |
| 38477 | 41 | |
| 38563 | 42 | override def toString = "[" + start.toString + ":" + stop.toString + "]" | 
| 43 | ||
| 38427 | 44 | def map(f: Offset => Offset): Range = Range(f(start), f(stop)) | 
| 45 | def +(i: Offset): Range = map(_ + i) | |
| 38570 | 46 | def -(i: Offset): Range = map(_ - i) | 
| 38662 | 47 | |
| 38725 | 48 | def is_singularity: Boolean = start == stop | 
| 38662 | 49 | |
| 38565 
32b924a832c4
further clarification/unification of Position.Range and Text.Range concerning singularities: start offset is always included;
 wenzelm parents: 
38564diff
changeset | 50 | 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 | 51 | 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 | 52 | 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 | 53 | def compare(that: Range): Int = if (overlaps(that)) 0 else this.start compare that.start | 
| 38485 | 54 | |
| 45240 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 55 | 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 | 56 | (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 | 57 | |
| 38564 | 58 | def restrict(that: Range): Range = | 
| 38485 | 59 | Range(this.start max that.start, this.stop min that.stop) | 
| 43428 
b41dea5772c6
more robust treatment of partial range restriction;
 wenzelm parents: 
43425diff
changeset | 60 | |
| 
b41dea5772c6
more robust treatment of partial range restriction;
 wenzelm parents: 
43425diff
changeset | 61 | 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 | 62 | 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 | 63 | 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 | 64 | |
| 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 65 | 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 | 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(Range(this.start min that.start, this.stop max that.stop)) | 
| 38427 | 68 | } | 
| 38426 | 69 | |
| 70 | ||
| 44379 | 71 | /* perspective */ | 
| 72 | ||
| 44473 | 73 | object Perspective | 
| 44379 | 74 |   {
 | 
| 44474 | 75 | val empty: Perspective = Perspective(Nil) | 
| 44379 | 76 | |
| 46576 
ae9286f64574
approximate Perspective.full within the bounds of the JVM;
 wenzelm parents: 
46207diff
changeset | 77 | def full: Perspective = Perspective(List(Range(0, Integer.MAX_VALUE / 2))) | 
| 
ae9286f64574
approximate Perspective.full within the bounds of the JVM;
 wenzelm parents: 
46207diff
changeset | 78 | |
| 44473 | 79 | def apply(ranges: Seq[Range]): Perspective = | 
| 44379 | 80 |     {
 | 
| 44473 | 81 | val result = new mutable.ListBuffer[Text.Range] | 
| 82 | var last: Option[Text.Range] = None | |
| 45240 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 83 |       def ship(next: Option[Range]) { result ++= last; last = next }
 | 
| 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 84 | |
| 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 85 | for (range <- ranges.sortBy(_.start)) | 
| 44473 | 86 |       {
 | 
| 87 |         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 | 88 | 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 | 89 | 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 | 90 |             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 | 91 | 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 | 92 | 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 | 93 | } | 
| 44473 | 94 | } | 
| 44379 | 95 | } | 
| 45240 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 96 | ship(None) | 
| 44473 | 97 | new Perspective(result.toList) | 
| 44379 | 98 | } | 
| 44473 | 99 | } | 
| 100 | ||
| 46712 | 101 | final class Perspective private( | 
| 102 | val ranges: List[Range]) // visible text partitioning in canonical order | |
| 44473 | 103 |   {
 | 
| 104 | def is_empty: Boolean = ranges.isEmpty | |
| 105 | def range: Range = | |
| 106 | if (is_empty) Range(0) | |
| 107 | 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 | 108 | |
| 
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 | 109 | 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 | 110 | 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 | 111 |       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 | 112 | 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 | 113 | 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 | 114 | } | 
| 45240 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 wenzelm parents: 
44474diff
changeset | 115 | override def toString = ranges.toString | 
| 44379 | 116 | } | 
| 117 | ||
| 118 | ||
| 38577 | 119 | /* information associated with text range */ | 
| 120 | ||
| 43714 | 121 | sealed case class Info[A](val range: Text.Range, val info: A) | 
| 38577 | 122 |   {
 | 
| 123 | def restrict(r: Text.Range): Info[A] = Info(range.restrict(r), info) | |
| 46207 | 124 | def try_restrict(r: Text.Range): Option[Info[A]] = range.try_restrict(r).map(Info(_, info)) | 
| 38577 | 125 | } | 
| 126 | ||
| 45470 | 127 | type Markup = Info[XML.Elem] | 
| 45455 | 128 | |
| 38577 | 129 | |
| 38426 | 130 | /* editing */ | 
| 34286 | 131 | |
| 38425 | 132 | object Edit | 
| 133 |   {
 | |
| 38426 | 134 | def insert(start: Offset, text: String): Edit = new Edit(true, start, text) | 
| 135 | def remove(start: Offset, text: String): Edit = new Edit(false, start, text) | |
| 38425 | 136 | } | 
| 34286 | 137 | |
| 46712 | 138 | final class Edit private(val is_insert: Boolean, val start: Offset, val text: String) | 
| 38425 | 139 |   {
 | 
| 140 | override def toString = | |
| 141 |       (if (is_insert) "Insert(" else "Remove(") + (start, text).toString + ")"
 | |
| 34286 | 142 | |
| 143 | ||
| 38425 | 144 | /* transform offsets */ | 
| 34286 | 145 | |
| 38426 | 146 | private def transform(do_insert: Boolean, i: Offset): Offset = | 
| 147 | if (i < start) i | |
| 43425 | 148 | else if (do_insert) i + text.length | 
| 38426 | 149 | else (i - text.length) max start | 
| 34286 | 150 | |
| 43425 | 151 | def convert(i: Offset): Offset = transform(is_insert, i) | 
| 152 | def revert(i: Offset): Offset = transform(!is_insert, i) | |
| 153 | def convert(range: Range): Range = range.map(convert) | |
| 154 | def revert(range: Range): Range = range.map(revert) | |
| 38425 | 155 | |
| 34286 | 156 | |
| 38425 | 157 | /* edit strings */ | 
| 158 | ||
| 38426 | 159 | private def insert(i: Offset, string: String): String = | 
| 160 | string.substring(0, i) + text + string.substring(i) | |
| 34276 | 161 | |
| 38426 | 162 | private def remove(i: Offset, count: Int, string: String): String = | 
| 163 | string.substring(0, i) + string.substring(i + count) | |
| 38425 | 164 | |
| 165 | def can_edit(string: String, shift: Int): Boolean = | |
| 166 | shift <= start && start < shift + string.length | |
| 167 | ||
| 168 | def edit(string: String, shift: Int): (Option[Edit], String) = | |
| 169 | if (!can_edit(string, shift)) (Some(this), string) | |
| 170 | else if (is_insert) (None, insert(start - shift, string)) | |
| 171 |       else {
 | |
| 38426 | 172 | val i = start - shift | 
| 173 | val count = text.length min (string.length - i) | |
| 38425 | 174 | val rest = | 
| 175 | if (count == text.length) None | |
| 176 | else Some(Edit.remove(start, text.substring(count))) | |
| 38426 | 177 | (rest, remove(i, count, string)) | 
| 38425 | 178 | } | 
| 179 | } | |
| 34276 | 180 | } |