1 /* Title: Pure/PIDE/text.scala
2 Author: Fabian Immler, TU Munich
5 Basic operations on plain text.
11 import scala.collection.mutable
12 import scala.math.Ordering
13 import scala.util.Sorting
23 /* range -- with total quasi-ordering */
27 def apply(start: Offset): Range = Range(start, start)
29 object Ordering extends scala.math.Ordering[Text.Range]
31 def compare(r1: Text.Range, r2: Text.Range): Int = r1 compare r2
35 sealed case class Range(val start: Offset, val stop: Offset)
37 // denotation: {start} Un {i. start < i & i < stop}
39 error("Bad range: [" + start.toString + ":" + stop.toString + "]")
41 override def toString = "[" + start.toString + ":" + stop.toString + "]"
43 def map(f: Offset => Offset): Range = Range(f(start), f(stop))
44 def +(i: Offset): Range = map(_ + i)
45 def -(i: Offset): Range = map(_ - i)
47 def is_singularity: Boolean = start == stop
49 def contains(i: Offset): Boolean = start == i || start < i && i < stop
50 def contains(that: Range): Boolean = this.contains(that.start) && that.stop <= this.stop
51 def overlaps(that: Range): Boolean = this.contains(that.start) || that.contains(this.start)
52 def compare(that: Range): Int = if (overlaps(that)) 0 else this.start compare that.start
54 def apart(that: Range): Boolean =
55 (this.start max that.start) > (this.stop min that.stop)
57 def restrict(that: Range): Range =
58 Range(this.start max that.start, this.stop min that.stop)
60 def try_restrict(that: Range): Option[Range] =
61 if (this apart that) None
62 else Some(restrict(that))
64 def try_join(that: Range): Option[Range] =
65 if (this apart that) None
66 else Some(Range(this.start min that.start, this.stop max that.stop))
74 val empty: Perspective = Perspective(Nil)
76 def apply(ranges: Seq[Range]): Perspective =
78 val result = new mutable.ListBuffer[Text.Range]
79 var last: Option[Text.Range] = None
80 def ship(next: Option[Range]) { result ++= last; last = next }
82 for (range <- ranges.sortBy(_.start))
85 case None => ship(Some(range))
86 case Some(last_range) =>
87 last_range.try_join(range) match {
88 case None => ship(Some(range))
89 case joined => last = joined
94 new Perspective(result.toList)
98 class Perspective private(val ranges: List[Range]) // visible text partitioning in canonical order
100 def is_empty: Boolean = ranges.isEmpty
102 if (is_empty) Range(0)
103 else Range(ranges.head.start, ranges.last.stop)
104 override def toString = ranges.toString
108 /* information associated with text range */
110 sealed case class Info[A](val range: Text.Range, val info: A)
112 def restrict(r: Text.Range): Info[A] = Info(range.restrict(r), info)
113 def try_restrict(r: Text.Range): Option[Info[A]] =
114 try { Some(Info(range.restrict(r), info)) }
115 catch { case ERROR(_) => None }
118 type Markup = Info[XML.Elem]
125 def insert(start: Offset, text: String): Edit = new Edit(true, start, text)
126 def remove(start: Offset, text: String): Edit = new Edit(false, start, text)
129 class Edit private(val is_insert: Boolean, val start: Offset, val text: String)
131 override def toString =
132 (if (is_insert) "Insert(" else "Remove(") + (start, text).toString + ")"
135 /* transform offsets */
137 private def transform(do_insert: Boolean, i: Offset): Offset =
139 else if (do_insert) i + text.length
140 else (i - text.length) max start
142 def convert(i: Offset): Offset = transform(is_insert, i)
143 def revert(i: Offset): Offset = transform(!is_insert, i)
144 def convert(range: Range): Range = range.map(convert)
145 def revert(range: Range): Range = range.map(revert)
150 private def insert(i: Offset, string: String): String =
151 string.substring(0, i) + text + string.substring(i)
153 private def remove(i: Offset, count: Int, string: String): String =
154 string.substring(0, i) + string.substring(i + count)
156 def can_edit(string: String, shift: Int): Boolean =
157 shift <= start && start < shift + string.length
159 def edit(string: String, shift: Int): (Option[Edit], String) =
160 if (!can_edit(string, shift)) (Some(this), string)
161 else if (is_insert) (None, insert(start - shift, string))
163 val i = start - shift
164 val count = text.length min (string.length - i)
166 if (count == text.length) None
167 else Some(Edit.remove(start, text.substring(count)))
168 (rest, remove(i, count, string))