# HG changeset patch # User Christian Urban # Date 1613964146 0 # Node ID e752d84225ec7b818ac3f07d57b785ebcc4ff617 # Parent fedc16924b76ecd5a4b4c1b6735bfd592dfe6362 updated diff -r fedc16924b76 -r e752d84225ec progs/scala/re-annotated.sc --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/progs/scala/re-annotated.sc Mon Feb 22 03:22:26 2021 +0000 @@ -0,0 +1,407 @@ +// uses bitcode sequences and annotated regular +// expressions +// +// for basic regular expressions and RECD +// +// uses a non-standard extraction method for generating +// tokens (this is tail-recursive) +// +// can match 60 copies of the fib-program (size 10500) +// in about 20 secs + + +import scala.language.implicitConversions +import scala.language.reflectiveCalls +import scala.annotation.tailrec + +// standard regular expressions +abstract class Rexp +case object ZERO extends Rexp +case object ONE extends Rexp +case class CHAR(c: Char) extends Rexp +case class ALT(r1: Rexp, r2: Rexp) extends Rexp +case class SEQ(r1: Rexp, r2: Rexp) extends Rexp +case class STAR(r: Rexp) extends Rexp +case class RECD(x: String, r: Rexp) extends Rexp + +abstract class Bit +case object Z extends Bit +case object S extends Bit + +type Bits = List[Bit] + +// annotated regular expressions +abstract class ARexp +case object AZERO extends ARexp +case class AONE(bs: Bits) extends ARexp +case class ACHAR(bs: Bits, c: Char) extends ARexp +case class AALTS(bs: Bits, rs: List[ARexp]) extends ARexp +case class ASEQ(bs: Bits, r1: ARexp, r2: ARexp) extends ARexp +case class ASTAR(bs: Bits, r: ARexp) extends ARexp + +// an abbreviation for binary alternatives +def AALT(bs: Bits, r1: ARexp, r2: ARexp) = AALTS(bs, List(r1, r2)) + +abstract class Val +case object Empty extends Val +case class Chr(c: Char) extends Val +case class Sequ(v1: Val, v2: Val) extends Val +case class Left(v: Val) extends Val +case class Right(v: Val) extends Val +case class Stars(vs: List[Val]) extends Val +case class Recd(x: String, v: Val) extends Val + +// some convenience for typing in regular expressions +def charlist2rexp(s: List[Char]): Rexp = s match { + case Nil => ONE + case c::Nil => CHAR(c) + case c::s => SEQ(CHAR(c), charlist2rexp(s)) +} +implicit def string2rexp(s: String) : Rexp = charlist2rexp(s.toList) + +implicit def RexpOps(r: Rexp) = new { + def | (s: Rexp) = ALT(r, s) + def % = STAR(r) + def ~ (s: Rexp) = SEQ(r, s) +} + +implicit def stringOps(s: String) = new { + def | (r: Rexp) = ALT(s, r) + def | (r: String) = ALT(s, r) + def % = STAR(s) + def ~ (r: Rexp) = SEQ(s, r) + def ~ (r: String) = SEQ(s, r) + def $ (r: Rexp) = RECD(s, r) +} + +def size(r: Rexp) : Int = r match { + case ZERO => 1 + case ONE => 1 + case CHAR(_) => 1 + case ALT(r1, r2) => 1 + size(r1) + size(r2) + case SEQ(r1, r2) => 1 + size(r1) + size(r2) + case STAR(r) => 1 + size(r) + case RECD(_, r) => 1 + size(r) +} + + + +// Bitcoded + Annotation +//======================= + +//erase function: extracts the Rexp from ARexp +def erase(r:ARexp): Rexp = r match{ + case AZERO => ZERO + case AONE(_) => ONE + case ACHAR(bs, c) => CHAR(c) + case AALTS(bs, Nil) => ZERO + case AALTS(bs, r::Nil) => erase(r) + case AALTS(bs, r::rs) => ALT(erase(r), erase(AALTS(bs, rs))) + case ASEQ(bs, r1, r2) => SEQ (erase(r1), erase(r2)) + case ASTAR(cs, r)=> STAR(erase(r)) +} + +def fuse(bs: Bits, r: ARexp) : ARexp = r match { + case AZERO => AZERO + case AONE(cs) => AONE(bs ++ cs) + case ACHAR(cs, c) => ACHAR(bs ++ cs, c) + case AALTS(cs, rs) => AALTS(bs ++ cs, rs) + case ASEQ(cs, r1, r2) => ASEQ(bs ++ cs, r1, r2) + case ASTAR(cs, r) => ASTAR(bs ++ cs, r) +} + +def internalise(r: Rexp) : ARexp = r match { + case ZERO => AZERO + case ONE => AONE(Nil) + case CHAR(c) => ACHAR(Nil, c) + case ALT(r1, r2) => AALT(Nil, fuse(List(Z), internalise(r1)), fuse(List(S), internalise(r2))) + case SEQ(r1, r2) => ASEQ(Nil, internalise(r1), internalise(r2)) + case STAR(r) => ASTAR(Nil, internalise(r)) + case RECD(_, r) => internalise(r) +} + +// example +// internalise(("a" | "ab") ~ ("b" | "")) + + +// decoding of a value from a bitsequence +// (this is not tail-recursive and therefore a potential bottleneck) +def vdecode_aux(r: Rexp, bs: Bits) : (Val, Bits) = (r, bs) match { + case (ONE, bs) => (Empty, bs) + case (CHAR(c), bs) => (Chr(c), bs) + case (ALT(r1, r2), Z::bs) => { + val (v, bs1) = vdecode_aux(r1, bs) + (Left(v), bs1) + } + case (ALT(r1, r2), S::bs) => { + val (v, bs1) = vdecode_aux(r2, bs) + (Right(v), bs1) + } + case (SEQ(r1, r2), bs) => { + val (v1, bs1) = vdecode_aux(r1, bs) + val (v2, bs2) = vdecode_aux(r2, bs1) + (Sequ(v1, v2), bs2) + } + case (STAR(r1), Z::bs) => { + val (v, bs1) = vdecode_aux(r1, bs) + val (Stars(vs), bs2) = vdecode_aux(STAR(r1), bs1) + (Stars(v::vs), bs2) + } + case (STAR(_), S::bs) => (Stars(Nil), bs) + case (RECD(s, r1), bs) => + val (v, bs1) = vdecode_aux(r1, bs) + (Recd(s, v), bs1) +} + +def vdecode(r: Rexp, bs: Bits) = vdecode_aux(r, bs) match { + case (v, Nil) => v + case _ => throw new Exception("Not decodable") +} + +// decoding of sequence of string tokens from a bitsequence +// tail-recursive version using an accumulator (alternative for +// vdecode) +@tailrec +def sdecode_aux(rs: List[Rexp], bs: Bits, acc: List[String]) : List[String] = (rs, bs) match { + case (Nil, _) => acc + case (_, Nil) => acc + case (ONE::rest, bs) => sdecode_aux(rest, bs, acc) + case (CHAR(c)::rest, bs) => + sdecode_aux(rest, bs, (acc.head ++ c.toString)::acc.tail) + case (ALT(r1, r2)::rest, Z::bs) => sdecode_aux(r1::rest, bs, acc) + case (ALT(r1, r2)::rest, S::bs) => sdecode_aux(r2::rest, bs, acc) + case (SEQ(r1, r2)::rest, bs) => sdecode_aux(r1::r2::rest, bs, acc) + case (STAR(r1)::rest, Z::bs) => sdecode_aux(r1::STAR(r1)::rest, bs, acc) + case (STAR(_)::rest, S::bs) => sdecode_aux(rest, bs, acc) + case (RECD(s, r1)::rest, bs) => + sdecode_aux(r1::rest, bs, s"$s:"::acc) +} + +def sdecode(r: Rexp, bs: Bits) : List[String] = + sdecode_aux(List(r), bs, List("")).reverse.tail + + + +// nullable function: tests whether the an (annotated) +// regular expression can recognise the empty string +def bnullable (r: ARexp) : Boolean = r match { + case AZERO => false + case AONE(_) => true + case ACHAR(_,_) => false + case AALTS(_, rs) => rs.exists(bnullable) + case ASEQ(_, r1, r2) => bnullable(r1) && bnullable(r2) + case ASTAR(_, _) => true +} + +def bmkeps(r: ARexp) : Bits = r match { + case AONE(bs) => bs + case AALTS(bs, r::Nil) => bs ++ bmkeps(r) + case AALTS(bs, r::rs) => + if (bnullable(r)) bs ++ bmkeps(r) else bmkeps(AALTS(bs, rs)) + case ASEQ(bs, r1, r2) => bs ++ bmkeps(r1) ++ bmkeps(r2) + case ASTAR(bs, r) => bs ++ List(S) +} + +// derivative of a regular expression w.r.t. a character +def bder(c: Char, r: ARexp) : ARexp = r match { + case AZERO => AZERO + case AONE(_) => AZERO + case ACHAR(bs, d) => if (c == d) AONE(bs) else AZERO + case AALTS(bs, rs) => AALTS(bs, rs.map(bder(c, _))) + case ASEQ(bs, r1, r2) => + if (bnullable(r1)) AALT(bs, ASEQ(Nil, bder(c, r1), r2), fuse(bmkeps(r1), bder(c, r2))) + else ASEQ(bs, bder(c, r1), r2) + case ASTAR(bs, r) => ASEQ(bs, fuse(List(Z), bder(c, r)), ASTAR(Nil, r)) +} + +// derivative w.r.t. a string (iterates bder) +@tailrec +def bders (s: List[Char], r: ARexp) : ARexp = s match { + case Nil => r + case c::s => bders(s, bder(c, r)) +} + +// main unsimplified lexing function (produces a bitsequence) +def blex(r: ARexp, s: List[Char]) : Bits = s match { + case Nil => if (bnullable(r)) bmkeps(r) else throw new Exception("Not matched") + case c::cs => blex(bder(c, r), cs) +} + +// calls blex and decodes the value +def blexing(r: Rexp, s: String) : Val = + vdecode(r, blex(internalise(r), s.toList)) + + +// example by Tudor +//val reg = (STAR("a") ~ ("b" | "c")).% + +//println(blexing(reg, "aab")) + + +//======================= +// simplification +// + +def flts(rs: List[ARexp]) : List[ARexp] = rs match { + case Nil => Nil + case AZERO :: rs => flts(rs) + case AALTS(bs, rs1) :: rs => rs1.map(fuse(bs, _)) ++ flts(rs) + case r1 :: rs => r1 :: flts(rs) +} + +def bsimp(r: ARexp): ARexp = r match { + case ASEQ(bs1, r1, r2) => (bsimp(r1), bsimp(r2)) match { + case (AZERO, _) => AZERO + case (_, AZERO) => AZERO + case (AONE(bs2), r2s) => fuse(bs1 ++ bs2, r2s) + // needed in order to keep the size down + case (AALTS(bs, rs), r2) => AALTS(bs1 ++ bs, rs.map(ASEQ(Nil, _, r2))) + case (r1s, r2s) => ASEQ(bs1, r1s, r2s) + } + // distinctBy deletes copies of the same "erased" regex + case AALTS(bs1, rs) => (flts(rs.map(bsimp))).distinctBy(erase) match { + case Nil => AZERO + case r::Nil => fuse(bs1, r) + case rs => AALTS(bs1, rs) + } + case r => r +} + +def blex_simp(r: ARexp, s: List[Char]) : Bits = s match { + case Nil => if (bnullable(r)) bmkeps(r) + else throw new Exception("Not matched") + case c::cs => blex_simp(bsimp(bder(c, r)), cs) +} + +// blexing_simp decodes a value from the bitsequence (potentially slow) +// blexing2_simp decodes a string-list from the bitsequence +def blexing_simp(r: Rexp, s: String) : Val = + vdecode(r, blex_simp(internalise(r), s.toList)) + +def blexing2_simp(r: Rexp, s: String) : List[String] = + sdecode(r, blex_simp(internalise(r), s.toList)) + + +//println(blexing_simp(reg, "aab")) + + +// extracts a string from value +def flatten(v: Val) : String = v match { + case Empty => "" + case Chr(c) => c.toString + case Left(v) => flatten(v) + case Right(v) => flatten(v) + case Sequ(v1, v2) => flatten(v1) + flatten(v2) + case Stars(vs) => vs.map(flatten).mkString +} + +// extracts an environment from a value +def env(v: Val) : List[(String, String)] = v match { + case Empty => Nil + case Chr(c) => Nil + case Left(v) => env(v) + case Right(v) => env(v) + case Sequ(v1, v2) => env(v1) ::: env(v2) + case Stars(vs) => vs.flatMap(env) + case Recd(x, v) => (x, flatten(v))::env(v) +} + +def bsize(a: ARexp) = size(erase(a)) + +// Some Tests +//============ + + +def time_needed[T](i: Int, code: => T) = { + val start = System.nanoTime() + for (j <- 1 to i) code + val end = System.nanoTime() + (end - start)/(i * 1.0e9) +} + +/* +val evil1 = STAR(STAR("a")) ~ "b" +val evil2 = STAR(STAR(STAR("a"))) ~ "b" +val evil3 = STAR("aa" | "a") + +for(i <- 0 to 10000 by 1000) { + println(time_needed(1, blexing2_simp(evil1, "a"*i ++ "b"))) +} + +for(i <- 0 to 10000 by 1000) { + println(time_needed(1, blexing2_simp(evil2, "a"*i ++ "b"))) +} + +for(i <- 0 to 10000 by 1000) { + println(time_needed(1, blexing2_simp(evil3, "a"*i))) +} +*/ + +// WHILE LANGUAGE +//================ +def PLUS(r: Rexp) = r ~ r.% + +def Range(s : List[Char]) : Rexp = s match { + case Nil => ZERO + case c::Nil => CHAR(c) + case c::s => ALT(CHAR(c), Range(s)) +} +def RANGE(s: String) = Range(s.toList) + +val SYM = RANGE("ABCDEFGHIJKLMNOPQRSTUVXYZabcdefghijklmnopqrstuvwxyz_") +val DIGIT = RANGE("0123456789") +val ID = SYM ~ (SYM | DIGIT).% +val NUM = PLUS(DIGIT) +val KEYWORD : Rexp = "skip" | "while" | "do" | "if" | "then" | "else" | "read" | "write" +val SEMI: Rexp = ";" +val OP: Rexp = ":=" | "=" | "-" | "+" | "*" | "!=" | "<" | ">" +val WHITESPACE = PLUS(" " | "\n" | "\t") +val RPAREN: Rexp = "{" +val LPAREN: Rexp = "}" +val STRING: Rexp = "\"" ~ SYM.% ~ "\"" + + +val WHILE_REGS = (("k" $ KEYWORD) | + ("i" $ ID) | + ("o" $ OP) | + ("n" $ NUM) | + ("s" $ SEMI) | + ("str" $ STRING) | + ("p" $ (LPAREN | RPAREN)) | + ("w" $ WHITESPACE)).% + + +// Some Simple While Tests +//======================== + +val prog0 = """read n""" +println(s"test: $prog0") +println(env(blexing_simp(WHILE_REGS, prog0))) +println(blexing2_simp(WHILE_REGS, prog0)) + +val prog1 = """read n; write n""" +println(s"test: $prog1") +println(env(blexing_simp(WHILE_REGS, prog1))) +println(blexing2_simp(WHILE_REGS, prog1)) + +val prog2 = """ +write "Fib"; +read n; +minus1 := 0; +minus2 := 1; +while n > 0 do { + temp := minus2; + minus2 := minus1 + minus2; + minus1 := temp; + n := n - 1 +}; +write "Result"; +write minus2 +""" + +println("lexing fib program (once)") +println(blexing2_simp(WHILE_REGS, prog2).filter(s => s == "" || !s.startsWith("w"))) + +val n = 60 +println(s"lexing fib program ($n times, size ${prog2.length * n})") +println(time_needed(1, blexing2_simp(WHILE_REGS, prog2 * n))) diff -r fedc16924b76 -r e752d84225ec progs/scala/re-annotated2.sc --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/progs/scala/re-annotated2.sc Mon Feb 22 03:22:26 2021 +0000 @@ -0,0 +1,418 @@ +// uses bitcode sequences and annotated regular +// expressions +// +// for basic regular expressions and RECD +// +// the main difference is that it includes +// also character sets +// +// uses a non-standard extraction method for generating +// tokens (this is tail-recursive) +// +// can match 200 copies of the fib-program (size 35000) +// in about 20 secs + + +import scala.language.implicitConversions +import scala.language.reflectiveCalls +import scala.annotation.tailrec + +// standard regular expressions +abstract class Rexp +case object ZERO extends Rexp +case object ONE extends Rexp +case class ALT(r1: Rexp, r2: Rexp) extends Rexp +case class SEQ(r1: Rexp, r2: Rexp) extends Rexp +case class STAR(r: Rexp) extends Rexp +case class RECD(x: String, r: Rexp) extends Rexp +case class CHARSET(f: Char => Boolean) extends Rexp + +//abbreviation for single characters +def CHAR(c: Char) = CHARSET(_ == c) + +// bit-codes with additional information about +// matched characters +abstract class Bit +case object Z extends Bit +case object S extends Bit +case class C(c: Char) extends Bit + +type Bits = List[Bit] + +// annotated regular expressions +abstract class ARexp +case object AZERO extends ARexp +case class AONE(bs: Bits) extends ARexp +case class AALTS(bs: Bits, rs: List[ARexp]) extends ARexp +case class ASEQ(bs: Bits, r1: ARexp, r2: ARexp) extends ARexp +case class ASTAR(bs: Bits, r: ARexp) extends ARexp +case class ACHARSET(bs: Bits, f: Char => Boolean) extends ARexp + +// an abbreviation for binary alternatives +def AALT(bs: Bits, r1: ARexp, r2: ARexp) = AALTS(bs, List(r1, r2)) + +abstract class Val +case object Empty extends Val +case class Chr(c: Char) extends Val +case class Sequ(v1: Val, v2: Val) extends Val +case class Left(v: Val) extends Val +case class Right(v: Val) extends Val +case class Stars(vs: List[Val]) extends Val +case class Recd(x: String, v: Val) extends Val + +// some convenience for typing in regular expressions +def charlist2rexp(s: List[Char]): Rexp = s match { + case Nil => ONE + case c::Nil => CHAR(c) + case c::s => SEQ(CHAR(c), charlist2rexp(s)) +} +implicit def string2rexp(s: String) : Rexp = charlist2rexp(s.toList) + +implicit def RexpOps(r: Rexp) = new { + def | (s: Rexp) = ALT(r, s) + def % = STAR(r) + def ~ (s: Rexp) = SEQ(r, s) +} + +implicit def stringOps(s: String) = new { + def | (r: Rexp) = ALT(s, r) + def | (r: String) = ALT(s, r) + def % = STAR(s) + def ~ (r: Rexp) = SEQ(s, r) + def ~ (r: String) = SEQ(s, r) + def $ (r: Rexp) = RECD(s, r) +} + +def size(r: Rexp) : Int = r match { + case ZERO => 1 + case ONE => 1 + case ALT(r1, r2) => 1 + size(r1) + size(r2) + case SEQ(r1, r2) => 1 + size(r1) + size(r2) + case STAR(r) => 1 + size(r) + case RECD(_, r) => 1 + size(r) + case CHARSET(_) => 1 +} + + + +// Bitcoded + Annotation +//======================= + +//erase function: extracts the Rexp from ARexp +def erase(r:ARexp): Rexp = r match{ + case AZERO => ZERO + case AONE(_) => ONE + case AALTS(bs, Nil) => ZERO + case AALTS(bs, r::Nil) => erase(r) + case AALTS(bs, r::rs) => ALT(erase(r), erase(AALTS(bs, rs))) + case ASEQ(bs, r1, r2) => SEQ (erase(r1), erase(r2)) + case ASTAR(cs, r)=> STAR(erase(r)) + case ACHARSET(bs, f) => CHARSET(f) +} + +def fuse(bs: Bits, r: ARexp) : ARexp = r match { + case AZERO => AZERO + case AONE(cs) => AONE(bs ++ cs) + case AALTS(cs, rs) => AALTS(bs ++ cs, rs) + case ASEQ(cs, r1, r2) => ASEQ(bs ++ cs, r1, r2) + case ASTAR(cs, r) => ASTAR(bs ++ cs, r) + case ACHARSET(cs, f) => ACHARSET(bs ++ cs, f) +} + +def internalise(r: Rexp) : ARexp = r match { + case ZERO => AZERO + case ONE => AONE(Nil) + case ALT(r1, r2) => + AALT(Nil, fuse(List(Z), internalise(r1)), fuse(List(S), internalise(r2))) + case SEQ(r1, r2) => ASEQ(Nil, internalise(r1), internalise(r2)) + case STAR(r) => ASTAR(Nil, internalise(r)) + case RECD(_, r) => internalise(r) + case CHARSET(f) => ACHARSET(Nil, f) +} + +// example +// internalise(("a" | "ab") ~ ("b" | "")) + + +// decoding of a value from a bitsequence +// (this is not tail-recursive and therefore a potential bottleneck) +def vdecode_aux(r: Rexp, bs: Bits) : (Val, Bits) = (r, bs) match { + case (ONE, bs) => (Empty, bs) + case (ALT(r1, r2), Z::bs) => { + val (v, bs1) = vdecode_aux(r1, bs) + (Left(v), bs1) + } + case (ALT(r1, r2), S::bs) => { + val (v, bs1) = vdecode_aux(r2, bs) + (Right(v), bs1) + } + case (SEQ(r1, r2), bs) => { + val (v1, bs1) = vdecode_aux(r1, bs) + val (v2, bs2) = vdecode_aux(r2, bs1) + (Sequ(v1, v2), bs2) + } + case (STAR(r1), Z::bs) => { + val (v, bs1) = vdecode_aux(r1, bs) + val (Stars(vs), bs2) = vdecode_aux(STAR(r1), bs1) + (Stars(v::vs), bs2) + } + case (STAR(_), S::bs) => (Stars(Nil), bs) + case (RECD(s, r1), bs) => + val (v, bs1) = vdecode_aux(r1, bs) + (Recd(s, v), bs1) + case (CHARSET(_), C(c)::bs) => (Chr(c), bs) +} + +def vdecode(r: Rexp, bs: Bits) = vdecode_aux(r, bs) match { + case (v, Nil) => v + case _ => throw new Exception("Not decodable") +} + +// decoding of sequence of string tokens from a bitsequence +// tail-recursive version using an accumulator (alternative for +// vdecode) +@tailrec +def sdecode_aux(rs: List[Rexp], bs: Bits, acc: List[String]) : List[String] = (rs, bs) match { + case (Nil, _) => acc + case (_, Nil) => acc + case (ONE::rest, bs) => sdecode_aux(rest, bs, acc) + case (ALT(r1, r2)::rest, Z::bs) => sdecode_aux(r1::rest, bs, acc) + case (ALT(r1, r2)::rest, S::bs) => sdecode_aux(r2::rest, bs, acc) + case (SEQ(r1, r2)::rest, bs) => sdecode_aux(r1::r2::rest, bs, acc) + case (STAR(r1)::rest, Z::bs) => sdecode_aux(r1::STAR(r1)::rest, bs, acc) + case (STAR(_)::rest, S::bs) => sdecode_aux(rest, bs, acc) + case (RECD(s, r1)::rest, bs) => sdecode_aux(r1::rest, bs, s"$s:"::acc) + case (CHARSET(_)::rest, C(c)::bs) => + sdecode_aux(rest, bs, (acc.head :+ c)::acc.tail) +} + +def sdecode(r: Rexp, bs: Bits) : List[String] = + sdecode_aux(List(r), bs, List("")).reverse.tail + + + +// nullable function: tests whether the a (annotated) +// regular expression can recognise the empty string +def bnullable (r: ARexp) : Boolean = r match { + case AZERO => false + case AONE(_) => true + case AALTS(_, rs) => rs.exists(bnullable) + case ASEQ(_, r1, r2) => bnullable(r1) && bnullable(r2) + case ASTAR(_, _) => true + case ACHARSET(_, _) => false +} + +def bmkeps(r: ARexp) : Bits = r match { + case AONE(bs) => bs + case AALTS(bs, r::Nil) => bs ++ bmkeps(r) + case AALTS(bs, r::rs) => + if (bnullable(r)) bs ++ bmkeps(r) else bmkeps(AALTS(bs, rs)) + case ASEQ(bs, r1, r2) => bs ++ bmkeps(r1) ++ bmkeps(r2) + case ASTAR(bs, r) => bs ++ List(S) + case _ => throw new Exception("trying to bmkeps from non-nullable regex") +} + +// derivative of a regular expression w.r.t. a character +def bder(c: Char, r: ARexp) : ARexp = r match { + case AZERO => AZERO + case AONE(_) => AZERO + case AALTS(bs, rs) => AALTS(bs, rs.map(bder(c, _))) + case ASEQ(bs, r1, r2) => + if (bnullable(r1)) AALT(bs, ASEQ(Nil, bder(c, r1), r2), fuse(bmkeps(r1), bder(c, r2))) + else ASEQ(bs, bder(c, r1), r2) + case ASTAR(bs, r) => ASEQ(bs, fuse(List(Z), bder(c, r)), ASTAR(Nil, r)) + case ACHARSET(bs, f) => if(f(c)) AONE(bs :+ C(c)) else AZERO +} + +// derivative w.r.t. a string (iterates bder) +@tailrec +def bders (s: List[Char], r: ARexp) : ARexp = s match { + case Nil => r + case c::s => bders(s, bder(c, r)) +} + +// main unsimplified lexing function (produces a bitsequence) +def blex(r: ARexp, s: List[Char]) : Bits = s match { + case Nil => if (bnullable(r)) bmkeps(r) else throw new Exception("Not matched") + case c::cs => blex(bder(c, r), cs) +} + +// calls blex and decodes the value +def blexing(r: Rexp, s: String) : Val = + vdecode(r, blex(internalise(r), s.toList)) + + +// example by Tudor +//val reg = (STAR("a") ~ ("b" | "c")).% + +//println(blexing(reg, "aab")) + + +//======================= +// simplification +// + +def flts(rs: List[ARexp]) : List[ARexp] = rs match { + case Nil => Nil + case AZERO :: rs => flts(rs) + case AALTS(bs, rs1) :: rs => rs1.map(fuse(bs, _)) ++ flts(rs) + case r1 :: rs => r1 :: flts(rs) +} + +def bsimp(r: ARexp): ARexp = r match { + case ASEQ(bs1, r1, r2) => (bsimp(r1), bsimp(r2)) match { + case (AZERO, _) => AZERO + case (_, AZERO) => AZERO + case (AONE(bs2), r2s) => fuse(bs1 ++ bs2, r2s) + // needed in order to keep the size down + case (AALTS(bs, rs), r2) => AALTS(bs1 ++ bs, rs.map(ASEQ(Nil, _, r2))) + case (r1s, r2s) => ASEQ(bs1, r1s, r2s) + } + // distinctBy deletes copies of the same "erased" regex + case AALTS(bs1, rs) => (flts(rs.map(bsimp))).distinctBy(erase) match { + case Nil => AZERO + case r::Nil => fuse(bs1, r) + case rs => AALTS(bs1, rs) + } + case r => r +} + +def blex_simp(r: ARexp, s: List[Char]) : Bits = s match { + case Nil => if (bnullable(r)) bmkeps(r) + else throw new Exception("Not matched") + case c::cs => blex_simp(bsimp(bder(c, r)), cs) +} + +// blexing_simp decodes a value from the bitsequence (potentially slow) +// blexing2_simp decodes a string-list from the bitsequence +def blexing_simp(r: Rexp, s: String) : Val = + vdecode(r, blex_simp(internalise(r), s.toList)) + +def blexing2_simp(r: Rexp, s: String) : List[String] = + sdecode(r, blex_simp(internalise(r), s.toList)) + + +//println(blexing_simp(reg, "aab")) + + +// extracts a string from value +def flatten(v: Val) : String = v match { + case Empty => "" + case Chr(c) => c.toString + case Left(v) => flatten(v) + case Right(v) => flatten(v) + case Sequ(v1, v2) => flatten(v1) + flatten(v2) + case Stars(vs) => vs.map(flatten).mkString +} + +// extracts an environment from a value +def env(v: Val) : List[(String, String)] = v match { + case Empty => Nil + case Chr(c) => Nil + case Left(v) => env(v) + case Right(v) => env(v) + case Sequ(v1, v2) => env(v1) ::: env(v2) + case Stars(vs) => vs.flatMap(env) + case Recd(x, v) => (x, flatten(v))::env(v) +} + +def bsize(a: ARexp) = size(erase(a)) + +// Some Tests +//============ + + +def time_needed[T](i: Int, code: => T) = { + val start = System.nanoTime() + for (j <- 1 to i) code + val end = System.nanoTime() + (end - start)/(i * 1.0e9) +} + +val evil1 = STAR(STAR("a")) ~ "b" +val evil2 = STAR(STAR(STAR("a"))) ~ "b" +val evil3 = STAR("aa" | "a") + +/* +println("evil1") +for(i <- 0 to 10000 by 1000) { + println(time_needed(1, blexing2_simp(evil1, "a"*i ++ "b"))) +} +*/ + +/* +println("evil2") +for(i <- 0 to 10000 by 1000) { + println(time_needed(1, blexing2_simp(evil2, "a"*i ++ "b"))) +} +*/ + +/* +println("evil3") +for(i <- 0 to 10000 by 1000) { + println(time_needed(1, blexing2_simp(evil3, "a"*i))) +} +*/ + +// WHILE LANGUAGE +//================ +def PLUS(r: Rexp) = r ~ r.% +def RANGE(s: String) = CHARSET(s.toSet) + +val SYM = RANGE("ABCDEFGHIJKLMNOPQRSTUVXYZabcdefghijklmnopqrstuvwxyz_") +val DIGIT = RANGE("0123456789") +val ID = SYM ~ (SYM | DIGIT).% +val NUM = PLUS(DIGIT) +val KEYWORD : Rexp = "skip" | "while" | "do" | "if" | "then" | "else" | "read" | "write" +val SEMI: Rexp = ";" +val OP: Rexp = ":=" | "=" | "-" | "+" | "*" | "!=" | "<" | ">" +val WHITESPACE = PLUS(" " | "\n" | "\t") +val RPAREN: Rexp = "{" +val LPAREN: Rexp = "}" +val STRING: Rexp = "\"" ~ SYM.% ~ "\"" + + +val WHILE_REGS = (("k" $ KEYWORD) | + ("i" $ ID) | + ("o" $ OP) | + ("n" $ NUM) | + ("s" $ SEMI) | + ("str" $ STRING) | + ("p" $ (LPAREN | RPAREN)) | + ("w" $ WHITESPACE)).% + + +// Some Simple While Tests +//======================== + +val prog0 = """read n""" +println(s"test: $prog0") +println(env(blexing_simp(WHILE_REGS, prog0))) +println(blexing2_simp(WHILE_REGS, prog0)) + +val prog1 = """read n; write n""" +println(s"test: $prog1") +println(env(blexing_simp(WHILE_REGS, prog1))) +println(blexing2_simp(WHILE_REGS, prog1)) + +val prog2 = """ +write "Fib"; +read n; +minus1 := 0; +minus2 := 1; +while n > 0 do { + temp := minus2; + minus2 := minus1 + minus2; + minus1 := temp; + n := n - 1 +}; +write "Result"; +write minus2 +""" + +println("lexing fib program (once)") +println(blexing2_simp(WHILE_REGS, prog2).filter(s => s == "" || !s.startsWith("w"))) + +val n = 200 +println(s"lexing fib program ($n times, size ${prog2.length * n})") +println(time_needed(1, blexing2_simp(WHILE_REGS, prog2 * n))) diff -r fedc16924b76 -r e752d84225ec progs/scala/re-bit2.scala --- a/progs/scala/re-bit2.scala Sat Oct 24 12:13:39 2020 +0100 +++ b/progs/scala/re-bit2.scala Mon Feb 22 03:22:26 2021 +0000 @@ -2,7 +2,7 @@ import scala.language.reflectiveCalls import scala.annotation.tailrec - +// standard regular expressions abstract class Rexp case object ZERO extends Rexp case object ONE extends Rexp @@ -11,14 +11,13 @@ case class SEQ(r1: Rexp, r2: Rexp) extends Rexp case class STAR(r: Rexp) extends Rexp - - abstract class Bit case object Z extends Bit case object S extends Bit type Bits = List[Bit] +// annotated regular expressions abstract class ARexp case object AZERO extends ARexp case class AONE(bs: Bits) extends ARexp @@ -27,6 +26,7 @@ case class ASEQ(bs: Bits, r1: ARexp, r2: ARexp) extends ARexp case class ASTAR(bs: Bits, r: ARexp) extends ARexp +// an abbreviation for binary alternatives def AALT(bs: Bits, r1: ARexp, r2: ARexp) = AALTS(bs, List(r1, r2)) abstract class Val @@ -36,8 +36,7 @@ case class Left(v: Val) extends Val case class Right(v: Val) extends Val case class Stars(vs: List[Val]) extends Val -case class Position(i: Int, v: Val) extends Val // for testing purposes -case object Undefined extends Val // for testing purposes + // some convenience for typing in regular expressions def charlist2rexp(s: List[Char]): Rexp = s match { @@ -62,71 +61,10 @@ } -// nullable function: tests whether the regular -// expression can recognise the empty string -def nullable(r: Rexp) : Boolean = r match { - case ZERO => false - case ONE => true - case CHAR(_) => false - case ALT(r1, r2) => nullable(r1) || nullable(r2) - case SEQ(r1, r2) => nullable(r1) && nullable(r2) - case STAR(_) => true -} - -// derivative of a regular expression w.r.t. a character -def der(c: Char, r: Rexp) : Rexp = r match { - case ZERO => ZERO - case ONE => ZERO - case CHAR(d) => if (c == d) ONE else ZERO - case ALT(r1, r2) => ALT(der(c, r1), der(c, r2)) - case SEQ(r1, r2) => - if (nullable(r1)) ALT(SEQ(der(c, r1), r2), der(c, r2)) - else SEQ(der(c, r1), r2) - case STAR(r) => SEQ(der(c, r), STAR(r)) -} - -// derivative w.r.t. a string (iterates der) -def ders(s: List[Char], r: Rexp) : Rexp = s match { - case Nil => r - case c::s => ders(s, der(c, r)) -} - -// mkeps and injection part -def mkeps(r: Rexp) : Val = r match { - case ONE => Empty - case ALT(r1, r2) => - if (nullable(r1)) Left(mkeps(r1)) else Right(mkeps(r2)) - case SEQ(r1, r2) => Sequ(mkeps(r1), mkeps(r2)) - case STAR(r) => Stars(Nil) -} - - -def inj(r: Rexp, c: Char, v: Val) : Val = (r, v) match { - case (STAR(r), Sequ(v1, Stars(vs))) => Stars(inj(r, c, v1)::vs) - case (SEQ(r1, r2), Sequ(v1, v2)) => Sequ(inj(r1, c, v1), v2) - case (SEQ(r1, r2), Left(Sequ(v1, v2))) => Sequ(inj(r1, c, v1), v2) - case (SEQ(r1, r2), Right(v2)) => Sequ(mkeps(r1), inj(r2, c, v2)) - case (ALT(r1, r2), Left(v1)) => Left(inj(r1, c, v1)) - case (ALT(r1, r2), Right(v2)) => Right(inj(r2, c, v2)) - case (CHAR(d), Empty) => Chr(c) -} - -// main lexing function (produces a value) -// - no simplification -def lex(r: Rexp, s: List[Char]) : Val = s match { - case Nil => if (nullable(r)) mkeps(r) - else throw new Exception("Not matched") - case c::cs => inj(r, c, lex(der(c, r), cs)) -} - -def lexing(r: Rexp, s: String) : Val = lex(r, s.toList) - - - // Bitcoded + Annotation //======================= -//erase function: extracts the regx from Aregex +//erase function: extracts the Rexp from ARexp def erase(r:ARexp): Rexp = r match{ case AZERO => ZERO case AONE(_) => ONE @@ -138,7 +76,6 @@ case ASTAR(cs, r)=> STAR(erase(r)) } -// translation into ARexps def fuse(bs: Bits, r: ARexp) : ARexp = r match { case AZERO => AZERO case AONE(cs) => AONE(bs ++ cs) @@ -161,7 +98,7 @@ internalise(("a" | "ab") ~ ("b" | "")) - +// decoding of a value from a bitsequence def decode_aux(r: Rexp, bs: Bits) : (Val, Bits) = (r, bs) match { case (ONE, bs) => (Empty, bs) case (CHAR(c), bs) => (Chr(c), bs) @@ -202,8 +139,8 @@ } -// nullable function: tests whether the aregular -// expression can recognise the empty string +// nullable function: tests whether the an (annotated) +// regular expression can recognise the empty string def bnullable (r: ARexp) : Boolean = r match { case AZERO => false case AONE(_) => true @@ -234,7 +171,7 @@ case ASTAR(bs, r) => ASEQ(bs, fuse(List(Z), bder(c, r)), ASTAR(Nil, r)) } -// derivative w.r.t. a string (iterates der) +// derivative w.r.t. a string (iterates bder) @tailrec def bders (s: List[Char], r: ARexp) : ARexp = s match { case Nil => r @@ -252,9 +189,9 @@ // example by Tudor -val reg = ((("a".%)) ~ ("b" | "c")).% +//val reg = ((("a".%)) ~ ("b" | "c")).% -println(blexing(reg, "aab")) +//println(blexing(reg, "aab")) //======================= @@ -289,7 +226,7 @@ //case (AALTS(bs, rs), r2) => AALTS(bs, rs.map(ASEQ(Nil, _, r2))) case (r1s, r2s) => ASEQ(bs1, r1s, r2s) } - case AALTS(bs1, rs) => distinctBy(flts(rs.map(bsimp)), erase) match { + case AALTS(bs1, rs) => (flts(rs.map(bsimp))).distinctBy(erase) match { case Nil => AZERO case r::Nil => fuse(bs1, r) case rs => AALTS(bs1, rs) @@ -306,141 +243,8 @@ def blexing_simp(r: Rexp, s: String) : Val = decode(r, blex_simp(internalise(r), s.toList)) -println(blexing_simp(reg, "aab")) - -// bsimp2 by Chengsong - -def pos_i(rs: List[ARexp], v: Val): Int = (rs, v) match { - case (r::Nil, v1) => 0 - case ( r::rs1, Right(v)) => pos_i(rs1, v) + 1 - case ( r::rs1, Left(v) ) => 0 -} - -def pos_v(rs: List[ARexp], v: Val): Val = (rs, v) match { - case (r::Nil, v1) => v1 - case (r::rs1, Right(v)) => pos_v(rs1, v) - case (r::rs1, Left(v) ) => v -} - -def unify(rs: List[ARexp], v: Val): Val = { - Position(pos_i(rs, v), pos_v(rs, v)) -} - -// coat does the job of "coating" a value -// given the number of right outside it -def coat(v: Val, i: Int) : Val = i match { - case 0 => v - case i => if (i > 0) coat(Right(v), i - 1) else { println(v, i); throw new Exception("coat minus")} -} - -def distinctBy2[B, C](v: Val, xs: List[B], f: B => C, acc: List[C] = Nil, res: List[B] = Nil): (List[B], Val) = xs match { - case Nil => (res, v) - case (x::xs) => { - val re = f(x) - if (acc.contains(re)) v match { - case Position(i, v0) => distinctBy2(Position(i - 1, v0), xs, f, acc, res) - case _ => throw new Exception("dB2") - } - else distinctBy2(v, xs, f, re::acc, x::res) - } - } - - -def flats(rs: List[ARexp]): List[ARexp] = rs match { - case Nil => Nil - case AZERO :: rs1 => flats(rs1) - case AALTS(bs, rs1) :: rs2 => rs1.map(fuse(bs, _)) ::: flats(rs2) - case r1 :: rs2 => r1 :: flats(rs2) - } - - +//println(blexing_simp(reg, "aab")) -def flats2(front: List[ARexp], i: Int, rs: List[ARexp], v: Val): (List[ARexp], Val) = v match { - case Position(j, v0) => { - if (i < 0) (front ::: flats(rs), Position(j, v0) ) - else if(i == 0){ - rs match { - case AALTS(bs, rs1) :: rs2 => ( (front ::: rs1.map(fuse(bs, _))):::flats(rs2), Position(j + rs1.length - 1, pos_v(rs1, v0))) - case r::rs2 => (front ::: List(r) ::: flats(rs2), Position(j, v0)) - case _ => throw new Exception("flats2 i = 0") - } - } - else{ - rs match { - case AZERO::rs1 => flats2(front, i - 1, rs1, Position(j - 1, v0)) - case AALTS(bs, rs1) ::rs2 => flats2(front:::rs1.map(fuse(bs, _)), i - 1, rs2, Position(j + rs1.length - 1, v0)) - case r::rs1 => flats2(front:::List(r), i - 1, rs1, Position(j, v0)) - case _ => throw new Exception("flats2 i>0") - } - } - } - case _ => throw new Exception("flats2 error") - } - -def deunify(rs_length: Int, v: Val): Val = v match{ - case Position(i, v0) => { - if (i <0 ) { println(rs_length, v); throw new Exception("deunify minus") } - else if (rs_length == 1) {println(v); v} - else if (rs_length - 1 == i) coat(v0, i) - else coat(Left(v0), i) - } - case _ => throw new Exception("deunify error") -} - - -def bsimp2(r: ARexp, v: Val): (ARexp, Val) = (r, v) match { - case (ASEQ(bs1, r1, r2), Sequ(v1, v2)) => (bsimp2(r1, v1), bsimp2(r2, v2)) match { - case ((AZERO, _), (_, _)) => (AZERO, Undefined) - case ((_, _), (AZERO, _)) => (AZERO, Undefined) - case ((AONE(bs2), v1s) , (r2s, v2s)) => (fuse(bs1 ++ bs2, r2s), v2s) - // v2 tells how to retrieve bits in r2s, which is enough as the bits - // of the first part of the sequence has already been integrated to - // the top level of the second regx. - case ((r1s, v1s), (r2s, v2s)) => (ASEQ(bs1, r1s, r2s), Sequ(v1s, v2s)) - } - case (AALTS(bs1, rs), v) => { - val vlist = unify(rs, v) - vlist match { - case Position(i, v0) => { - val v_simp = bsimp2(rs(i), v0)._2 - val rs_simp = rs.map(bsimp) - val flat_res = flats2( Nil, i, rs_simp, Position(i, v_simp) ) - val dist_res = distinctBy2(flat_res._2, flat_res._1, erase) - val rs_new = dist_res._1 - val v_new = deunify(rs_new.length, dist_res._2) - rs_new match { - case Nil => (AZERO, Undefined) - case s :: Nil => (fuse(bs1, s), v_new) - case rs => (AALTS(bs1, rs), v_new) - } - } - case _ => throw new Exception("Funny vlist bsimp2") - } - } - // STAR case - // case ASTAR(bs, r) => ASTAR(bs, bsimp(r)) - case (r, v) => (r, v) - } - - - -val dr = ASEQ(List(),AALTS(List(S),List(AONE(List(Z)), AONE(List(S)))),ASTAR(List(),AALTS(List(),List(ACHAR(List(Z),'a'), ACHAR(List(S),'a'))))) -val dv = Sequ(Left(Empty),Stars(List())) -println(bsimp2(dr, dv)) - - -/* -def blex_simp2(r: ARexp, s: List[Char]) : Bits = s match { - case Nil => if (bnullable(r)) bmkeps(r) - else throw new Exception("Not matched") - case c::cs => blex(bsimp2(bder(c, r)), cs) -} - -def blexing_simp2(r: Rexp, s: String) : Val = - decode(r, blex_simp2(internalise(r), s.toList)) - -println(blexing_simp2(reg, "aab")) -*/ // extracts a string from value def flatten(v: Val) : String = v match {