RegexExplosionPlot/evilForBsimp.sc
author Chengsong
Mon, 24 Jul 2023 11:09:48 +0100
changeset 666 6da4516ea87d
parent 545 333013923c5a
permissions -rw-r--r--
more changes to figures & benchmarking
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
517
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
     1
// A simple lexer inspired by work of Sulzmann & Lu
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
     2
//==================================================
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
     3
//
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
     4
// Call the test cases with 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
     5
//
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
     6
//   amm lexer.sc small
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
     7
//   amm lexer.sc fib
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
     8
//   amm lexer.sc loops
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
     9
//   amm lexer.sc email
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    10
//
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    11
//   amm lexer.sc all
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    12
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    13
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    14
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    15
// regular expressions including records
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    16
abstract class Rexp 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    17
case object ZERO extends Rexp
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    18
case object ONE extends Rexp
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    19
case object ANYCHAR extends Rexp
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    20
case class CHAR(c: Char) extends Rexp
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    21
case class ALTS(r1: Rexp, r2: Rexp) extends Rexp 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    22
case class SEQ(r1: Rexp, r2: Rexp) extends Rexp 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    23
case class STAR(r: Rexp) extends Rexp 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    24
case class RECD(x: String, r: Rexp) extends Rexp  
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    25
case class NTIMES(n: Int, r: Rexp) extends Rexp
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    26
case class OPTIONAL(r: Rexp) extends Rexp
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    27
case class NOT(r: Rexp) extends Rexp
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    28
                // records for extracting strings or tokens
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    29
  
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    30
// values  
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    31
abstract class Val
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    32
case object Failure extends Val
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    33
case object Empty extends Val
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    34
case class Chr(c: Char) extends Val
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    35
case class Sequ(v1: Val, v2: Val) extends Val
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    36
case class Left(v: Val) extends Val
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    37
case class Right(v: Val) extends Val
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    38
case class Stars(vs: List[Val]) extends Val
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    39
case class Rec(x: String, v: Val) extends Val
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    40
case class Ntime(vs: List[Val]) extends Val
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    41
case class Optionall(v: Val) extends Val
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    42
case class Nots(s: String) extends Val
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    43
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    44
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    45
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    46
abstract class Bit
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    47
case object Z extends Bit
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    48
case object S extends Bit
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    49
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    50
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    51
type Bits = List[Bit]
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    52
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    53
abstract class ARexp 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    54
case object AZERO extends ARexp
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    55
case class AONE(bs: Bits) extends ARexp
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    56
case class ACHAR(bs: Bits, c: Char) extends ARexp
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    57
case class AALTS(bs: Bits, rs: List[ARexp]) extends ARexp 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    58
case class ASEQ(bs: Bits, r1: ARexp, r2: ARexp) extends ARexp 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    59
case class ASTAR(bs: Bits, r: ARexp) extends ARexp 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    60
case class ANOT(bs: Bits, r: ARexp) extends ARexp
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    61
case class AANYCHAR(bs: Bits) extends ARexp
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    62
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    63
import scala.util.Try
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    64
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    65
trait Generator[+T] {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    66
    self => // an alias for "this"
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    67
    def generate(): T
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    68
  
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    69
    def gen(n: Int) : List[T] = 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    70
      if (n == 0) Nil else self.generate() :: gen(n - 1)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    71
    
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    72
    def map[S](f: T => S): Generator[S] = new Generator[S] {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    73
      def generate = f(self.generate())  
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    74
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    75
    def flatMap[S](f: T => Generator[S]): Generator[S] = new Generator[S] {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    76
      def generate = f(self.generate()).generate()
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    77
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    78
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    79
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    80
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    81
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    82
  // tests a property according to a given random generator
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    83
  def test[T](r: Generator[T], amount: Int = 100)(pred: T => Boolean) {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    84
    for (_ <- 0 until amount) {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    85
      val value = r.generate()
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    86
      assert(pred(value), s"Test failed for: $value")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    87
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    88
    println(s"Test passed $amount times")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    89
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    90
  def test2[T, S](r: Generator[T], s: Generator[S], amount: Int = 100)(pred: (T, S) => Boolean) {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    91
    for (_ <- 0 until amount) {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    92
      val valueR = r.generate()
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    93
      val valueS = s.generate()
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    94
      assert(pred(valueR, valueS), s"Test failed for: $valueR, $valueS")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    95
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    96
    println(s"Test passed $amount times")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    97
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    98
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
    99
// random integers
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   100
val integers = new Generator[Int] {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   101
  val rand = new java.util.Random
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   102
  def generate() = rand.nextInt()
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   103
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   104
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   105
// random booleans
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   106
val booleans = integers.map(_ > 0)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   107
  
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   108
// random integers in the range lo and high  
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   109
def range(lo: Int, hi: Int): Generator[Int] = 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   110
  for (x <- integers) yield (lo + x.abs % (hi - lo)).abs
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   111
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   112
// random characters
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   113
def chars_range(lo: Char, hi: Char) = range(lo, hi).map(_.toChar)  
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   114
val chars = chars_range('a', 'z')
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   115
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   116
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   117
def oneOf[T](xs: T*): Generator[T] = 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   118
  for (idx <- range(0, xs.length)) yield xs(idx)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   119
  
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   120
def single[T](x: T) = new Generator[T] {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   121
  def generate() = x
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   122
}   
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   123
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   124
def pairs[T, U](t: Generator[T], u: Generator[U]): Generator[(T, U)] = 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   125
  for (x <- t; y <- u) yield (x, y)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   126
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   127
// lists
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   128
def emptyLists = single(Nil) 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   129
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   130
def nonEmptyLists : Generator[List[Int]] = 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   131
  for (head <- integers; tail <- lists) yield head :: tail
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   132
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   133
def lists: Generator[List[Int]] = for {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   134
  kind <- booleans
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   135
  list <- if (kind) emptyLists else nonEmptyLists
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   136
} yield list
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   137
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   138
def char_list(len: Int): Generator[List[Char]] = {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   139
  if(len <= 0) single(Nil)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   140
  else{
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   141
    for { 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   142
      c <- chars_range('a', 'c')
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   143
      tail <- char_list(len - 1)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   144
    } yield c :: tail
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   145
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   146
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   147
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   148
def strings(len: Int): Generator[String] = for(cs <- char_list(len)) yield cs.toString
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   149
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   150
def sampleString(r: Rexp) : List[String] = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   151
  case STAR(r) => stringsFromRexp(r).flatMap(s => List("", s, s ++ s))//only generate 0, 1, 2 reptitions
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   152
  case SEQ(r1, r2) => stringsFromRexp(r1).flatMap(s1 => stringsFromRexp(r2).map(s2 => s1 ++ s2) )
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   153
  case ALTS(r1, r2) => throw new Error(s" Rexp ${r} not expected: all alternatives are supposed to have been opened up")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   154
  case ONE => "" :: Nil
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   155
  case ZERO => Nil
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   156
  case CHAR(c) => c.toString :: Nil
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   157
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   158
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   159
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   160
def stringsFromRexp(r: Rexp) : List[String] = 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   161
  breakIntoTerms(r).flatMap(r => sampleString(r))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   162
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   163
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   164
// (simple) binary trees
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   165
trait Tree[T]
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   166
case class Inner[T](left: Tree[T], right: Tree[T]) extends Tree[T]
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   167
case class Leaf[T](x: T) extends Tree[T]
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   168
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   169
def leafs[T](t: Generator[T]): Generator[Leaf[T]] = 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   170
  for (x <- t) yield Leaf(x)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   171
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   172
def inners[T](t: Generator[T]): Generator[Inner[T]] = 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   173
  for (l <- trees(t); r <- trees(t)) yield Inner(l, r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   174
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   175
def trees[T](t: Generator[T]): Generator[Tree[T]] = 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   176
  for (kind <- range(0, 2);  
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   177
       tree <- if (kind == 0) leafs(t) else inners(t)) yield tree
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   178
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   179
// regular expressions
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   180
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   181
// generates random leaf-regexes; prefers CHAR-regexes
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   182
def leaf_rexp() : Generator[Rexp] =
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   183
  for (kind <- range(0, 5);
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   184
       c <- chars_range('a', 'd')) yield
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   185
    kind match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   186
      case 0 => ZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   187
      case 1 => ONE
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   188
      case _ => CHAR(c) 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   189
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   190
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   191
// generates random inner regexes with maximum depth d
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   192
def inner_rexp(d: Int) : Generator[Rexp] =
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   193
  for (kind <- range(0, 3);
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   194
       l <- rexp(d); 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   195
       r <- rexp(d))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   196
  yield kind match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   197
    case 0 => ALTS(l, r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   198
    case 1 => SEQ(l, r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   199
    case 2 => STAR(r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   200
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   201
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   202
// generates random regexes with maximum depth d;
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   203
// prefers inner regexes in 2/3 of the cases
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   204
def rexp(d: Int = 100): Generator[Rexp] = 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   205
  for (kind <- range(0, 3);
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   206
       r <- if (d <= 0 || kind == 0) leaf_rexp() else inner_rexp(d - 1)) yield r
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   207
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   208
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   209
// some test functions for rexps
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   210
def height(r: Rexp) : Int = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   211
  case ZERO => 1
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   212
  case ONE => 1
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   213
  case CHAR(_) => 1
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   214
  case ALTS(r1, r2) => 1 + List(height(r1), height(r2)).max
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   215
  case SEQ(r1, r2) =>  1 + List(height(r1), height(r2)).max
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   216
  case STAR(r) => 1 + height(r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   217
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   218
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   219
// def size(r: Rexp) : Int = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   220
//   case ZERO => 1
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   221
//   case ONE => 1
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   222
//   case CHAR(_) => 1
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   223
//   case ALTS(r1, r2) => 1 + size(r1) + size(r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   224
//   case SEQ(r1, r2) =>  1 + size(r1) + size(r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   225
//   case STAR(r) => 1 + size(r) 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   226
// }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   227
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   228
// randomly subtracts 1 or 2 from the STAR case
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   229
def size_faulty(r: Rexp) : Int = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   230
  case ZERO => 1
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   231
  case ONE => 1
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   232
  case CHAR(_) => 1
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   233
  case ALTS(r1, r2) => 1 + size_faulty(r1) + size_faulty(r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   234
  case SEQ(r1, r2) =>  1 + size_faulty(r1) + size_faulty(r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   235
  case STAR(r) => 1 + size_faulty(r) - range(0, 2).generate
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   236
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   237
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   238
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   239
   
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   240
// some convenience for typing in regular expressions
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   241
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   242
def charlist2rexp(s : List[Char]): Rexp = s match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   243
  case Nil => ONE
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   244
  case c::Nil => CHAR(c)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   245
  case c::s => SEQ(CHAR(c), charlist2rexp(s))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   246
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   247
implicit def string2rexp(s : String) : Rexp = 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   248
  charlist2rexp(s.toList)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   249
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   250
implicit def RexpOps(r: Rexp) = new {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   251
  def | (s: Rexp) = ALTS(r, s)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   252
  def % = STAR(r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   253
  def ~ (s: Rexp) = SEQ(r, s)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   254
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   255
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   256
implicit def stringOps(s: String) = new {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   257
  def | (r: Rexp) = ALTS(s, r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   258
  def | (r: String) = ALTS(s, r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   259
  def % = STAR(s)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   260
  def ~ (r: Rexp) = SEQ(s, r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   261
  def ~ (r: String) = SEQ(s, r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   262
  def $ (r: Rexp) = RECD(s, r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   263
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   264
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   265
def nullable(r: Rexp) : Boolean = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   266
  case ZERO => false
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   267
  case ONE => true
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   268
  case CHAR(_) => false
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   269
  case ANYCHAR => false
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   270
  case ALTS(r1, r2) => nullable(r1) || nullable(r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   271
  case SEQ(r1, r2) => nullable(r1) && nullable(r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   272
  case STAR(_) => true
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   273
  case RECD(_, r1) => nullable(r1)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   274
  case NTIMES(n, r) => if (n == 0) true else nullable(r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   275
  case OPTIONAL(r) => true
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   276
  case NOT(r) => !nullable(r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   277
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   278
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   279
def der(c: Char, r: Rexp) : Rexp = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   280
  case ZERO => ZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   281
  case ONE => ZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   282
  case CHAR(d) => if (c == d) ONE else ZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   283
  case ANYCHAR => ONE 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   284
  case ALTS(r1, r2) => ALTS(der(c, r1), der(c, r2))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   285
  case SEQ(r1, r2) => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   286
    if (nullable(r1)) ALTS(SEQ(der(c, r1), r2), der(c, r2))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   287
    else SEQ(der(c, r1), r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   288
  case STAR(r) => SEQ(der(c, r), STAR(r))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   289
  case RECD(_, r1) => der(c, r1)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   290
  case NTIMES(n, r) => if(n > 0) SEQ(der(c, r), NTIMES(n - 1, r)) else ZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   291
  case OPTIONAL(r) => der(c, r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   292
  case NOT(r) =>  NOT(der(c, r))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   293
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   294
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   295
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   296
// extracts a string from a value
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   297
def flatten(v: Val) : String = v match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   298
  case Empty => ""
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   299
  case Chr(c) => c.toString
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   300
  case Left(v) => flatten(v)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   301
  case Right(v) => flatten(v)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   302
  case Sequ(v1, v2) => flatten(v1) ++ flatten(v2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   303
  case Stars(vs) => vs.map(flatten).mkString
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   304
  case Ntime(vs) => vs.map(flatten).mkString
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   305
  case Optionall(v) => flatten(v)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   306
  case Rec(_, v) => flatten(v)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   307
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   308
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   309
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   310
// extracts an environment from a value;
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   311
// used for tokenising a string
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   312
def env(v: Val) : List[(String, String)] = v match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   313
  case Empty => Nil
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   314
  case Chr(c) => Nil
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   315
  case Left(v) => env(v)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   316
  case Right(v) => env(v)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   317
  case Sequ(v1, v2) => env(v1) ::: env(v2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   318
  case Stars(vs) => vs.flatMap(env)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   319
  case Ntime(vs) => vs.flatMap(env)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   320
  case Rec(x, v) => (x, flatten(v))::env(v)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   321
  case Optionall(v) => env(v)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   322
  case Nots(s) => ("Negative", s) :: Nil
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   323
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   324
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   325
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   326
// The injection and mkeps part of the lexer
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   327
//===========================================
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   328
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   329
def mkeps(r: Rexp) : Val = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   330
  case ONE => Empty
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   331
  case ALTS(r1, r2) => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   332
    if (nullable(r1)) Left(mkeps(r1)) else Right(mkeps(r2))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   333
  case SEQ(r1, r2) => Sequ(mkeps(r1), mkeps(r2))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   334
  case STAR(r) => Stars(Nil)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   335
  case RECD(x, r) => Rec(x, mkeps(r))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   336
  case NTIMES(n, r) => Ntime(List.fill(n)(mkeps(r)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   337
  case OPTIONAL(r) => Optionall(Empty)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   338
  case NOT(rInner) => if(nullable(rInner)) throw new Exception("error")  
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   339
                         else Nots("")//Nots(s.reverse.toString)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   340
//   case NOT(ZERO) => Empty
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   341
//   case NOT(CHAR(c)) => Empty
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   342
//   case NOT(SEQ(r1, r2)) => Sequ(mkeps(NOT(r1)), mkeps(NOT(r2)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   343
//   case NOT(ALTS(r1, r2)) => if(!nullable(r1)) Left(mkeps(NOT(r1))) else Right(mkeps(NOT(r2)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   344
//   case NOT(STAR(r)) => Stars(Nil) 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   345
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   346
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   347
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   348
def inj(r: Rexp, c: Char, v: Val) : Val = (r, v) match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   349
  case (STAR(r), Sequ(v1, Stars(vs))) => Stars(inj(r, c, v1)::vs)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   350
  case (SEQ(r1, r2), Sequ(v1, v2)) => Sequ(inj(r1, c, v1), v2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   351
  case (SEQ(r1, r2), Left(Sequ(v1, v2))) => Sequ(inj(r1, c, v1), v2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   352
  case (SEQ(r1, r2), Right(v2)) => Sequ(mkeps(r1), inj(r2, c, v2))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   353
  case (ALTS(r1, r2), Left(v1)) => Left(inj(r1, c, v1))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   354
  case (ALTS(r1, r2), Right(v2)) => Right(inj(r2, c, v2))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   355
  case (CHAR(d), Empty) => Chr(c) 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   356
  case (RECD(x, r1), _) => Rec(x, inj(r1, c, v))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   357
  case (NTIMES(n, r), Sequ(v1, Ntime(vs))) => Ntime(inj(r, c, v1)::vs)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   358
  case (OPTIONAL(r), v) => Optionall(inj(r, c, v))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   359
  case (NOT(r), Nots(s)) => Nots(c.toString ++ s)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   360
  case (ANYCHAR, Empty) => Chr(c)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   361
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   362
666
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   363
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   364
def print_lexer(r: Rexp, s: List[Char]): Val = s match{
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   365
  case Nil => {println("mkeps!"); println(r); val v = mkeps(r); println(v); v}
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   366
  case c::cs => { 
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   367
    val derc = der(c, r);
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   368
    println("derivative is: "+ derc);
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   369
    val v = print_lexer(derc, cs); println("Injection!");println("r: "+r); println("c: "+c); println("v:"+v); inj(r, c, v ) }
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   370
}
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   371
val r = STAR("a"|"aa") ~ ("c")
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   372
val v = print_lexer(r, "aac".toList)
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   373
println(v)
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   374
517
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   375
// some "rectification" functions for simplification
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   376
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   377
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   378
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   379
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   380
// The Lexing Rules for the WHILE Language
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   381
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   382
  // bnullable function: tests whether the aregular 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   383
  // expression can recognise the empty string
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   384
def bnullable (r: ARexp) : Boolean = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   385
    case AZERO => false
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   386
    case AONE(_) => true
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   387
    case ACHAR(_,_) => false
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   388
    case AALTS(_, rs) => rs.exists(bnullable)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   389
    case ASEQ(_, r1, r2) => bnullable(r1) && bnullable(r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   390
    case ASTAR(_, _) => true
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   391
    case ANOT(_, rn) => !bnullable(rn)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   392
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   393
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   394
def mkepsBC(r: ARexp) : Bits = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   395
    case AONE(bs) => bs
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   396
    case AALTS(bs, rs) => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   397
      val n = rs.indexWhere(bnullable)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   398
      bs ++ mkepsBC(rs(n))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   399
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   400
    case ASEQ(bs, r1, r2) => bs ++ mkepsBC(r1) ++ mkepsBC(r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   401
    case ASTAR(bs, r) => bs ++ List(Z)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   402
    case ANOT(bs, rn) => bs
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   403
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   404
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   405
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   406
def bder(c: Char, r: ARexp) : ARexp = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   407
    case AZERO => AZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   408
    case AONE(_) => AZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   409
    case ACHAR(bs, f) => if (c == f) AONE(bs) else AZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   410
    case AALTS(bs, rs) => AALTS(bs, rs.map(bder(c, _)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   411
    case ASEQ(bs, r1, r2) => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   412
      if (bnullable(r1)) AALTS(bs, ASEQ(Nil, bder(c, r1), r2) :: fuse(mkepsBC(r1), bder(c, r2)) :: Nil )
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   413
      else ASEQ(bs, bder(c, r1), r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   414
    case ASTAR(bs, r) => ASEQ(bs, fuse(List(S), bder(c, r)), ASTAR(Nil, r))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   415
    case ANOT(bs, rn) => ANOT(bs, bder(c, rn))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   416
    case AANYCHAR(bs) => AONE(bs)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   417
  } 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   418
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   419
def fuse(bs: Bits, r: ARexp) : ARexp = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   420
    case AZERO => AZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   421
    case AONE(cs) => AONE(bs ++ cs)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   422
    case ACHAR(cs, f) => ACHAR(bs ++ cs, f)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   423
    case AALTS(cs, rs) => AALTS(bs ++ cs, rs)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   424
    case ASEQ(cs, r1, r2) => ASEQ(bs ++ cs, r1, r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   425
    case ASTAR(cs, r) => ASTAR(bs ++ cs, r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   426
    case ANOT(cs, r) => ANOT(bs ++ cs, r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   427
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   428
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   429
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   430
def internalise(r: Rexp) : ARexp = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   431
    case ZERO => AZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   432
    case ONE => AONE(Nil)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   433
    case CHAR(c) => ACHAR(Nil, c)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   434
    //case PRED(f) => APRED(Nil, f)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   435
    case ALTS(r1, r2) => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   436
      AALTS(Nil, List(fuse(List(Z), internalise(r1)), fuse(List(S), internalise(r2))))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   437
    // case ALTS(r1::rs) => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   438
    //   val AALTS(Nil, rs2) = internalise(ALTS(rs))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   439
    //   AALTS(Nil, fuse(List(Z), internalise(r1)) :: rs2.map(fuse(List(S), _)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   440
    // }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   441
    case SEQ(r1, r2) => ASEQ(Nil, internalise(r1), internalise(r2))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   442
    case STAR(r) => ASTAR(Nil, internalise(r))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   443
    case RECD(x, r) => internalise(r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   444
    case NOT(r) => ANOT(Nil, internalise(r))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   445
    case ANYCHAR => AANYCHAR(Nil)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   446
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   447
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   448
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   449
def bsimp(r: ARexp): ARexp = 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   450
  {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   451
    r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   452
      case ASEQ(bs1, r1, r2) => (bsimp(r1), bsimp(r2)) match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   453
          case (AZERO, _) => AZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   454
          case (_, AZERO) => AZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   455
          case (AONE(bs2), r2s) => fuse(bs1 ++ bs2, r2s)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   456
          case (r1s, r2s) => ASEQ(bs1, r1s, r2s)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   457
      }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   458
      case AALTS(bs1, rs) => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   459
            val rs_simp = rs.map(bsimp(_))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   460
            val flat_res = flats(rs_simp)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   461
            val dist_res = distinctBy(flat_res, erase)//strongDB(flat_res)//distinctBy(flat_res, erase)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   462
            dist_res match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   463
              case Nil => AZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   464
              case s :: Nil => fuse(bs1, s)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   465
              case rs => AALTS(bs1, rs)  
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   466
            }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   467
          
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   468
      }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   469
      case r => r
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   470
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   471
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   472
def strongBsimp(r: ARexp): ARexp =
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   473
{
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   474
  r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   475
    case ASEQ(bs1, r1, r2) => (strongBsimp(r1), strongBsimp(r2)) match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   476
        case (AZERO, _) => AZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   477
        case (_, AZERO) => AZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   478
        case (AONE(bs2), r2s) => fuse(bs1 ++ bs2, r2s)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   479
        case (r1s, r2s) => ASEQ(bs1, r1s, r2s)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   480
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   481
    case AALTS(bs1, rs) => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   482
          val rs_simp = rs.map(strongBsimp(_))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   483
          val flat_res = flats(rs_simp)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   484
          val dist_res = distinctBy4(flat_res)//distinctBy(flat_res, erase)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   485
          dist_res match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   486
            case Nil => AZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   487
            case s :: Nil => fuse(bs1, s)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   488
            case rs => AALTS(bs1, rs)  
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   489
          }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   490
        
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   491
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   492
    case r => r
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   493
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   494
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   495
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   496
def strongBsimp5(r: ARexp): ARexp =
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   497
{
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   498
  // println("was this called?")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   499
  r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   500
    case ASEQ(bs1, r1, r2) => (strongBsimp5(r1), strongBsimp5(r2)) match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   501
        case (AZERO, _) => AZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   502
        case (_, AZERO) => AZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   503
        case (AONE(bs2), r2s) => fuse(bs1 ++ bs2, r2s)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   504
        case (r1s, r2s) => ASEQ(bs1, r1s, r2s)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   505
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   506
    case AALTS(bs1, rs) => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   507
        // println("alts case")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   508
          val rs_simp = rs.map(strongBsimp5(_))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   509
          val flat_res = flats(rs_simp)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   510
          var dist_res = distinctBy5(flat_res)//distinctBy(flat_res, erase)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   511
          var dist2_res = distinctBy5(dist_res)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   512
          while(dist_res != dist2_res){
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   513
            dist_res = dist2_res
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   514
            dist2_res = distinctBy5(dist_res)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   515
          }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   516
          (dist2_res) match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   517
            case Nil => AZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   518
            case s :: Nil => fuse(bs1, s)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   519
            case rs => AALTS(bs1, rs)  
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   520
          }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   521
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   522
    case r => r
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   523
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   524
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   525
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   526
def bders (s: List[Char], r: ARexp) : ARexp = s match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   527
  case Nil => r
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   528
  case c::s => bders(s, bder(c, r))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   529
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   530
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   531
def flats(rs: List[ARexp]): List[ARexp] = rs match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   532
    case Nil => Nil
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   533
    case AZERO :: rs1 => flats(rs1)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   534
    case AALTS(bs, rs1) :: rs2 => rs1.map(fuse(bs, _)) ::: flats(rs2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   535
    case r1 :: rs2 => r1 :: flats(rs2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   536
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   537
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   538
def distinctBy[B, C](xs: List[B], f: B => C, acc: List[C] = Nil): List[B] = xs match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   539
  case Nil => Nil
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   540
  case (x::xs) => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   541
    val res = f(x)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   542
    if (acc.contains(res)) distinctBy(xs, f, acc)  
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   543
    else x::distinctBy(xs, f, res::acc)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   544
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   545
} 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   546
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   547
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   548
def pruneRexp(r: ARexp, allowableTerms: List[Rexp]) : ARexp = {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   549
  r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   550
    case ASEQ(bs, r1, r2) => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   551
      val termsTruncated = allowableTerms.collect(rt => rt match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   552
        case SEQ(r1p, r2p) if(r2p == erase(r2)) => r1p//if(r2p == erase(r2)) 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   553
      })
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   554
      val pruned : ARexp = pruneRexp(r1, termsTruncated)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   555
      pruned match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   556
        case AZERO => AZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   557
        case AONE(bs1) => fuse(bs ++ bs1, r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   558
        case pruned1 => ASEQ(bs, pruned1, r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   559
      }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   560
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   561
    case AALTS(bs, rs) => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   562
      //allowableTerms.foreach(a => println(shortRexpOutput(a)))        
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   563
      val rsp = rs.map(r => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   564
                    pruneRexp(r, allowableTerms)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   565
                  )
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   566
                  .filter(r => r != AZERO)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   567
      rsp match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   568
        case Nil => AZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   569
        case r1::Nil => fuse(bs, r1)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   570
        case rs1 => AALTS(bs, rs1)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   571
      }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   572
    case r => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   573
      if(allowableTerms.contains(erase(r))) r else AZERO //assert(r != AZERO)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   574
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   575
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   576
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   577
def oneSimp(r: Rexp) : Rexp = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   578
  case SEQ(ONE, r) => r
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   579
  case SEQ(r1, r2) => SEQ(oneSimp(r1), r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   580
  case r => r//assert r != 0 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   581
    
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   582
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   583
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   584
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   585
def distinctBy4(xs: List[ARexp], acc: List[Rexp] = Nil) : List[ARexp] = xs match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   586
  case Nil => Nil
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   587
  case x :: xs =>
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   588
    //assert(acc.distinct == acc)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   589
    val erased = erase(x)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   590
    if(acc.contains(erased))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   591
      distinctBy4(xs, acc)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   592
    else{
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   593
      val addToAcc =  breakIntoTerms(erased).filter(r => !acc.contains(oneSimp(r)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   594
      //val xp = pruneRexp(x, addToAcc)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   595
      pruneRexp(x, addToAcc) match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   596
        case AZERO => distinctBy4(xs, addToAcc.map(oneSimp(_)) ::: acc)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   597
        case xPrime => xPrime :: distinctBy4(xs, addToAcc.map(oneSimp(_)) ::: acc)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   598
      }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   599
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   600
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   601
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   602
// fun pAKC_aux :: "arexp list ⇒ arexp ⇒ rexp ⇒ arexp"
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   603
//   where
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   604
// "pAKC_aux rsa r ctx = (if (L (SEQ (erase r) ( ctx) )) ⊆ (L (erase (AALTs [] rsa))) then AZERO else
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   605
//                           case r of (ASEQ bs r1 r2) ⇒ 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   606
//                             bsimp_ASEQ bs (pAKC_aux rsa r1 (SEQ  (erase r2) ( ctx) )) r2   |
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   607
//                                     (AALTs bs rs) ⇒ 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   608
//                             bsimp_AALTs bs (flts (map (λr. pAKC_aux rsa r ( ctx) ) rs) )    |
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   609
//                                     r             ⇒ r
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   610
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   611
// def canonicalSeq(rs: List[Rexp], acc: Rexp) = rs match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   612
//   case r::Nil => SEQ(r, acc)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   613
//   case Nil => acc
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   614
//   case r1::r2::Nil => SEQ(SEQ(r1, r2), acc)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   615
// }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   616
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   617
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   618
//the "fake" Language interpretation: just concatenates!
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   619
def L(acc: Rexp, rs: List[Rexp]) : Rexp = rs match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   620
  case Nil => acc
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   621
  case r :: rs1 => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   622
    // if(acc == ONE) 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   623
    //   L(r, rs1) 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   624
    // else
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   625
      L(SEQ(acc, r), rs1)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   626
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   627
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   628
def rprint(r: Rexp) : Unit = println(shortRexpOutput(r))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   629
def rsprint(rs: List[Rexp]) = rs.foreach(r => println(shortRexpOutput(r)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   630
def aprint(a: ARexp) = println(shortRexpOutput(erase(a)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   631
def asprint(as: List[ARexp]) = as.foreach(a => println(shortRexpOutput(erase(a))))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   632
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   633
def pAKC(acc: List[Rexp], r: ARexp, ctx: List[Rexp]) : ARexp = {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   634
  // println("pakc")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   635
  // println(shortRexpOutput(erase(r)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   636
  // println("acc")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   637
  // rsprint(acc)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   638
  // println("ctx---------")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   639
  // rsprint(ctx)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   640
  // println("ctx---------end")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   641
  // rsprint(breakIntoTerms(L(erase(r), ctx)).map(oneSimp))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   642
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   643
  if (breakIntoTerms(L(erase(r), ctx)).map(oneSimp).forall(acc.contains)) {//acc.flatMap(breakIntoTerms
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   644
    AZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   645
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   646
  else{
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   647
    r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   648
      case ASEQ(bs, r1, r2) => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   649
      (pAKC(acc, r1, erase(r2) :: ctx)) match{
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   650
        case AZERO => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   651
          AZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   652
        case AONE(bs1) => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   653
          fuse(bs1, r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   654
        case r1p => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   655
          ASEQ(bs, r1p, r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   656
      }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   657
      case AALTS(bs, rs0) => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   658
        // println("before pruning")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   659
        // println(s"ctx is ")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   660
        // ctx.foreach(r => println(shortRexpOutput(r)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   661
        // println(s"rs0 is ")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   662
        // rs0.foreach(r => println(shortRexpOutput(erase(r))))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   663
        // println(s"acc is ")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   664
        // acc.foreach(r => println(shortRexpOutput(r)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   665
        rs0.map(r => pAKC(acc, r, ctx)).filter(_ != AZERO) match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   666
          case Nil => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   667
            // println("after pruning Nil")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   668
            AZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   669
          case r :: Nil => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   670
            // println("after pruning singleton")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   671
            // println(shortRexpOutput(erase(r)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   672
            r 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   673
          case rs0p => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   674
          // println("after pruning non-singleton")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   675
            AALTS(bs, rs0p)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   676
        }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   677
      case r => r
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   678
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   679
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   680
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   681
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   682
def distinctBy5(xs: List[ARexp], acc: List[Rexp] = Nil) : List[ARexp] = xs match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   683
  case Nil => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   684
    Nil
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   685
  case x :: xs => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   686
    val erased = erase(x)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   687
    if(acc.contains(erased)){
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   688
      distinctBy5(xs, acc)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   689
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   690
    else{
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   691
      val pruned = pAKC(acc, x, Nil)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   692
      val newTerms = breakIntoTerms(erase(pruned))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   693
      pruned match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   694
        case AZERO => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   695
          distinctBy5(xs, acc)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   696
        case xPrime => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   697
          xPrime :: distinctBy5(xs, newTerms.map(oneSimp) ::: acc)//distinctBy5(xs, addToAcc.map(oneSimp(_)) ::: acc)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   698
      }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   699
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   700
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   701
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   702
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   703
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   704
def breakIntoTerms(r: Rexp) : List[Rexp] = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   705
  case SEQ(r1, r2)  => breakIntoTerms(r1).map(r11 => SEQ(r11, r2))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   706
  case ALTS(r1, r2) => breakIntoTerms(r1) ::: breakIntoTerms(r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   707
  case ZERO => Nil
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   708
  case _ => r::Nil
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   709
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   710
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   711
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   712
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   713
def decode_aux(r: Rexp, bs: Bits) : (Val, Bits) = (r, bs) match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   714
  case (ONE, bs) => (Empty, bs)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   715
  case (CHAR(f), bs) => (Chr(f), bs)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   716
  case (ALTS(r1, r2), Z::bs1) => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   717
      val (v, bs2) = decode_aux(r1, bs1)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   718
      (Left(v), bs2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   719
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   720
  case (ALTS(r1, r2), S::bs1) => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   721
      val (v, bs2) = decode_aux(r2, bs1)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   722
      (Right(v), bs2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   723
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   724
  case (SEQ(r1, r2), bs) => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   725
    val (v1, bs1) = decode_aux(r1, bs)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   726
    val (v2, bs2) = decode_aux(r2, bs1)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   727
    (Sequ(v1, v2), bs2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   728
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   729
  case (STAR(r1), S::bs) => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   730
    val (v, bs1) = decode_aux(r1, bs)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   731
    //(v)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   732
    val (Stars(vs), bs2) = decode_aux(STAR(r1), bs1)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   733
    (Stars(v::vs), bs2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   734
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   735
  case (STAR(_), Z::bs) => (Stars(Nil), bs)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   736
  case (RECD(x, r1), bs) => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   737
    val (v, bs1) = decode_aux(r1, bs)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   738
    (Rec(x, v), bs1)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   739
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   740
  case (NOT(r), bs) => (Nots(r.toString), bs)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   741
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   742
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   743
def decode(r: Rexp, bs: Bits) = decode_aux(r, bs) match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   744
  case (v, Nil) => v
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   745
  case _ => throw new Exception("Not decodable")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   746
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   747
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   748
666
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   749
def blexing(r: Rexp, s: String) : Option[Val] = {
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   750
    val bit_code = blex(internalise(r), s.toList)
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   751
    bit_code match {
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   752
      case Some(bits) => Some(decode(r, bits))
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   753
      case None => None
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   754
    }
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   755
}
517
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   756
666
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   757
def blexing_simp(r: Rexp, s: String) : Option[Val] = {
517
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   758
    val bit_code = blex_simp(internalise(r), s.toList)
666
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   759
    bit_code match {
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   760
      case Some(bits) => Some(decode(r, bits))
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   761
      case None => None
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   762
    }
517
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   763
}
666
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   764
//def simpBlexer(r: Rexp, s: String) : Val = {
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   765
//  Try(blexing_simp(r, s)).getOrElse(Failure)
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   766
//}
517
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   767
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   768
def strong_blexing_simp(r: Rexp, s: String) : Val = {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   769
  decode(r, strong_blex_simp(internalise(r), s.toList))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   770
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   771
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   772
def strong_blexing_simp5(r: Rexp, s: String) : Val = {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   773
  decode(r, strong_blex_simp5(internalise(r), s.toList))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   774
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   775
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   776
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   777
def strongBlexer(r: Rexp, s: String) : Val = {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   778
  Try(strong_blexing_simp(r, s)).getOrElse(Failure)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   779
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   780
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   781
def strongBlexer5(r: Rexp, s: String): Val = {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   782
  Try(strong_blexing_simp5(r, s)).getOrElse(Failure)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   783
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   784
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   785
def strong_blex_simp(r: ARexp, s: List[Char]) : Bits = s match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   786
  case Nil => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   787
    if (bnullable(r)) {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   788
      //println(asize(r))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   789
      val bits = mkepsBC(r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   790
      bits
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   791
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   792
    else 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   793
      throw new Exception("Not matched")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   794
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   795
  case c::cs => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   796
    val der_res = bder(c,r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   797
    val simp_res = strongBsimp(der_res)  
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   798
    strong_blex_simp(simp_res, cs)      
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   799
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   800
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   801
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   802
def strong_blex_simp5(r: ARexp, s: List[Char]) : Bits = s match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   803
  case Nil => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   804
    if (bnullable(r)) {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   805
      //println(asize(r))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   806
      val bits = mkepsBC(r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   807
      bits
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   808
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   809
    else 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   810
      throw new Exception("Not matched")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   811
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   812
  case c::cs => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   813
    val der_res = bder(c,r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   814
    val simp_res = strongBsimp5(der_res)  
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   815
    strong_blex_simp5(simp_res, cs)      
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   816
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   817
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   818
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   819
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   820
  def bders_simp(s: List[Char], r: ARexp) : ARexp = s match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   821
    case Nil => r
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   822
    case c::s => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   823
      //println(erase(r))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   824
      bders_simp(s, bsimp(bder(c, r)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   825
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   826
  
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   827
  def bdersStrong5(s: List[Char], r: ARexp) : ARexp = s match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   828
    case Nil => r
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   829
    case c::s => bdersStrong5(s, strongBsimp5(bder(c, r)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   830
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   831
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   832
  def bdersSimp(s: String, r: Rexp) : ARexp = bders_simp(s.toList, internalise(r))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   833
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   834
  def bdersStrong(s: List[Char], r: ARexp) : ARexp = s match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   835
    case Nil => r 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   836
    case c::s => bdersStrong(s, strongBsimp(bder(c, r)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   837
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   838
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   839
  def bdersStrongRexp(s: String, r: Rexp) : ARexp = bdersStrong(s.toList, internalise(r))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   840
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   841
  def erase(r:ARexp): Rexp = r match{
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   842
    case AZERO => ZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   843
    case AONE(_) => ONE
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   844
    case ACHAR(bs, c) => CHAR(c)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   845
    case AALTS(bs, Nil) => ZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   846
    case AALTS(bs, a::Nil) => erase(a)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   847
    case AALTS(bs, a::as) => ALTS(erase(a), erase(AALTS(bs, as)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   848
    case ASEQ(bs, r1, r2) => SEQ (erase(r1), erase(r2))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   849
    case ASTAR(cs, r)=> STAR(erase(r))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   850
    case ANOT(bs, r) => NOT(erase(r))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   851
    case AANYCHAR(bs) => ANYCHAR
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   852
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   853
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   854
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   855
  def allCharSeq(r: Rexp) : Boolean = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   856
    case CHAR(c) => true
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   857
    case SEQ(r1, r2) => allCharSeq(r1) && allCharSeq(r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   858
    case _ => false
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   859
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   860
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   861
  def flattenSeq(r: Rexp) : String = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   862
    case CHAR(c) => c.toString
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   863
    case SEQ(r1, r2) => flattenSeq(r1) ++ flattenSeq(r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   864
    case _ => throw new Error("flatten unflattenable rexp")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   865
  } 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   866
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   867
  def shortRexpOutput(r: Rexp) : String = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   868
      case CHAR(c) => c.toString
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   869
      case ONE => "1"
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   870
      case ZERO => "0"
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   871
      case SEQ(r1, r2) if(allCharSeq(r)) => flattenSeq(r)//"[" ++ shortRexpOutput(r1) ++ "~" ++ shortRexpOutput(r2) ++ "]"
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   872
      case SEQ(r1, r2) => "[" ++ shortRexpOutput(r1) ++ "~" ++ shortRexpOutput(r2) ++ "]"
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   873
      case ALTS(r1, r2) => "(" ++ shortRexpOutput(r1) ++ "+" ++ shortRexpOutput(r2) ++ ")"
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   874
      case STAR(STAR(r)) => "(..)*"// ++ shortRexpOutput(r) ++ "]*"
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   875
      case STAR(r) => "STAR(" ++ shortRexpOutput(r) ++ ")"
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   876
      //case RTOP => "RTOP"
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   877
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   878
666
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   879
  def blex(r: ARexp, s: List[Char]) : Option[Bits]= s match {
517
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   880
      case Nil => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   881
        if (bnullable(r)) {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   882
          val bits = mkepsBC(r)
666
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   883
          Some(bits)
517
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   884
        }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   885
        else 
666
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   886
          None
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   887
          //throw new Exception("Not matched")
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   888
      }
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   889
      case c::cs => {
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   890
        val der_res = bder(c,r)
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   891
        //val simp_res = bsimp(der_res)  
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   892
        blex(der_res, cs)      
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   893
      }
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   894
  }
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   895
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   896
  def blex_simp(r: ARexp, s: List[Char]) : Option[Bits]= s match {
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   897
      case Nil => {
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   898
        if (bnullable(r)) {
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   899
          val bits = mkepsBC(r)
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   900
          Some(bits)
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   901
        }
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   902
        else 
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   903
          None
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
   904
          //throw new Exception("Not matched")
517
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   905
      }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   906
      case c::cs => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   907
        val der_res = bder(c,r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   908
        val simp_res = bsimp(der_res)  
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   909
        blex_simp(simp_res, cs)      
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   910
      }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   911
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   912
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   913
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   914
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   915
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   916
    def size(r: Rexp) : Int = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   917
      case ZERO => 1
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   918
      case ONE => 1
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   919
      case CHAR(_) => 1
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   920
      case ANYCHAR => 1
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   921
      case NOT(r0) => 1 + size(r0)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   922
      case SEQ(r1, r2) => 1 + size(r1) + size(r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   923
      case ALTS(r1, r2) => 1 + List(r1, r2).map(size).sum
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   924
      case STAR(r) => 1 + size(r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   925
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   926
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   927
    def asize(a: ARexp) = size(erase(a))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   928
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   929
//pder related
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   930
type Mon = (Char, Rexp)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   931
type Lin = Set[Mon]
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   932
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   933
def dot_prod(rs: Set[Rexp], r: Rexp): Set[Rexp] = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   934
    case ZERO => Set()
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   935
    case ONE => rs
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   936
    case r => rs.map((re) => if (re == ONE) r else SEQ(re, r)  )   
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   937
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   938
  def cir_prod(l: Lin, t: Rexp): Lin = t match {//remember this Lin is different from the Lin in Antimirov's paper. Here it does not mean the set of all subsets of linear forms that does not contain zero, but rather the type a set of linear forms
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   939
    case ZERO => Set()
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   940
    case ONE => l
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   941
    case t => l.map( m => m._2 match {case ZERO => m case ONE => (m._1, t) case p => (m._1, SEQ(p, t)) }  )
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   942
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   943
  def lf(r: Rexp): Lin = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   944
    case ZERO => Set()
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   945
    case ONE => Set()
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   946
    case CHAR(f) => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   947
      //val Some(c) = alphabet.find(f) 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   948
      Set((f, ONE))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   949
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   950
    case ALTS(r1, r2) => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   951
      lf(r1 ) ++ lf(r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   952
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   953
    case STAR(r1) => cir_prod(lf(r1), STAR(r1)) //may try infix notation later......
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   954
    case SEQ(r1, r2) =>{
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   955
      if (nullable(r1))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   956
        cir_prod(lf(r1), r2) ++ lf(r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   957
      else
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   958
        cir_prod(lf(r1), r2) 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   959
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   960
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   961
  def lfs(r: Set[Rexp]): Lin = {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   962
    r.foldLeft(Set[Mon]())((acc, r) => acc ++ lf(r))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   963
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   964
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   965
  def pder(x: Char, t: Rexp): Set[Rexp] = {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   966
    val lft = lf(t)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   967
    (lft.filter(mon => if(mon._1 == x) true else false)).map(mon => mon._2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   968
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   969
  def pders_single(s: List[Char], t: Rexp) : Set[Rexp] = s match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   970
    case x::xs => pders(xs, pder(x, t))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   971
    case Nil => Set(t)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   972
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   973
  def pders(s: List[Char], ts: Set[Rexp]) : Set[Rexp] = s match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   974
    case x::xs => pders(xs, ts.foldLeft(Set[Rexp]())((acc, t) => acc ++ pder(x, t)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   975
    case Nil => ts 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   976
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   977
  def pderss(ss: List[List[Char]], t: Rexp): Set[Rexp] = ss.foldLeft( Set[Rexp]() )( (acc, s) => pders_single(s, t) ++ acc )
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   978
  def pdera(t: Rexp): Set[Rexp] = lf(t).map(mon => mon._2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   979
  //all implementation of partial derivatives that involve set union are potentially buggy
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   980
  //because they don't include the original regular term before they are pdered.
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   981
  //now only pderas is fixed.
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   982
  def pderas(t: Set[Rexp], d: Int): Set[Rexp] = if(d > 0) pderas(lfs(t).map(mon => mon._2), d - 1) ++ t else lfs(t).map(mon => mon._2) ++ t//repeated application of pderas over the newest set of pders.
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   983
  def pderUNIV(r: Rexp) : Set[Rexp] = pderas(Set(r), awidth(r) + 1)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   984
  def awidth(r: Rexp) : Int = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   985
    case CHAR(c) => 1
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   986
    case SEQ(r1, r2) => awidth(r1) + awidth(r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   987
    case ALTS(r1, r2) => awidth(r1) + awidth(r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   988
    case ONE => 0
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   989
    case ZERO => 0
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   990
    case STAR(r) => awidth(r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   991
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   992
  //def sigma_lf(l: Set[Mon]) : Rexp = ALTS(l.map(mon => SEQ(CHAR(mon._1),mon._2)).toList)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   993
  //def sigma(rs: Set[Rexp]) : Rexp = ALTS(rs.toList)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   994
  def o(r: Rexp) = if (nullable(r)) ONE else ZERO
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   995
  //def nlf(t: Rexp) : Rexp = ALTS(List( o(t), sigma_lf(lf(t)) ))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   996
  def pdp(x: Char, r: Rexp) : Set[Rexp] = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   997
    case ZERO => Set[Rexp]()
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   998
    case ONE => Set[Rexp]()
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
   999
    case CHAR(f) => if(x == f) Set(ONE) else Set[Rexp]()
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1000
    case ALTS(r1, r2) => pdp(x, r1) ++ pdp(x, r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1001
    case STAR(r1) => pdp(x, r).map(a => SEQ(a, STAR(r1)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1002
    case SEQ(a0, b) => if(nullable(a0)) pdp(x, a0).map(a => SEQ(a, b)) ++ pdp(x, b) else pdp(x, a0).map(a => SEQ(a, b))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1003
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1004
  def pdps(s: List[Char], ts: Set[Rexp]): Set[Rexp] = s match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1005
    case x::xs => pdps(xs, ts.foldLeft(Set[Rexp]())((acc, t) => acc ++ pder(x, t)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1006
    case Nil => ts   
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1007
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1008
  def pdpss(ss: List[List[Char]], t: Rexp): Set[Rexp] = ss.foldLeft( Set[Rexp]())((acc, s) => pdps(s, Set(t)) ++ acc)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1009
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1010
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1011
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1012
def starPrint(r: Rexp) : Unit = r match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1013
        
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1014
          case SEQ(head, rstar) =>
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1015
            println(shortRexpOutput(head) ++ "~STARREG")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1016
          case STAR(rstar) =>
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1017
            println("STARREG")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1018
          case ALTS(r1, r2) =>  
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1019
            println("(")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1020
            starPrint(r1)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1021
            println("+")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1022
            starPrint(r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1023
            println(")")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1024
          case ZERO => println("0")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1025
      }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1026
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1027
// @arg(doc = "small tests")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1028
def n_astar_list(d: Int) : Rexp = {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1029
  if(d == 0) STAR("a") 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1030
  else ALTS(STAR("a" * d), n_astar_list(d - 1))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1031
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1032
def n_astar_alts(d: Int) : Rexp = d match {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1033
  case 0 => n_astar_list(0)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1034
  case d => STAR(n_astar_list(d))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1035
  // case r1 :: r2 :: Nil => ALTS(r1, r2)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1036
  // case r1 :: Nil => r1
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1037
  // case r :: rs => ALTS(r, n_astar_alts(rs))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1038
  // case Nil => throw new Error("should give at least 1 elem")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1039
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1040
def n_astar_aux(d: Int) = {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1041
  if(d == 0) n_astar_alts(0)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1042
  else ALTS(n_astar_alts(d), n_astar_alts(d - 1))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1043
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1044
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1045
def n_astar(d: Int) : Rexp = STAR(n_astar_aux(d))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1046
//val STARREG = n_astar(3)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1047
// ( STAR("a") | 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1048
//                  ("a" | "aa").% | 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1049
//                 ( "a" | "aa" | "aaa").% 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1050
//                 ).%
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1051
                //( "a" | "aa" | "aaa" | "aaaa").% |
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1052
                //( "a" | "aa" | "aaa" | "aaaa" | "aaaaa").% 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1053
(((STAR("a") | ( STAR("aaa")) | STAR("aaaaa"| ( STAR("aaaaaaa")) | STAR("aaaaaaaaaaa"))).%).%).%
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1054
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1055
// @main
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1056
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1057
def lcm(list: Seq[Int]):Int=list.foldLeft(1:Int){
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1058
  (a, b) => b * a /
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1059
  Stream.iterate((a,b)){case (x,y) => (y, x%y)}.dropWhile(_._2 != 0).head._1.abs
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1060
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1061
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1062
def small() = {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1063
  //val pderSTAR = pderUNIV(STARREG)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1064
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1065
  //val refSize = pderSTAR.map(size(_)).sum
518
ff7945a988a3 more to thesis
Chengsong
parents: 517
diff changeset
  1066
  for(n <- 1 to 12){
ff7945a988a3 more to thesis
Chengsong
parents: 517
diff changeset
  1067
    val STARREG = n_astar_alts(n)
517
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1068
    val iMax = (lcm((1 to n).toList))
545
333013923c5a more changes to blexersimp.thy
Chengsong
parents: 518
diff changeset
  1069
    for(i <- 1 to iMax + 1 ){// 100, 400, 800, 840, 841, 900 
517
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1070
      val prog0 = "a" * i
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1071
      //println(s"test: $prog0")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1072
      print(n)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1073
      print(" ")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1074
      // print(i)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1075
      // print(" ")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1076
      println(asize(bders_simp(prog0.toList, internalise(STARREG))))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1077
    }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1078
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1079
}
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1080
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1081
def generator_test() {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1082
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1083
  // test(rexp(7), 1000) { (r: Rexp) => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1084
  //   val ss = stringsFromRexp(r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1085
  //   val boolList = ss.map(s => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1086
  //     val simpVal = simpBlexer(r, s)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1087
  //     val strongVal = strongBlexer(r, s)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1088
  //     // println(simpVal)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1089
  //     // println(strongVal)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1090
  //     (simpVal == strongVal) && (simpVal != None) 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1091
  //   })
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1092
  //   !boolList.exists(b => b == false)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1093
  // }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1094
  // test(single(STAR(ALTS(STAR(CHAR('c')),ALTS(CHAR('c'),ZERO)))), 100000) { (r: Rexp) => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1095
  //   val ss = stringsFromRexp(r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1096
  //   val boolList = ss.map(s => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1097
  //     val bdStrong = bdersStrong(s.toList, internalise(r))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1098
  //     val bdStrong5 = bdersStrong5(s.toList, internalise(r))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1099
  //     // println(shortRexpOutput(r))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1100
  //     // println(s)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1101
  //     // println(strongVal)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1102
  //     // println(strongVal5)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1103
  //     // (strongVal == strongVal5) 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1104
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1105
  //     if(asize(bdStrong5) > asize(bdStrong)){
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1106
  //       println(s)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1107
  //       println(shortRexpOutput(erase(bdStrong5)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1108
  //       println(shortRexpOutput(erase(bdStrong)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1109
  //     }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1110
  //     asize(bdStrong5) <= asize(bdStrong)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1111
  //   })
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1112
  //   !boolList.exists(b => b == false)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1113
  // }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1114
  //*** example where bdStrong5 has a smaller size than bdStrong
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1115
  // test(single(STAR(SEQ(ALTS(SEQ(STAR(CHAR('a')),ALTS(ALTS(ONE,ZERO),SEQ(ONE,ONE))),CHAR('a')),ONE))), 1) { (r: Rexp) => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1116
  //*** depth 5 example bdStrong5 larger size than bdStrong
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1117
  // test(single(STAR(SEQ(SEQ(ALTS(CHAR('b'),STAR(CHAR('b'))),CHAR('b')),(ALTS(STAR(CHAR('c')), ONE))))), 1) {(r: Rexp) =>
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1118
 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1119
 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1120
 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1121
  //sanity check from Christian's request
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1122
  // val r = ("a" | "ab") ~ ("bc" | "c")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1123
  // val a = internalise(r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1124
  // val aval = blexing_simp(r, "abc")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1125
  // println(aval)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1126
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1127
  //sample counterexample:(depth 7)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1128
  //STAR(SEQ(ALTS(STAR(STAR(STAR(STAR(CHAR(c))))),ALTS(CHAR(c),CHAR(b))),ALTS(ZERO,SEQ(ALTS(ALTS(STAR(CHAR(c)),SEQ(CHAR(b),CHAR(a))),CHAR(c)),STAR(ALTS(ALTS(ONE,CHAR(a)),STAR(CHAR(c))))))))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1129
  //(depth5)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1130
  //STAR(SEQ(ALTS(ALTS(STAR(CHAR(b)),SEQ(ONE,CHAR(b))),SEQ(STAR(CHAR(a)),CHAR(b))),ALTS(ZERO,ALTS(STAR(CHAR(b)),STAR(CHAR(a))))))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1131
  test(rexp(4), 10000000) { (r: Rexp) => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1132
  // ALTS(SEQ(SEQ(ONE,CHAR('a')),STAR(CHAR('a'))),SEQ(ALTS(CHAR('c'),ONE),STAR(ZERO))))))), 1) { (r: Rexp) => 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1133
    val ss = stringsFromRexp(r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1134
    val boolList = ss.map(s => {
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1135
      val bdStrong = bdersStrong(s.toList, internalise(r))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1136
      val bdStrong5 = bdersStrong5(s.toList, internalise(r))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1137
      val bdFurther5 = strongBsimp5(bdStrong5)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1138
      val sizeFurther5 = asize(bdFurther5)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1139
      val pdersSet = pderUNIV(r)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1140
      val size5 = asize(bdStrong5)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1141
      val size0 = asize(bdStrong)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1142
      // println(s)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1143
      // println("pdersSet size")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1144
      // println(pdersSet.size)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1145
      // pdersSet.foreach(r => println(shortRexpOutput(r)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1146
      // println("after bdStrong5")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1147
      
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1148
      // println(shortRexpOutput(erase(bdStrong5)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1149
      // println(breakIntoTerms(erase(bdStrong5)).size)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1150
      
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1151
      // println("after bdStrong")
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1152
      // println(shortRexpOutput(erase(bdStrong)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1153
      // println(breakIntoTerms(erase(bdStrong)).size)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1154
      // println(size5, size0, sizeFurther5)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1155
      // aprint(strongBsimp5(bdStrong))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1156
      // println(asize(strongBsimp5(bdStrong5)))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1157
      // println(s)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1158
      size5 <= size0 
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1159
    })
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1160
    // println(boolList)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1161
    //!boolList.exists(b => b == false)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1162
    !boolList.exists(b => b == false)
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1163
  }
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1164
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1165
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1166
}
666
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
  1167
//small()
517
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1168
// generator_test()
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1169
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1170
// 1
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1171
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1172
// SEQ(SEQ(SEQ(ONE,CHAR('b')),STAR(CHAR('b'))),
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1173
// SEQ(ALTS(ALTS(ZERO,STAR(CHAR('b'))),
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1174
// STAR(ALTS(CHAR('a'),SEQ(SEQ(STAR(ALTS(STAR(CHAR('c')),CHAR('a'))),
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1175
// SEQ(CHAR('a'),SEQ(ALTS(CHAR('b'),ZERO),SEQ(ONE,CHAR('b'))))),ONE)))),ONE))
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1176
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1177
3c5e58d08939 for plotting
Chengsong
parents:
diff changeset
  1178
// Sequ(Sequ(Sequ(Empty,Chr(b)),Stars(List(Chr(b), Chr(b)))),Sequ(Right(Stars(List(Right(Sequ(Sequ(Stars(List(Right(Chr(a)), Right(Chr(a)))),Sequ(Chr(a),Sequ(Left(Chr(b)),Sequ(Empty,Chr(b))))),Empty)), Right(Sequ(Sequ(Stars(List(Right(Chr(a)), Right(Chr(a)))),Sequ(Chr(a),Sequ(Left(Chr(b)),Sequ(Empty,Chr(b))))),Empty))))),Empty))
666
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
  1179
// Sequ(Sequ(Sequ(Empty,Chr(b)),Stars(List(Chr(b), Chr(b)))),Sequ(Right(Stars(List(Right(Sequ(Sequ(Stars(List(Right(Chr(a)), Right(Chr(a)))),Sequ(Chr(a),Sequ(Left(Chr(b)),Sequ(Empty,Chr(b))))),Empty)), Right(Sequ(Sequ(Stars(List(Right(Chr(a)), Right(Chr(a)))),Sequ(Chr(a),Sequ(Left(Chr(b)),Sequ(Empty,Chr(b))))),Empty))))),Empty))
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
  1180
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
  1181
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
  1182
//def time[R](block: => R): R = {
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
  1183
//    val t0 = System.nanoTime()
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
  1184
//    val result = block    // call-by-name
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
  1185
//    val t1 = System.nanoTime()
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
  1186
//    println(" " + (t1 - t0)/1e9 + "")
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
  1187
//    result
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
  1188
//}
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
  1189
//val astarstarB  = STAR(STAR("a"))~("b")
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
  1190
//val data_points_x_axis = List.range(0, 60000, 3000)
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
  1191
//for (i <- data_points_x_axis) {print(i); time { blexing(astarstarB, ("a"*i)) } }
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
  1192
////println(data_points_x_axis.zip( List(0,1,3,4)))
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
  1193
//for(i <- List.range(0, 40000, 3000)) print(i + " ")
6da4516ea87d more changes to figures & benchmarking
Chengsong
parents: 545
diff changeset
  1194