| author | wenzelm | 
| Sat, 16 Jul 2011 17:11:49 +0200 | |
| changeset 43845 | d89353d17f54 | 
| parent 43696 | 58bb7ca5c7a2 | 
| child 45252 | 6de58d947e57 | 
| permissions | -rw-r--r-- | 
| 31648 | 1 | /* Title: Pure/General/scan.scala | 
| 2 | Author: Makarius | |
| 3 | ||
| 40523 
1050315f6ee2
simplified/robustified treatment of malformed symbols, which are now fully internalized (total Symbol.explode etc.);
 wenzelm parents: 
40290diff
changeset | 4 | Efficient scanning of keywords and tokens. | 
| 31648 | 5 | */ | 
| 6 | ||
| 7 | package isabelle | |
| 8 | ||
| 34187 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 9 | |
| 36011 
3ff725ac13a4
adapted to Scala 2.8.0 Beta1 -- with notable changes to scala.collection;
 wenzelm parents: 
34316diff
changeset | 10 | import scala.collection.generic.Addable | 
| 36679 
ac021aed685e
use IndexedSeq instead of deprecated RandomAccessSeq, which is merely an alias;
 wenzelm parents: 
36012diff
changeset | 11 | import scala.collection.IndexedSeq | 
| 34187 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 12 | import scala.collection.immutable.PagedSeq | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 13 | import scala.util.parsing.input.{OffsetPosition, Position => InputPosition, Reader}
 | 
| 31648 | 14 | import scala.util.parsing.combinator.RegexParsers | 
| 15 | ||
| 34187 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 16 | import java.io.{File, InputStream, BufferedInputStream, FileInputStream}
 | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 17 | |
| 31648 | 18 | |
| 19 | object Scan | |
| 20 | {
 | |
| 43411 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 21 | /** context of partial scans **/ | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 22 | |
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 23 | sealed abstract class Context | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 24 | case object Finished extends Context | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 25 | case class Quoted(quote: String) extends Context | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 26 | case object Verbatim extends Context | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 27 | case class Comment(depth: Int) extends Context | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 28 | |
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 29 | |
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 30 | |
| 31649 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 31 | /** Lexicon -- position tree **/ | 
| 31648 | 32 | |
| 31649 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 33 | object Lexicon | 
| 31648 | 34 |   {
 | 
| 42718 | 35 | protected case class Tree(val branches: Map[Char, (String, Tree)]) | 
| 31649 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 36 | private val empty_tree = Tree(Map()) | 
| 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 37 | |
| 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 38 | val empty: Lexicon = new Lexicon | 
| 31762 
20745ab5b79a
more precise implementation of trait methods -- oddly this seems to require copy/paste for +, ++;
 wenzelm parents: 
31753diff
changeset | 39 | def apply(elems: String*): Lexicon = empty ++ elems | 
| 31648 | 40 | } | 
| 41 | ||
| 36011 
3ff725ac13a4
adapted to Scala 2.8.0 Beta1 -- with notable changes to scala.collection;
 wenzelm parents: 
34316diff
changeset | 42 | class Lexicon extends Addable[String, Lexicon] with RegexParsers | 
| 31648 | 43 |   {
 | 
| 31649 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 44 | /* representation */ | 
| 31648 | 45 | |
| 31649 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 46 | import Lexicon.Tree | 
| 31764 | 47 | protected val main_tree: Tree = Lexicon.empty_tree | 
| 31648 | 48 | |
| 49 | ||
| 31650 | 50 | /* auxiliary operations */ | 
| 51 | ||
| 52 | private def content(tree: Tree, result: List[String]): List[String] = | |
| 53 | (result /: tree.branches.toList) ((res, entry) => | |
| 54 |         entry match { case (_, (s, tr)) =>
 | |
| 55 | if (s.isEmpty) content(tr, res) else content(tr, s :: res) }) | |
| 56 | ||
| 57 | private def lookup(str: CharSequence): Option[(Boolean, Tree)] = | |
| 58 |     {
 | |
| 59 | val len = str.length | |
| 60 | def look(tree: Tree, tip: Boolean, i: Int): Option[(Boolean, Tree)] = | |
| 61 |       {
 | |
| 62 |         if (i < len) {
 | |
| 63 |           tree.branches.get(str.charAt(i)) match {
 | |
| 64 | case Some((s, tr)) => look(tr, !s.isEmpty, i + 1) | |
| 65 | case None => None | |
| 66 | } | |
| 67 | } else Some(tip, tree) | |
| 68 | } | |
| 69 | look(main_tree, false, 0) | |
| 70 | } | |
| 71 | ||
| 72 | def completions(str: CharSequence): List[String] = | |
| 31764 | 73 |       lookup(str) match {
 | 
| 31650 | 74 | case Some((true, tree)) => content(tree, List(str.toString)) | 
| 75 | case Some((false, tree)) => content(tree, Nil) | |
| 76 | case None => Nil | |
| 31764 | 77 | } | 
| 31650 | 78 | |
| 79 | ||
| 36011 
3ff725ac13a4
adapted to Scala 2.8.0 Beta1 -- with notable changes to scala.collection;
 wenzelm parents: 
34316diff
changeset | 80 | /* pseudo Set methods */ | 
| 31649 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 81 | |
| 36012 | 82 | def iterator: Iterator[String] = content(main_tree, Nil).sortWith(_ <= _).iterator | 
| 31649 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 83 | |
| 36011 
3ff725ac13a4
adapted to Scala 2.8.0 Beta1 -- with notable changes to scala.collection;
 wenzelm parents: 
34316diff
changeset | 84 |     override def toString: String = iterator.mkString("Lexicon(", ", ", ")")
 | 
| 31649 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 85 | |
| 36011 
3ff725ac13a4
adapted to Scala 2.8.0 Beta1 -- with notable changes to scala.collection;
 wenzelm parents: 
34316diff
changeset | 86 | def empty: Lexicon = Lexicon.empty | 
| 
3ff725ac13a4
adapted to Scala 2.8.0 Beta1 -- with notable changes to scala.collection;
 wenzelm parents: 
34316diff
changeset | 87 | def isEmpty: Boolean = main_tree.branches.isEmpty | 
| 31648 | 88 | |
| 31762 
20745ab5b79a
more precise implementation of trait methods -- oddly this seems to require copy/paste for +, ++;
 wenzelm parents: 
31753diff
changeset | 89 | def contains(elem: String): Boolean = | 
| 
20745ab5b79a
more precise implementation of trait methods -- oddly this seems to require copy/paste for +, ++;
 wenzelm parents: 
31753diff
changeset | 90 |       lookup(elem) match {
 | 
| 31650 | 91 | case Some((tip, _)) => tip | 
| 92 | case _ => false | |
| 93 | } | |
| 31648 | 94 | |
| 36011 
3ff725ac13a4
adapted to Scala 2.8.0 Beta1 -- with notable changes to scala.collection;
 wenzelm parents: 
34316diff
changeset | 95 | |
| 
3ff725ac13a4
adapted to Scala 2.8.0 Beta1 -- with notable changes to scala.collection;
 wenzelm parents: 
34316diff
changeset | 96 | /* Addable methods */ | 
| 
3ff725ac13a4
adapted to Scala 2.8.0 Beta1 -- with notable changes to scala.collection;
 wenzelm parents: 
34316diff
changeset | 97 | |
| 
3ff725ac13a4
adapted to Scala 2.8.0 Beta1 -- with notable changes to scala.collection;
 wenzelm parents: 
34316diff
changeset | 98 | def repr = this | 
| 
3ff725ac13a4
adapted to Scala 2.8.0 Beta1 -- with notable changes to scala.collection;
 wenzelm parents: 
34316diff
changeset | 99 | |
| 31762 
20745ab5b79a
more precise implementation of trait methods -- oddly this seems to require copy/paste for +, ++;
 wenzelm parents: 
31753diff
changeset | 100 | def + (elem: String): Lexicon = | 
| 
20745ab5b79a
more precise implementation of trait methods -- oddly this seems to require copy/paste for +, ++;
 wenzelm parents: 
31753diff
changeset | 101 | if (contains(elem)) this | 
| 
20745ab5b79a
more precise implementation of trait methods -- oddly this seems to require copy/paste for +, ++;
 wenzelm parents: 
31753diff
changeset | 102 |       else {
 | 
| 
20745ab5b79a
more precise implementation of trait methods -- oddly this seems to require copy/paste for +, ++;
 wenzelm parents: 
31753diff
changeset | 103 | val len = elem.length | 
| 
20745ab5b79a
more precise implementation of trait methods -- oddly this seems to require copy/paste for +, ++;
 wenzelm parents: 
31753diff
changeset | 104 | def extend(tree: Tree, i: Int): Tree = | 
| 
20745ab5b79a
more precise implementation of trait methods -- oddly this seems to require copy/paste for +, ++;
 wenzelm parents: 
31753diff
changeset | 105 |           if (i < len) {
 | 
| 
20745ab5b79a
more precise implementation of trait methods -- oddly this seems to require copy/paste for +, ++;
 wenzelm parents: 
31753diff
changeset | 106 | val c = elem.charAt(i) | 
| 
20745ab5b79a
more precise implementation of trait methods -- oddly this seems to require copy/paste for +, ++;
 wenzelm parents: 
31753diff
changeset | 107 | val end = (i + 1 == len) | 
| 
20745ab5b79a
more precise implementation of trait methods -- oddly this seems to require copy/paste for +, ++;
 wenzelm parents: 
31753diff
changeset | 108 |             tree.branches.get(c) match {
 | 
| 
20745ab5b79a
more precise implementation of trait methods -- oddly this seems to require copy/paste for +, ++;
 wenzelm parents: 
31753diff
changeset | 109 | case Some((s, tr)) => | 
| 
20745ab5b79a
more precise implementation of trait methods -- oddly this seems to require copy/paste for +, ++;
 wenzelm parents: 
31753diff
changeset | 110 | Tree(tree.branches + | 
| 
20745ab5b79a
more precise implementation of trait methods -- oddly this seems to require copy/paste for +, ++;
 wenzelm parents: 
31753diff
changeset | 111 | (c -> (if (end) elem else s, extend(tr, i + 1)))) | 
| 
20745ab5b79a
more precise implementation of trait methods -- oddly this seems to require copy/paste for +, ++;
 wenzelm parents: 
31753diff
changeset | 112 | case None => | 
| 
20745ab5b79a
more precise implementation of trait methods -- oddly this seems to require copy/paste for +, ++;
 wenzelm parents: 
31753diff
changeset | 113 | Tree(tree.branches + | 
| 
20745ab5b79a
more precise implementation of trait methods -- oddly this seems to require copy/paste for +, ++;
 wenzelm parents: 
31753diff
changeset | 114 | (c -> (if (end) elem else "", extend(Lexicon.empty_tree, i + 1)))) | 
| 
20745ab5b79a
more precise implementation of trait methods -- oddly this seems to require copy/paste for +, ++;
 wenzelm parents: 
31753diff
changeset | 115 | } | 
| 34135 | 116 | } | 
| 117 | else tree | |
| 31762 
20745ab5b79a
more precise implementation of trait methods -- oddly this seems to require copy/paste for +, ++;
 wenzelm parents: 
31753diff
changeset | 118 | val old = this | 
| 
20745ab5b79a
more precise implementation of trait methods -- oddly this seems to require copy/paste for +, ++;
 wenzelm parents: 
31753diff
changeset | 119 |         new Lexicon { override val main_tree = extend(old.main_tree, 0) }
 | 
| 31648 | 120 | } | 
| 31762 
20745ab5b79a
more precise implementation of trait methods -- oddly this seems to require copy/paste for +, ++;
 wenzelm parents: 
31753diff
changeset | 121 | |
| 31649 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 122 | |
| 34301 | 123 | |
| 34137 | 124 | /** RegexParsers methods **/ | 
| 31649 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 125 | |
| 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 126 | override val whiteSpace = "".r | 
| 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 127 | |
| 34137 | 128 | |
| 43411 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 129 | /* optional termination */ | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 130 | |
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 131 | def opt_term[T](p: => Parser[T]): Parser[Option[T]] = | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 132 | p ^^ (x => Some(x)) | """\z""".r ^^ (_ => None) | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 133 | |
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 134 | |
| 34137 | 135 | /* keywords from lexicon */ | 
| 136 | ||
| 34140 | 137 | def keyword: Parser[String] = new Parser[String] | 
| 34135 | 138 |     {
 | 
| 31649 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 139 | def apply(in: Input) = | 
| 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 140 |       {
 | 
| 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 141 | val source = in.source | 
| 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 142 | val offset = in.offset | 
| 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 143 | val len = source.length - offset | 
| 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 144 | |
| 34140 | 145 | def scan(tree: Tree, result: String, i: Int): String = | 
| 31649 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 146 |         {
 | 
| 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 147 |           if (i < len) {
 | 
| 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 148 |             tree.branches.get(source.charAt(offset + i)) match {
 | 
| 34140 | 149 | case Some((s, tr)) => scan(tr, if (s.isEmpty) result else s, i + 1) | 
| 34135 | 150 | case None => result | 
| 31649 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 151 | } | 
| 34135 | 152 | } | 
| 153 | else result | |
| 31649 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 154 | } | 
| 34140 | 155 | val result = scan(main_tree, "", 0) | 
| 156 |         if (result.isEmpty) Failure("keyword expected", in)
 | |
| 157 | else Success(result, in.drop(result.length)) | |
| 31648 | 158 | } | 
| 31649 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 159 |     }.named("keyword")
 | 
| 31648 | 160 | |
| 34137 | 161 | |
| 34157 
0a0a19153626
explicit representation of Token_Kind -- cannot really depend on runtime types due to erasure;
 wenzelm parents: 
34144diff
changeset | 162 | /* symbol range */ | 
| 34137 | 163 | |
| 43696 | 164 | def symbol_range(pred: Symbol.Symbol => Boolean, min_count: Int, max_count: Int): Parser[String] = | 
| 34140 | 165 | new Parser[String] | 
| 34135 | 166 |       {
 | 
| 34137 | 167 | def apply(in: Input) = | 
| 168 |         {
 | |
| 169 | val start = in.offset | |
| 170 | val end = in.source.length | |
| 171 | val matcher = new Symbol.Matcher(in.source) | |
| 172 | ||
| 173 | var i = start | |
| 174 | var count = 0 | |
| 175 | var finished = false | |
| 176 |           while (!finished) {
 | |
| 177 |             if (i < end && count < max_count) {
 | |
| 178 | val n = matcher(i, end) | |
| 179 | val sym = in.source.subSequence(i, i + n).toString | |
| 180 |               if (pred(sym)) { i += n; count += 1 }
 | |
| 181 | else finished = true | |
| 182 | } | |
| 183 | else finished = true | |
| 184 | } | |
| 34140 | 185 |           if (count < min_count) Failure("bad input", in)
 | 
| 186 | else Success(in.source.subSequence(start, i).toString, in.drop(i - start)) | |
| 34137 | 187 | } | 
| 34157 
0a0a19153626
explicit representation of Token_Kind -- cannot really depend on runtime types due to erasure;
 wenzelm parents: 
34144diff
changeset | 188 |       }.named("symbol_range")
 | 
| 34137 | 189 | |
| 43696 | 190 | def one(pred: Symbol.Symbol => Boolean): Parser[String] = | 
| 191 | symbol_range(pred, 1, 1) | |
| 192 | ||
| 193 | def many(pred: Symbol.Symbol => Boolean): Parser[String] = | |
| 194 | symbol_range(pred, 0, Integer.MAX_VALUE) | |
| 195 | ||
| 196 | def many1(pred: Symbol.Symbol => Boolean): Parser[String] = | |
| 197 | symbol_range(pred, 1, Integer.MAX_VALUE) | |
| 34140 | 198 | |
| 199 | ||
| 34143 | 200 | /* quoted strings */ | 
| 34140 | 201 | |
| 43696 | 202 | private def quoted_body(quote: Symbol.Symbol): Parser[String] = | 
| 43411 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 203 |     {
 | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 204 | rep(many1(sym => sym != quote && sym != "\\") | "\\" + quote | "\\\\" | | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 205 |         (("""\\\d\d\d""".r) ^? { case x if x.substring(1, 4).toInt <= 255 => x })) ^^ (_.mkString)
 | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 206 | } | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 207 | |
| 43696 | 208 | private def quoted(quote: Symbol.Symbol): Parser[String] = | 
| 34143 | 209 |     {
 | 
| 43411 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 210 |       quote ~ quoted_body(quote) ~ quote ^^ { case x ~ y ~ z => x + y + z }
 | 
| 34143 | 211 |     }.named("quoted")
 | 
| 34140 | 212 | |
| 43696 | 213 | def quoted_content(quote: Symbol.Symbol, source: String): String = | 
| 34140 | 214 |     {
 | 
| 215 | require(parseAll(quoted(quote), source).successful) | |
| 34189 | 216 | val body = source.substring(1, source.length - 1) | 
| 217 |       if (body.exists(_ == '\\')) {
 | |
| 218 | val content = | |
| 40523 
1050315f6ee2
simplified/robustified treatment of malformed symbols, which are now fully internalized (total Symbol.explode etc.);
 wenzelm parents: 
40290diff
changeset | 219 | rep(many1(sym => sym != quote && sym != "\\") | | 
| 34189 | 220 |               "\\" ~> (quote | "\\" | """\d\d\d""".r ^^ { case x => x.toInt.toChar.toString }))
 | 
| 221 | parseAll(content ^^ (_.mkString), body).get | |
| 222 | } | |
| 223 | else body | |
| 34140 | 224 | } | 
| 225 | ||
| 43696 | 226 | def quoted_context(quote: Symbol.Symbol, ctxt: Context): Parser[(String, Context)] = | 
| 43411 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 227 |     {
 | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 228 |       ctxt match {
 | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 229 | case Finished => | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 230 | quote ~ quoted_body(quote) ~ opt_term(quote) ^^ | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 231 |             { case x ~ y ~ Some(z) => (x + y + z, Finished)
 | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 232 | case x ~ y ~ None => (x + y, Quoted(quote)) } | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 233 | case Quoted(q) if q == quote => | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 234 | quoted_body(quote) ~ opt_term(quote) ^^ | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 235 |             { case x ~ Some(y) => (x + y, Finished)
 | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 236 | case x ~ None => (x, ctxt) } | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 237 |         case _ => failure("")
 | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 238 | } | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 239 |     }.named("quoted_context")
 | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 240 | |
| 34140 | 241 | |
| 34143 | 242 | /* verbatim text */ | 
| 243 | ||
| 43411 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 244 | private def verbatim_body: Parser[String] = | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 245 | rep(many1(sym => sym != "*") | """\*(?!\})""".r) ^^ (_.mkString) | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 246 | |
| 34143 | 247 | private def verbatim: Parser[String] = | 
| 248 |     {
 | |
| 43411 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 249 |       "{*" ~ verbatim_body ~ "*}" ^^ { case x ~ y ~ z => x + y + z }
 | 
| 34143 | 250 |     }.named("verbatim")
 | 
| 34140 | 251 | |
| 252 | def verbatim_content(source: String): String = | |
| 253 |     {
 | |
| 254 | require(parseAll(verbatim, source).successful) | |
| 255 | source.substring(2, source.length - 2) | |
| 256 | } | |
| 257 | ||
| 43411 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 258 | def verbatim_context(ctxt: Context): Parser[(String, Context)] = | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 259 |     {
 | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 260 |       ctxt match {
 | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 261 | case Finished => | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 262 |           "{*" ~ verbatim_body ~ opt_term("*}") ^^
 | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 263 |             { case x ~ y ~ Some(z) => (x + y + z, Finished)
 | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 264 | case x ~ y ~ None => (x + y, Verbatim) } | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 265 | case Verbatim => | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 266 |           verbatim_body ~ opt_term("*}") ^^
 | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 267 |             { case x ~ Some(y) => (x + y, Finished)
 | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 268 | case x ~ None => (x, Verbatim) } | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 269 |         case _ => failure("")
 | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 270 | } | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 271 |     }.named("verbatim_context")
 | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 272 | |
| 34140 | 273 | |
| 34143 | 274 | /* nested comments */ | 
| 275 | ||
| 43412 | 276 | private def comment_depth(depth: Int): Parser[(String, Int)] = new Parser[(String, Int)] | 
| 34143 | 277 |     {
 | 
| 43412 | 278 | require(depth >= 0) | 
| 279 | ||
| 34143 | 280 | val comment_text = | 
| 42441 
781c622af16a
more robust scanning of iterated comments, such as "(* (**) (**) *)";
 wenzelm parents: 
40523diff
changeset | 281 |         rep1(many1(sym => sym != "*" && sym != "(") | """\*(?!\))|\((?!\*)""".r)
 | 
| 34143 | 282 | |
| 283 | def apply(in: Input) = | |
| 284 |       {
 | |
| 285 | var rest = in | |
| 42441 
781c622af16a
more robust scanning of iterated comments, such as "(* (**) (**) *)";
 wenzelm parents: 
40523diff
changeset | 286 | def try_parse[A](p: Parser[A]): Boolean = | 
| 34143 | 287 |         {
 | 
| 42441 
781c622af16a
more robust scanning of iterated comments, such as "(* (**) (**) *)";
 wenzelm parents: 
40523diff
changeset | 288 |           parse(p ^^^ (), rest) match {
 | 
| 34143 | 289 |             case Success(_, next) => { rest = next; true }
 | 
| 290 | case _ => false | |
| 291 | } | |
| 292 | } | |
| 43412 | 293 | var d = depth | 
| 34143 | 294 | var finished = false | 
| 295 |         while (!finished) {
 | |
| 43412 | 296 |           if (try_parse("(*")) d += 1
 | 
| 297 |           else if (d > 0 && try_parse("*)")) d -= 1
 | |
| 298 | else if (d == 0 || !try_parse(comment_text)) finished = true | |
| 34143 | 299 | } | 
| 43412 | 300 | if (in.offset < rest.offset) | 
| 301 | Success((in.source.subSequence(in.offset, rest.offset).toString, d), rest) | |
| 34143 | 302 |         else Failure("comment expected", in)
 | 
| 303 | } | |
| 43412 | 304 |     }.named("comment_depth")
 | 
| 305 | ||
| 306 | def comment: Parser[String] = | |
| 307 |       comment_depth(0) ^? { case (x, d) if d == 0 => x }
 | |
| 308 | ||
| 309 | def comment_context(ctxt: Context): Parser[(String, Context)] = | |
| 310 |     {
 | |
| 311 | val depth = | |
| 312 |         ctxt match {
 | |
| 313 | case Finished => 0 | |
| 314 | case Comment(d) => d | |
| 315 | case _ => -1 | |
| 316 | } | |
| 317 | if (depth >= 0) | |
| 318 | comment_depth(depth) ^^ | |
| 319 |           { case (x, 0) => (x, Finished)
 | |
| 320 | case (x, d) => (x, Comment(d)) } | |
| 321 |       else failure("")
 | |
| 322 | } | |
| 34143 | 323 | |
| 324 | def comment_content(source: String): String = | |
| 325 |     {
 | |
| 326 | require(parseAll(comment, source).successful) | |
| 327 | source.substring(2, source.length - 2) | |
| 328 | } | |
| 329 | ||
| 330 | ||
| 34140 | 331 | /* outer syntax tokens */ | 
| 332 | ||
| 43411 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 333 | private def delimited_token: Parser[Token] = | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 334 |     {
 | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 335 |       val string = quoted("\"") ^^ (x => Token(Token.Kind.STRING, x))
 | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 336 |       val alt_string = quoted("`") ^^ (x => Token(Token.Kind.ALT_STRING, x))
 | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 337 | val verb = verbatim ^^ (x => Token(Token.Kind.VERBATIM, x)) | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 338 | val cmt = comment ^^ (x => Token(Token.Kind.COMMENT, x)) | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 339 | |
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 340 | string | (alt_string | (verb | cmt)) | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 341 | } | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 342 | |
| 43695 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43420diff
changeset | 343 | private def other_token(is_command: String => Boolean) | 
| 43411 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 344 | : Parser[Token] = | 
| 34140 | 345 |     {
 | 
| 43695 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43420diff
changeset | 346 |       val id = one(Symbol.is_letter) ~ many(Symbol.is_letdig) ^^ { case x ~ y => x + y }
 | 
| 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43420diff
changeset | 347 | val nat = many1(Symbol.is_digit) | 
| 40290 
47f572aff50a
support for floating-point tokens in outer syntax (coinciding with inner syntax version);
 wenzelm parents: 
38367diff
changeset | 348 |       val natdot = nat ~ "." ~ nat ^^ { case x ~ y ~ z => x + y + z }
 | 
| 34140 | 349 |       val id_nat = id ~ opt("." ~ nat) ^^ { case x ~ Some(y ~ z) => x + y + z case x ~ None => x }
 | 
| 350 | ||
| 351 |       val ident = id ~ rep("." ~> id) ^^
 | |
| 36956 
21be4832c362
renamed class Outer_Lex to Token and Token_Kind to Token.Kind;
 wenzelm parents: 
36679diff
changeset | 352 |         { case x ~ Nil => Token(Token.Kind.IDENT, x)
 | 
| 
21be4832c362
renamed class Outer_Lex to Token and Token_Kind to Token.Kind;
 wenzelm parents: 
36679diff
changeset | 353 |           case x ~ ys => Token(Token.Kind.LONG_IDENT, (x :: ys).mkString(".")) }
 | 
| 34140 | 354 | |
| 36956 
21be4832c362
renamed class Outer_Lex to Token and Token_Kind to Token.Kind;
 wenzelm parents: 
36679diff
changeset | 355 |       val var_ = "?" ~ id_nat ^^ { case x ~ y => Token(Token.Kind.VAR, x + y) }
 | 
| 
21be4832c362
renamed class Outer_Lex to Token and Token_Kind to Token.Kind;
 wenzelm parents: 
36679diff
changeset | 356 |       val type_ident = "'" ~ id ^^ { case x ~ y => Token(Token.Kind.TYPE_IDENT, x + y) }
 | 
| 
21be4832c362
renamed class Outer_Lex to Token and Token_Kind to Token.Kind;
 wenzelm parents: 
36679diff
changeset | 357 |       val type_var = "?'" ~ id_nat ^^ { case x ~ y => Token(Token.Kind.TYPE_VAR, x + y) }
 | 
| 
21be4832c362
renamed class Outer_Lex to Token and Token_Kind to Token.Kind;
 wenzelm parents: 
36679diff
changeset | 358 | val nat_ = nat ^^ (x => Token(Token.Kind.NAT, x)) | 
| 40290 
47f572aff50a
support for floating-point tokens in outer syntax (coinciding with inner syntax version);
 wenzelm parents: 
38367diff
changeset | 359 | val float = | 
| 
47f572aff50a
support for floating-point tokens in outer syntax (coinciding with inner syntax version);
 wenzelm parents: 
38367diff
changeset | 360 |         ("-" ~ natdot ^^ { case x ~ y => x + y } | natdot) ^^ (x => Token(Token.Kind.FLOAT, x))
 | 
| 34140 | 361 | |
| 362 | val sym_ident = | |
| 43695 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43420diff
changeset | 363 | (many1(Symbol.is_symbolic_char) | one(sym => Symbol.is_symbolic(sym))) ^^ | 
| 36956 
21be4832c362
renamed class Outer_Lex to Token and Token_Kind to Token.Kind;
 wenzelm parents: 
36679diff
changeset | 364 | (x => Token(Token.Kind.SYM_IDENT, x)) | 
| 34140 | 365 | |
| 43695 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43420diff
changeset | 366 | val space = many1(Symbol.is_blank) ^^ (x => Token(Token.Kind.SPACE, x)) | 
| 34140 | 367 | |
| 43411 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 368 | // FIXME check | 
| 43695 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43420diff
changeset | 369 | val junk = many(sym => !(Symbol.is_blank(sym))) | 
| 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43420diff
changeset | 370 | val junk1 = many1(sym => !(Symbol.is_blank(sym))) | 
| 34140 | 371 | |
| 34267 | 372 | val bad_delimiter = | 
| 38367 | 373 |         ("\"" | "`" | "{*" | "(*") ~ junk ^^ { case x ~ y => Token(Token.Kind.UNPARSED, x + y) }
 | 
| 43411 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 374 | val bad = junk1 ^^ (x => Token(Token.Kind.UNPARSED, x)) | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 375 | |
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 376 | val command_keyword = | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 377 | keyword ^^ (x => Token(if (is_command(x)) Token.Kind.COMMAND else Token.Kind.KEYWORD, x)) | 
| 34140 | 378 | |
| 43411 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 379 | space | (bad_delimiter | | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 380 | (((ident | (var_ | (type_ident | (type_var | (float | (nat_ | sym_ident)))))) ||| | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 381 | command_keyword) | bad)) | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 382 | } | 
| 34140 | 383 | |
| 43695 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43420diff
changeset | 384 | def token(is_command: String => Boolean): Parser[Token] = | 
| 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43420diff
changeset | 385 | delimited_token | other_token(is_command) | 
| 34140 | 386 | |
| 43695 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43420diff
changeset | 387 | def token_context(is_command: String => Boolean, ctxt: Context): Parser[(Token, Context)] = | 
| 43411 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 388 |     {
 | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 389 | val string = | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 390 |         quoted_context("\"", ctxt) ^^ { case (x, c) => (Token(Token.Kind.STRING, x), c) }
 | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 391 | val alt_string = | 
| 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 392 |         quoted_context("`", ctxt) ^^ { case (x, c) => (Token(Token.Kind.ALT_STRING, x), c) }
 | 
| 43412 | 393 |       val verb = verbatim_context(ctxt) ^^ { case (x, c) => (Token(Token.Kind.VERBATIM, x), c) }
 | 
| 394 |       val cmt = comment_context(ctxt) ^^ { case (x, c) => (Token(Token.Kind.COMMENT, x), c) }
 | |
| 43695 
5130dfe1b7be
simplified Symbol based on lazy Symbol.Interpretation -- reduced odd "functorial style";
 wenzelm parents: 
43420diff
changeset | 395 |       val other = other_token(is_command) ^^ { case x => (x, Finished) }
 | 
| 43411 
0206466ee473
some support for partial scans with explicit context;
 wenzelm parents: 
42718diff
changeset | 396 | |
| 43412 | 397 | string | (alt_string | (verb | (cmt | other))) | 
| 34140 | 398 | } | 
| 31649 
a11ea667d676
reorganized and abstracted version, via Set trait;
 wenzelm parents: 
31648diff
changeset | 399 | } | 
| 34187 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 400 | |
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 401 | |
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 402 | |
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 403 | /** read file without decoding -- enables efficient length operation **/ | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 404 | |
| 36679 
ac021aed685e
use IndexedSeq instead of deprecated RandomAccessSeq, which is merely an alias;
 wenzelm parents: 
36012diff
changeset | 405 | private class Restricted_Seq(seq: IndexedSeq[Char], start: Int, end: Int) | 
| 34187 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 406 | extends CharSequence | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 407 |   {
 | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 408 | def charAt(i: Int): Char = | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 409 | if (0 <= i && i < length) seq(start + i) | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 410 | else throw new IndexOutOfBoundsException | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 411 | |
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 412 | def length: Int = end - start // avoid potentially expensive seq.length | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 413 | |
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 414 | def subSequence(i: Int, j: Int): CharSequence = | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 415 | if (0 <= i && i <= j && j <= length) new Restricted_Seq(seq, start + i, start + j) | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 416 | else throw new IndexOutOfBoundsException | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 417 | |
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 418 | override def toString: String = | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 419 |     {
 | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 420 | val buf = new StringBuilder(length) | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 421 | for (offset <- start until end) buf.append(seq(offset)) | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 422 | buf.toString | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 423 | } | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 424 | } | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 425 | |
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 426 |   abstract class Byte_Reader extends Reader[Char] { def close: Unit }
 | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 427 | |
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 428 | def byte_reader(file: File): Byte_Reader = | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 429 |   {
 | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 430 | val stream = new BufferedInputStream(new FileInputStream(file)) | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 431 | val seq = new PagedSeq( | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 432 | (buf: Array[Char], offset: Int, length: Int) => | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 433 |         {
 | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 434 | var i = 0 | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 435 | var c = 0 | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 436 | var eof = false | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 437 |           while (!eof && i < length) {
 | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 438 | c = stream.read | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 439 | if (c == -1) eof = true | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 440 |             else { buf(offset + i) = c.toChar; i += 1 }
 | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 441 | } | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 442 | if (i > 0) i else -1 | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 443 | }) | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 444 | val restricted_seq = new Restricted_Seq(seq, 0, file.length.toInt) | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 445 | |
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 446 | class Paged_Reader(override val offset: Int) extends Byte_Reader | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 447 |     {
 | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 448 | override lazy val source: CharSequence = restricted_seq | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 449 | def first: Char = if (seq.isDefinedAt(offset)) seq(offset) else '\032' | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 450 | def rest: Paged_Reader = if (seq.isDefinedAt(offset)) new Paged_Reader(offset + 1) else this | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 451 | def pos: InputPosition = new OffsetPosition(source, offset) | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 452 | def atEnd: Boolean = !seq.isDefinedAt(offset) | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 453 | override def drop(n: Int): Paged_Reader = new Paged_Reader(offset + n) | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 454 |       def close { stream.close }
 | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 455 | } | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 456 | new Paged_Reader(0) | 
| 
7b659c1561f1
added byte_reader, which works without decoding and enables efficient length operation (for scala.util.parsing.input.Reader);
 wenzelm parents: 
34157diff
changeset | 457 | } | 
| 31648 | 458 | } |