64611
|
1 |
/* Title: Pure/PIDE/line.scala
|
64605
|
2 |
Author: Makarius
|
|
3 |
|
|
4 |
Line-oriented text.
|
|
5 |
*/
|
|
6 |
|
64611
|
7 |
package isabelle
|
64605
|
8 |
|
|
9 |
|
|
10 |
import scala.annotation.tailrec
|
|
11 |
|
|
12 |
|
|
13 |
object Line
|
|
14 |
{
|
64615
|
15 |
/* length wrt. encoding */
|
|
16 |
|
|
17 |
trait Length { def apply(text: String): Int }
|
|
18 |
object Length extends Length { def apply(text: String): Int = text.length }
|
|
19 |
|
|
20 |
|
64605
|
21 |
/* position */
|
|
22 |
|
|
23 |
object Position
|
|
24 |
{
|
|
25 |
val zero: Position = Position(0, 0)
|
|
26 |
}
|
|
27 |
|
|
28 |
sealed case class Position(line: Int, column: Int)
|
|
29 |
{
|
|
30 |
def line1: Int = line + 1
|
|
31 |
def column1: Int = column + 1
|
|
32 |
def print: String = line1.toString + ":" + column1.toString
|
|
33 |
|
|
34 |
def compare(that: Position): Int =
|
|
35 |
line compare that.line match {
|
|
36 |
case 0 => column compare that.column
|
|
37 |
case i => i
|
|
38 |
}
|
|
39 |
|
|
40 |
private def advance[A](iterator: Iterator[A], is_newline: A => Boolean): Position =
|
|
41 |
{
|
|
42 |
var l = line
|
|
43 |
var c = column
|
|
44 |
for (x <- iterator) {
|
|
45 |
if (is_newline(x)) { l += 1; c = 0 } else c += 1
|
|
46 |
}
|
|
47 |
Position(l, c)
|
|
48 |
}
|
|
49 |
|
|
50 |
def advance(text: String): Position =
|
|
51 |
if (text.isEmpty) this else advance[Char](text.iterator, _ == '\n')
|
|
52 |
|
|
53 |
def advance_codepoints(text: String): Position =
|
64610
|
54 |
if (text.isEmpty) this else advance[Int](Codepoint.iterator(text), _ == 10)
|
64605
|
55 |
}
|
|
56 |
|
|
57 |
|
|
58 |
/* range (right-open interval) */
|
|
59 |
|
|
60 |
sealed case class Range(start: Position, stop: Position)
|
|
61 |
{
|
|
62 |
if (start.compare(stop) > 0)
|
|
63 |
error("Bad line range: " + start.print + ".." + stop.print)
|
|
64 |
|
|
65 |
def print: String = start.print + ".." + stop.print
|
|
66 |
}
|
|
67 |
|
|
68 |
|
|
69 |
/* document with newline as separator (not terminator) */
|
|
70 |
|
|
71 |
object Document
|
|
72 |
{
|
|
73 |
val empty: Document = new Document("", Nil)
|
|
74 |
|
|
75 |
def apply(lines: List[Line]): Document =
|
|
76 |
if (lines.isEmpty) empty
|
|
77 |
else new Document(lines.mkString("", "\n", ""), lines)
|
|
78 |
|
|
79 |
def apply(text: String): Document =
|
|
80 |
if (text.contains('\r'))
|
|
81 |
apply(Library.cat_lines(Library.split_lines(text).map(Library.trim_line(_))))
|
|
82 |
else if (text == "") Document.empty
|
|
83 |
else new Document(text, Library.split_lines(text).map(Line(_)))
|
|
84 |
}
|
|
85 |
|
|
86 |
final class Document private(val text: String, val lines: List[Line])
|
|
87 |
{
|
|
88 |
override def toString: String = text
|
|
89 |
|
|
90 |
override def equals(that: Any): Boolean =
|
|
91 |
that match {
|
|
92 |
case other: Document => lines == other.lines
|
|
93 |
case _ => false
|
|
94 |
}
|
|
95 |
override def hashCode(): Int = lines.hashCode
|
|
96 |
|
|
97 |
private def position(advance: (Position, String) => Position, offset: Text.Offset): Position =
|
|
98 |
{
|
|
99 |
@tailrec def move(i: Text.Offset, lines_count: Int, lines_rest: List[Line]): Position =
|
|
100 |
{
|
|
101 |
lines_rest match {
|
|
102 |
case Nil => require(i == 0); Position(lines_count, 0)
|
|
103 |
case line :: ls =>
|
|
104 |
val n = line.text.length
|
|
105 |
if (ls.isEmpty || i <= n) advance(Position(lines_count, 0), line.text.substring(n - i))
|
|
106 |
else move(i - (n + 1), lines_count + 1, ls)
|
|
107 |
}
|
|
108 |
}
|
|
109 |
move(offset, 0, lines)
|
|
110 |
}
|
|
111 |
|
|
112 |
def position(i: Text.Offset): Position = position(_.advance(_), i)
|
|
113 |
def position_codepoints(i: Text.Offset): Position = position(_.advance_codepoints(_), i)
|
|
114 |
|
64614
|
115 |
// FIXME codepoints
|
64605
|
116 |
def offset(pos: Position): Option[Text.Offset] =
|
|
117 |
{
|
|
118 |
val l = pos.line
|
|
119 |
val c = pos.column
|
|
120 |
if (0 <= l && l < lines.length) {
|
|
121 |
val line_offset =
|
|
122 |
if (l == 0) 0
|
|
123 |
else (0 /: lines.take(l - 1))({ case (n, line) => n + line.text.length + 1 })
|
|
124 |
if (c <= lines(l).text.length) Some(line_offset + c) else None
|
|
125 |
}
|
|
126 |
else None
|
|
127 |
}
|
|
128 |
}
|
|
129 |
|
|
130 |
|
|
131 |
/* line text */
|
|
132 |
|
|
133 |
val empty: Line = new Line("")
|
|
134 |
def apply(text: String): Line = if (text == "") empty else new Line(text)
|
|
135 |
}
|
|
136 |
|
|
137 |
final class Line private(val text: String)
|
|
138 |
{
|
|
139 |
require(text.forall(c => c != '\r' && c != '\n'))
|
|
140 |
|
64610
|
141 |
lazy val length_codepoints: Int = Codepoint.iterator(text).length
|
64605
|
142 |
|
|
143 |
override def equals(that: Any): Boolean =
|
|
144 |
that match {
|
|
145 |
case other: Line => text == other.text
|
|
146 |
case _ => false
|
|
147 |
}
|
|
148 |
override def hashCode(): Int = text.hashCode
|
|
149 |
override def toString: String = text
|
|
150 |
}
|