| author | wenzelm | 
| Sat, 10 Dec 2016 17:22:47 +0100 | |
| changeset 64546 | 134ae7da2ccf | 
| parent 64370 | 865b39487b5d | 
| child 64678 | 914dffe59cc5 | 
| 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  | 
||
| 38425 | 15  | 
object Text  | 
| 34276 | 16  | 
{
 | 
| 38477 | 17  | 
/* offset */  | 
| 38426 | 18  | 
|
19  | 
type Offset = Int  | 
|
20  | 
||
| 38477 | 21  | 
|
| 
38565
 
32b924a832c4
further clarification/unification of Position.Range and Text.Range concerning singularities: start offset is always included;
 
wenzelm 
parents: 
38564 
diff
changeset
 | 
22  | 
/* range -- with total quasi-ordering */  | 
| 38477 | 23  | 
|
| 38568 | 24  | 
object Range  | 
25  | 
  {
 | 
|
26  | 
def apply(start: Offset): Range = Range(start, start)  | 
|
| 44379 | 27  | 
|
| 56172 | 28  | 
val offside: Range = apply(-1)  | 
29  | 
||
| 44379 | 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  | 
||
| 60215 | 36  | 
sealed case class Range(start: Offset, stop: Offset)  | 
| 38427 | 37  | 
  {
 | 
| 
38565
 
32b924a832c4
further clarification/unification of Position.Range and Text.Range concerning singularities: start offset is always included;
 
wenzelm 
parents: 
38564 
diff
changeset
 | 
38  | 
    // denotation: {start} Un {i. start < i & i < stop}
 | 
| 43425 | 39  | 
if (start > stop)  | 
| 
64546
 
134ae7da2ccf
clarified output: avoid confusion with line:column notation;
 
wenzelm 
parents: 
64370 
diff
changeset
 | 
40  | 
      error("Bad range: [" + start.toString + ".." + stop.toString + "]")
 | 
| 38477 | 41  | 
|
| 
64546
 
134ae7da2ccf
clarified output: avoid confusion with line:column notation;
 
wenzelm 
parents: 
64370 
diff
changeset
 | 
42  | 
override def toString: String = "[" + start.toString + ".." + stop.toString + "]"  | 
| 38563 | 43  | 
|
| 
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: 
46712 
diff
changeset
 | 
44  | 
def length: Int = stop - start  | 
| 
 
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: 
46712 
diff
changeset
 | 
45  | 
|
| 38427 | 46  | 
def map(f: Offset => Offset): Range = Range(f(start), f(stop))  | 
| 56308 | 47  | 
def +(i: Offset): Range = if (i == 0) this else map(_ + i)  | 
48  | 
def -(i: Offset): Range = if (i == 0) this else map(_ - i)  | 
|
| 38662 | 49  | 
|
| 38725 | 50  | 
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: 
56473 
diff
changeset
 | 
51  | 
def inflate_singularity: Range = if (is_singularity) Range(start, start + 1) else this  | 
| 38662 | 52  | 
|
| 
58749
 
83b0f633190e
some structure matching, based on line token iterators;
 
wenzelm 
parents: 
57912 
diff
changeset
 | 
53  | 
def touches(i: Offset): Boolean = start <= i && i <= stop  | 
| 
 
83b0f633190e
some structure matching, based on line token iterators;
 
wenzelm 
parents: 
57912 
diff
changeset
 | 
54  | 
|
| 
38565
 
32b924a832c4
further clarification/unification of Position.Range and Text.Range concerning singularities: start offset is always included;
 
wenzelm 
parents: 
38564 
diff
changeset
 | 
55  | 
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: 
38564 
diff
changeset
 | 
56  | 
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: 
38564 
diff
changeset
 | 
57  | 
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: 
38564 
diff
changeset
 | 
58  | 
def compare(that: Range): Int = if (overlaps(that)) 0 else this.start compare that.start  | 
| 38485 | 59  | 
|
| 
45240
 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 
wenzelm 
parents: 
44474 
diff
changeset
 | 
60  | 
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: 
44474 
diff
changeset
 | 
61  | 
(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: 
44474 
diff
changeset
 | 
62  | 
|
| 38564 | 63  | 
def restrict(that: Range): Range =  | 
| 38485 | 64  | 
Range(this.start max that.start, this.stop min that.stop)  | 
| 
43428
 
b41dea5772c6
more robust treatment of partial range restriction;
 
wenzelm 
parents: 
43425 
diff
changeset
 | 
65  | 
|
| 
 
b41dea5772c6
more robust treatment of partial range restriction;
 
wenzelm 
parents: 
43425 
diff
changeset
 | 
66  | 
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: 
44474 
diff
changeset
 | 
67  | 
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: 
44474 
diff
changeset
 | 
68  | 
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: 
44474 
diff
changeset
 | 
69  | 
|
| 
 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 
wenzelm 
parents: 
44474 
diff
changeset
 | 
70  | 
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: 
44474 
diff
changeset
 | 
71  | 
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: 
44474 
diff
changeset
 | 
72  | 
else Some(Range(this.start min that.start, this.stop max that.stop))  | 
| 38427 | 73  | 
}  | 
| 38426 | 74  | 
|
75  | 
||
| 44379 | 76  | 
/* perspective */  | 
77  | 
||
| 44473 | 78  | 
object Perspective  | 
| 44379 | 79  | 
  {
 | 
| 44474 | 80  | 
val empty: Perspective = Perspective(Nil)  | 
| 44379 | 81  | 
|
| 
46576
 
ae9286f64574
approximate Perspective.full within the bounds of the JVM;
 
wenzelm 
parents: 
46207 
diff
changeset
 | 
82  | 
def full: Perspective = Perspective(List(Range(0, Integer.MAX_VALUE / 2)))  | 
| 
 
ae9286f64574
approximate Perspective.full within the bounds of the JVM;
 
wenzelm 
parents: 
46207 
diff
changeset
 | 
83  | 
|
| 44473 | 84  | 
def apply(ranges: Seq[Range]): Perspective =  | 
| 44379 | 85  | 
    {
 | 
| 44473 | 86  | 
val result = new mutable.ListBuffer[Text.Range]  | 
87  | 
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: 
44474 
diff
changeset
 | 
88  | 
      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: 
44474 
diff
changeset
 | 
89  | 
|
| 
 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 
wenzelm 
parents: 
44474 
diff
changeset
 | 
90  | 
for (range <- ranges.sortBy(_.start))  | 
| 44473 | 91  | 
      {
 | 
92  | 
        last match {
 | 
|
| 
45240
 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 
wenzelm 
parents: 
44474 
diff
changeset
 | 
93  | 
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: 
44474 
diff
changeset
 | 
94  | 
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: 
44474 
diff
changeset
 | 
95  | 
            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: 
44474 
diff
changeset
 | 
96  | 
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: 
44474 
diff
changeset
 | 
97  | 
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: 
44474 
diff
changeset
 | 
98  | 
}  | 
| 44473 | 99  | 
}  | 
| 44379 | 100  | 
}  | 
| 
45240
 
9d97bd3c086a
proper normal form for Perspective.ranges (overlapping ranges could be joined in wrong order, crashing multiple editor views);
 
wenzelm 
parents: 
44474 
diff
changeset
 | 
101  | 
ship(None)  | 
| 44473 | 102  | 
new Perspective(result.toList)  | 
| 44379 | 103  | 
}  | 
| 44473 | 104  | 
}  | 
105  | 
||
| 46712 | 106  | 
final class Perspective private(  | 
107  | 
val ranges: List[Range]) // visible text partitioning in canonical order  | 
|
| 44473 | 108  | 
  {
 | 
109  | 
def is_empty: Boolean = ranges.isEmpty  | 
|
110  | 
def range: Range =  | 
|
111  | 
if (is_empty) Range(0)  | 
|
112  | 
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: 
45470 
diff
changeset
 | 
113  | 
|
| 
 
6bdf8b926f50
recovered structural equality from 9d97bd3c086a, otherwise update_perspective might be issued over and over again, canceling a pending/slow execution;
 
wenzelm 
parents: 
45470 
diff
changeset
 | 
114  | 
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: 
45470 
diff
changeset
 | 
115  | 
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: 
45470 
diff
changeset
 | 
116  | 
      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: 
45470 
diff
changeset
 | 
117  | 
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: 
45470 
diff
changeset
 | 
118  | 
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: 
45470 
diff
changeset
 | 
119  | 
}  | 
| 57912 | 120  | 
override def toString: String = ranges.toString  | 
| 44379 | 121  | 
}  | 
122  | 
||
123  | 
||
| 38577 | 124  | 
/* information associated with text range */  | 
125  | 
||
| 60215 | 126  | 
sealed case class Info[A](range: Text.Range, info: A)  | 
| 38577 | 127  | 
  {
 | 
128  | 
def restrict(r: Text.Range): Info[A] = Info(range.restrict(r), info)  | 
|
| 46207 | 129  | 
def try_restrict(r: Text.Range): Option[Info[A]] = range.try_restrict(r).map(Info(_, info))  | 
| 38577 | 130  | 
}  | 
131  | 
||
| 45470 | 132  | 
type Markup = Info[XML.Elem]  | 
| 45455 | 133  | 
|
| 38577 | 134  | 
|
| 38426 | 135  | 
/* editing */  | 
| 34286 | 136  | 
|
| 38425 | 137  | 
object Edit  | 
138  | 
  {
 | 
|
| 38426 | 139  | 
def insert(start: Offset, text: String): Edit = new Edit(true, start, text)  | 
140  | 
def remove(start: Offset, text: String): Edit = new Edit(false, start, text)  | 
|
| 38425 | 141  | 
}  | 
| 34286 | 142  | 
|
| 46712 | 143  | 
final class Edit private(val is_insert: Boolean, val start: Offset, val text: String)  | 
| 38425 | 144  | 
  {
 | 
| 57912 | 145  | 
override def toString: String =  | 
| 38425 | 146  | 
      (if (is_insert) "Insert(" else "Remove(") + (start, text).toString + ")"
 | 
| 34286 | 147  | 
|
148  | 
||
| 38425 | 149  | 
/* transform offsets */  | 
| 34286 | 150  | 
|
| 38426 | 151  | 
private def transform(do_insert: Boolean, i: Offset): Offset =  | 
152  | 
if (i < start) i  | 
|
| 43425 | 153  | 
else if (do_insert) i + text.length  | 
| 38426 | 154  | 
else (i - text.length) max start  | 
| 34286 | 155  | 
|
| 43425 | 156  | 
def convert(i: Offset): Offset = transform(is_insert, i)  | 
157  | 
def revert(i: Offset): Offset = transform(!is_insert, i)  | 
|
| 38425 | 158  | 
|
| 34286 | 159  | 
|
| 38425 | 160  | 
/* edit strings */  | 
161  | 
||
| 38426 | 162  | 
private def insert(i: Offset, string: String): String =  | 
163  | 
string.substring(0, i) + text + string.substring(i)  | 
|
| 34276 | 164  | 
|
| 38426 | 165  | 
private def remove(i: Offset, count: Int, string: String): String =  | 
166  | 
string.substring(0, i) + string.substring(i + count)  | 
|
| 38425 | 167  | 
|
168  | 
def can_edit(string: String, shift: Int): Boolean =  | 
|
169  | 
shift <= start && start < shift + string.length  | 
|
170  | 
||
171  | 
def edit(string: String, shift: Int): (Option[Edit], String) =  | 
|
172  | 
if (!can_edit(string, shift)) (Some(this), string)  | 
|
173  | 
else if (is_insert) (None, insert(start - shift, string))  | 
|
174  | 
      else {
 | 
|
| 38426 | 175  | 
val i = start - shift  | 
176  | 
val count = text.length min (string.length - i)  | 
|
| 38425 | 177  | 
val rest =  | 
178  | 
if (count == text.length) None  | 
|
179  | 
else Some(Edit.remove(start, text.substring(count)))  | 
|
| 38426 | 180  | 
(rest, remove(i, count, string))  | 
| 38425 | 181  | 
}  | 
182  | 
}  | 
|
| 34276 | 183  | 
}  |