| 518 |      1 | //Strong Bsimp to obtain Antimirov's cubic bound
 | 
|  |      2 | 
 | 
| 431 |      3 | // A simple lexer inspired by work of Sulzmann & Lu
 | 
|  |      4 | //==================================================
 | 
|  |      5 | //
 | 
|  |      6 | // Call the test cases with 
 | 
|  |      7 | //
 | 
|  |      8 | //   amm lexer.sc small
 | 
|  |      9 | //   amm lexer.sc fib
 | 
|  |     10 | //   amm lexer.sc loops
 | 
|  |     11 | //   amm lexer.sc email
 | 
|  |     12 | //
 | 
|  |     13 | //   amm lexer.sc all
 | 
|  |     14 | 
 | 
| 514 |     15 | 
 | 
| 431 |     16 | 
 | 
|  |     17 | // regular expressions including records
 | 
|  |     18 | abstract class Rexp 
 | 
|  |     19 | case object ZERO extends Rexp
 | 
|  |     20 | case object ONE extends Rexp
 | 
|  |     21 | case object ANYCHAR extends Rexp
 | 
|  |     22 | case class CHAR(c: Char) extends Rexp
 | 
|  |     23 | case class ALTS(r1: Rexp, r2: Rexp) extends Rexp 
 | 
| 539 |     24 | case class ALTSS(rs: List[Rexp]) extends Rexp
 | 
| 431 |     25 | case class SEQ(r1: Rexp, r2: Rexp) extends Rexp 
 | 
|  |     26 | case class STAR(r: Rexp) extends Rexp 
 | 
|  |     27 | case class RECD(x: String, r: Rexp) extends Rexp  
 | 
|  |     28 | case class NTIMES(n: Int, r: Rexp) extends Rexp
 | 
|  |     29 | case class OPTIONAL(r: Rexp) extends Rexp
 | 
|  |     30 | case class NOT(r: Rexp) extends Rexp
 | 
|  |     31 |                 // records for extracting strings or tokens
 | 
|  |     32 |   
 | 
|  |     33 | // values  
 | 
|  |     34 | abstract class Val
 | 
| 492 |     35 | case object Failure extends Val
 | 
| 431 |     36 | case object Empty extends Val
 | 
|  |     37 | case class Chr(c: Char) extends Val
 | 
|  |     38 | case class Sequ(v1: Val, v2: Val) extends Val
 | 
|  |     39 | case class Left(v: Val) extends Val
 | 
|  |     40 | case class Right(v: Val) extends Val
 | 
|  |     41 | case class Stars(vs: List[Val]) extends Val
 | 
|  |     42 | case class Rec(x: String, v: Val) extends Val
 | 
|  |     43 | case class Ntime(vs: List[Val]) extends Val
 | 
|  |     44 | case class Optionall(v: Val) extends Val
 | 
|  |     45 | case class Nots(s: String) extends Val
 | 
|  |     46 | 
 | 
|  |     47 | 
 | 
|  |     48 | 
 | 
|  |     49 | abstract class Bit
 | 
|  |     50 | case object Z extends Bit
 | 
|  |     51 | case object S extends Bit
 | 
|  |     52 | 
 | 
|  |     53 | 
 | 
|  |     54 | type Bits = List[Bit]
 | 
|  |     55 | 
 | 
|  |     56 | abstract class ARexp 
 | 
|  |     57 | case object AZERO extends ARexp
 | 
|  |     58 | case class AONE(bs: Bits) extends ARexp
 | 
|  |     59 | case class ACHAR(bs: Bits, c: Char) extends ARexp
 | 
|  |     60 | case class AALTS(bs: Bits, rs: List[ARexp]) extends ARexp 
 | 
|  |     61 | case class ASEQ(bs: Bits, r1: ARexp, r2: ARexp) extends ARexp 
 | 
|  |     62 | case class ASTAR(bs: Bits, r: ARexp) extends ARexp 
 | 
|  |     63 | case class ANOT(bs: Bits, r: ARexp) extends ARexp
 | 
|  |     64 | case class AANYCHAR(bs: Bits) extends ARexp
 | 
|  |     65 | 
 | 
| 514 |     66 | import scala.util.Try
 | 
|  |     67 | 
 | 
| 492 |     68 | trait Generator[+T] {
 | 
|  |     69 |     self => // an alias for "this"
 | 
|  |     70 |     def generate(): T
 | 
|  |     71 |   
 | 
|  |     72 |     def gen(n: Int) : List[T] = 
 | 
|  |     73 |       if (n == 0) Nil else self.generate() :: gen(n - 1)
 | 
|  |     74 |     
 | 
|  |     75 |     def map[S](f: T => S): Generator[S] = new Generator[S] {
 | 
|  |     76 |       def generate = f(self.generate())  
 | 
|  |     77 |     }
 | 
|  |     78 |     def flatMap[S](f: T => Generator[S]): Generator[S] = new Generator[S] {
 | 
|  |     79 |       def generate = f(self.generate()).generate()
 | 
|  |     80 |     }
 | 
|  |     81 | 
 | 
|  |     82 | 
 | 
|  |     83 | }
 | 
|  |     84 | 
 | 
|  |     85 |   // tests a property according to a given random generator
 | 
|  |     86 |   def test[T](r: Generator[T], amount: Int = 100)(pred: T => Boolean) {
 | 
|  |     87 |     for (_ <- 0 until amount) {
 | 
|  |     88 |       val value = r.generate()
 | 
|  |     89 |       assert(pred(value), s"Test failed for: $value")
 | 
|  |     90 |     }
 | 
|  |     91 |     println(s"Test passed $amount times")
 | 
|  |     92 |   }
 | 
|  |     93 |   def test2[T, S](r: Generator[T], s: Generator[S], amount: Int = 100)(pred: (T, S) => Boolean) {
 | 
|  |     94 |     for (_ <- 0 until amount) {
 | 
|  |     95 |       val valueR = r.generate()
 | 
|  |     96 |       val valueS = s.generate()
 | 
|  |     97 |       assert(pred(valueR, valueS), s"Test failed for: $valueR, $valueS")
 | 
|  |     98 |     }
 | 
|  |     99 |     println(s"Test passed $amount times")
 | 
|  |    100 |   }
 | 
|  |    101 | 
 | 
|  |    102 | // random integers
 | 
|  |    103 | val integers = new Generator[Int] {
 | 
|  |    104 |   val rand = new java.util.Random
 | 
|  |    105 |   def generate() = rand.nextInt()
 | 
|  |    106 | }
 | 
|  |    107 | 
 | 
|  |    108 | // random booleans
 | 
|  |    109 | val booleans = integers.map(_ > 0)
 | 
|  |    110 |   
 | 
|  |    111 | // random integers in the range lo and high  
 | 
|  |    112 | def range(lo: Int, hi: Int): Generator[Int] = 
 | 
|  |    113 |   for (x <- integers) yield (lo + x.abs % (hi - lo)).abs
 | 
|  |    114 | 
 | 
|  |    115 | // random characters
 | 
|  |    116 | def chars_range(lo: Char, hi: Char) = range(lo, hi).map(_.toChar)  
 | 
|  |    117 | val chars = chars_range('a', 'z')
 | 
|  |    118 | 
 | 
|  |    119 | 
 | 
|  |    120 | def oneOf[T](xs: T*): Generator[T] = 
 | 
|  |    121 |   for (idx <- range(0, xs.length)) yield xs(idx)
 | 
|  |    122 |   
 | 
|  |    123 | def single[T](x: T) = new Generator[T] {
 | 
|  |    124 |   def generate() = x
 | 
|  |    125 | }   
 | 
|  |    126 | 
 | 
|  |    127 | def pairs[T, U](t: Generator[T], u: Generator[U]): Generator[(T, U)] = 
 | 
|  |    128 |   for (x <- t; y <- u) yield (x, y)
 | 
|  |    129 | 
 | 
|  |    130 | // lists
 | 
|  |    131 | def emptyLists = single(Nil) 
 | 
|  |    132 | 
 | 
|  |    133 | def nonEmptyLists : Generator[List[Int]] = 
 | 
|  |    134 |   for (head <- integers; tail <- lists) yield head :: tail
 | 
|  |    135 | 
 | 
|  |    136 | def lists: Generator[List[Int]] = for {
 | 
|  |    137 |   kind <- booleans
 | 
|  |    138 |   list <- if (kind) emptyLists else nonEmptyLists
 | 
|  |    139 | } yield list
 | 
|  |    140 | 
 | 
|  |    141 | def char_list(len: Int): Generator[List[Char]] = {
 | 
|  |    142 |   if(len <= 0) single(Nil)
 | 
|  |    143 |   else{
 | 
|  |    144 |     for { 
 | 
| 500 |    145 |       c <- chars_range('a', 'c')
 | 
| 492 |    146 |       tail <- char_list(len - 1)
 | 
|  |    147 |     } yield c :: tail
 | 
|  |    148 |   }
 | 
|  |    149 | }
 | 
|  |    150 | 
 | 
|  |    151 | def strings(len: Int): Generator[String] = for(cs <- char_list(len)) yield cs.toString
 | 
|  |    152 | 
 | 
| 493 |    153 | def sampleString(r: Rexp) : List[String] = r match {
 | 
|  |    154 |   case STAR(r) => stringsFromRexp(r).flatMap(s => List("", s, s ++ s))//only generate 0, 1, 2 reptitions
 | 
|  |    155 |   case SEQ(r1, r2) => stringsFromRexp(r1).flatMap(s1 => stringsFromRexp(r2).map(s2 => s1 ++ s2) )
 | 
|  |    156 |   case ALTS(r1, r2) => throw new Error(s" Rexp ${r} not expected: all alternatives are supposed to have been opened up")
 | 
|  |    157 |   case ONE => "" :: Nil
 | 
|  |    158 |   case ZERO => Nil
 | 
|  |    159 |   case CHAR(c) => c.toString :: Nil
 | 
|  |    160 | 
 | 
|  |    161 | }
 | 
|  |    162 | 
 | 
|  |    163 | def stringsFromRexp(r: Rexp) : List[String] = 
 | 
|  |    164 |   breakIntoTerms(r).flatMap(r => sampleString(r))
 | 
|  |    165 | 
 | 
| 492 |    166 | 
 | 
|  |    167 | // (simple) binary trees
 | 
|  |    168 | trait Tree[T]
 | 
|  |    169 | case class Inner[T](left: Tree[T], right: Tree[T]) extends Tree[T]
 | 
|  |    170 | case class Leaf[T](x: T) extends Tree[T]
 | 
|  |    171 | 
 | 
|  |    172 | def leafs[T](t: Generator[T]): Generator[Leaf[T]] = 
 | 
|  |    173 |   for (x <- t) yield Leaf(x)
 | 
|  |    174 | 
 | 
|  |    175 | def inners[T](t: Generator[T]): Generator[Inner[T]] = 
 | 
|  |    176 |   for (l <- trees(t); r <- trees(t)) yield Inner(l, r)
 | 
|  |    177 | 
 | 
|  |    178 | def trees[T](t: Generator[T]): Generator[Tree[T]] = 
 | 
|  |    179 |   for (kind <- range(0, 2);  
 | 
|  |    180 |        tree <- if (kind == 0) leafs(t) else inners(t)) yield tree
 | 
|  |    181 | 
 | 
|  |    182 | // regular expressions
 | 
|  |    183 | 
 | 
|  |    184 | // generates random leaf-regexes; prefers CHAR-regexes
 | 
|  |    185 | def leaf_rexp() : Generator[Rexp] =
 | 
|  |    186 |   for (kind <- range(0, 5);
 | 
|  |    187 |        c <- chars_range('a', 'd')) yield
 | 
|  |    188 |     kind match {
 | 
|  |    189 |       case 0 => ZERO
 | 
|  |    190 |       case 1 => ONE
 | 
|  |    191 |       case _ => CHAR(c) 
 | 
|  |    192 |     }
 | 
|  |    193 | 
 | 
|  |    194 | // generates random inner regexes with maximum depth d
 | 
|  |    195 | def inner_rexp(d: Int) : Generator[Rexp] =
 | 
|  |    196 |   for (kind <- range(0, 3);
 | 
|  |    197 |        l <- rexp(d); 
 | 
|  |    198 |        r <- rexp(d))
 | 
|  |    199 |   yield kind match {
 | 
|  |    200 |     case 0 => ALTS(l, r)
 | 
|  |    201 |     case 1 => SEQ(l, r)
 | 
|  |    202 |     case 2 => STAR(r)
 | 
|  |    203 |   }
 | 
|  |    204 | 
 | 
|  |    205 | // generates random regexes with maximum depth d;
 | 
|  |    206 | // prefers inner regexes in 2/3 of the cases
 | 
|  |    207 | def rexp(d: Int = 100): Generator[Rexp] = 
 | 
|  |    208 |   for (kind <- range(0, 3);
 | 
|  |    209 |        r <- if (d <= 0 || kind == 0) leaf_rexp() else inner_rexp(d - 1)) yield r
 | 
|  |    210 | 
 | 
|  |    211 | 
 | 
|  |    212 | // some test functions for rexps
 | 
|  |    213 | def height(r: Rexp) : Int = r match {
 | 
|  |    214 |   case ZERO => 1
 | 
|  |    215 |   case ONE => 1
 | 
|  |    216 |   case CHAR(_) => 1
 | 
|  |    217 |   case ALTS(r1, r2) => 1 + List(height(r1), height(r2)).max
 | 
|  |    218 |   case SEQ(r1, r2) =>  1 + List(height(r1), height(r2)).max
 | 
|  |    219 |   case STAR(r) => 1 + height(r)
 | 
|  |    220 | }
 | 
|  |    221 | 
 | 
|  |    222 | // def size(r: Rexp) : Int = r match {
 | 
|  |    223 | //   case ZERO => 1
 | 
|  |    224 | //   case ONE => 1
 | 
|  |    225 | //   case CHAR(_) => 1
 | 
|  |    226 | //   case ALTS(r1, r2) => 1 + size(r1) + size(r2)
 | 
|  |    227 | //   case SEQ(r1, r2) =>  1 + size(r1) + size(r2)
 | 
|  |    228 | //   case STAR(r) => 1 + size(r) 
 | 
|  |    229 | // }
 | 
|  |    230 | 
 | 
|  |    231 | // randomly subtracts 1 or 2 from the STAR case
 | 
|  |    232 | def size_faulty(r: Rexp) : Int = r match {
 | 
|  |    233 |   case ZERO => 1
 | 
|  |    234 |   case ONE => 1
 | 
|  |    235 |   case CHAR(_) => 1
 | 
|  |    236 |   case ALTS(r1, r2) => 1 + size_faulty(r1) + size_faulty(r2)
 | 
|  |    237 |   case SEQ(r1, r2) =>  1 + size_faulty(r1) + size_faulty(r2)
 | 
|  |    238 |   case STAR(r) => 1 + size_faulty(r) - range(0, 2).generate
 | 
|  |    239 | }
 | 
|  |    240 | 
 | 
| 431 |    241 | 
 | 
|  |    242 |    
 | 
|  |    243 | // some convenience for typing in regular expressions
 | 
|  |    244 | 
 | 
|  |    245 | def charlist2rexp(s : List[Char]): Rexp = s match {
 | 
|  |    246 |   case Nil => ONE
 | 
|  |    247 |   case c::Nil => CHAR(c)
 | 
|  |    248 |   case c::s => SEQ(CHAR(c), charlist2rexp(s))
 | 
|  |    249 | }
 | 
|  |    250 | implicit def string2rexp(s : String) : Rexp = 
 | 
|  |    251 |   charlist2rexp(s.toList)
 | 
|  |    252 | 
 | 
|  |    253 | implicit def RexpOps(r: Rexp) = new {
 | 
|  |    254 |   def | (s: Rexp) = ALTS(r, s)
 | 
|  |    255 |   def % = STAR(r)
 | 
|  |    256 |   def ~ (s: Rexp) = SEQ(r, s)
 | 
|  |    257 | }
 | 
|  |    258 | 
 | 
|  |    259 | implicit def stringOps(s: String) = new {
 | 
|  |    260 |   def | (r: Rexp) = ALTS(s, r)
 | 
|  |    261 |   def | (r: String) = ALTS(s, r)
 | 
|  |    262 |   def % = STAR(s)
 | 
|  |    263 |   def ~ (r: Rexp) = SEQ(s, r)
 | 
|  |    264 |   def ~ (r: String) = SEQ(s, r)
 | 
|  |    265 |   def $ (r: Rexp) = RECD(s, r)
 | 
|  |    266 | }
 | 
|  |    267 | 
 | 
|  |    268 | def nullable(r: Rexp) : Boolean = r match {
 | 
|  |    269 |   case ZERO => false
 | 
|  |    270 |   case ONE => true
 | 
|  |    271 |   case CHAR(_) => false
 | 
|  |    272 |   case ANYCHAR => false
 | 
|  |    273 |   case ALTS(r1, r2) => nullable(r1) || nullable(r2)
 | 
| 539 |    274 |   case ALTSS(rs) => rs.exists(nullable)
 | 
| 431 |    275 |   case SEQ(r1, r2) => nullable(r1) && nullable(r2)
 | 
|  |    276 |   case STAR(_) => true
 | 
|  |    277 |   case RECD(_, r1) => nullable(r1)
 | 
|  |    278 |   case NTIMES(n, r) => if (n == 0) true else nullable(r)
 | 
|  |    279 |   case OPTIONAL(r) => true
 | 
|  |    280 |   case NOT(r) => !nullable(r)
 | 
|  |    281 | }
 | 
|  |    282 | 
 | 
|  |    283 | def der(c: Char, r: Rexp) : Rexp = r match {
 | 
|  |    284 |   case ZERO => ZERO
 | 
|  |    285 |   case ONE => ZERO
 | 
|  |    286 |   case CHAR(d) => if (c == d) ONE else ZERO
 | 
|  |    287 |   case ANYCHAR => ONE 
 | 
|  |    288 |   case ALTS(r1, r2) => ALTS(der(c, r1), der(c, r2))
 | 
| 539 |    289 |   case ALTSS(rs) => ALTSS(rs.map(der(c, _)))
 | 
| 431 |    290 |   case SEQ(r1, r2) => 
 | 
|  |    291 |     if (nullable(r1)) ALTS(SEQ(der(c, r1), r2), der(c, r2))
 | 
|  |    292 |     else SEQ(der(c, r1), r2)
 | 
|  |    293 |   case STAR(r) => SEQ(der(c, r), STAR(r))
 | 
|  |    294 |   case RECD(_, r1) => der(c, r1)
 | 
|  |    295 |   case NTIMES(n, r) => if(n > 0) SEQ(der(c, r), NTIMES(n - 1, r)) else ZERO
 | 
|  |    296 |   case OPTIONAL(r) => der(c, r)
 | 
|  |    297 |   case NOT(r) =>  NOT(der(c, r))
 | 
|  |    298 | }
 | 
|  |    299 | 
 | 
| 539 |    300 | def ders(s: List[Char], r: Rexp) : Rexp = s match {
 | 
|  |    301 |   case Nil => r
 | 
|  |    302 |   case c :: cs => ders(cs, der(c, r))
 | 
|  |    303 | }
 | 
|  |    304 | 
 | 
|  |    305 | def ders_simp(s: List[Char], r: Rexp) : Rexp = s match {
 | 
|  |    306 |   case Nil => simp(r)
 | 
|  |    307 |   case c :: cs => ders_simp(cs, simp(der(c, r)))
 | 
|  |    308 | }
 | 
|  |    309 | 
 | 
|  |    310 | 
 | 
|  |    311 | def simp2(r: Rexp) : Rexp = r match {
 | 
|  |    312 |   case SEQ(r1, r2) => 
 | 
|  |    313 |     (simp2(r1), simp2(r2)) match {
 | 
|  |    314 |       case (ZERO, _) => ZERO
 | 
|  |    315 |       case (_, ZERO) => ZERO
 | 
|  |    316 |       case (ONE, r2s) => r2s
 | 
|  |    317 |       case (r1s, ONE) => r1s
 | 
|  |    318 |       case (r1s, r2s) => 
 | 
|  |    319 |         SEQ(r1s, r2s)
 | 
|  |    320 |     }
 | 
|  |    321 |   case ALTS(r1, r2) => 
 | 
|  |    322 |     val r1r2s = fltsplain(simp2(r1) :: simp2(r2) :: Nil).distinct 
 | 
|  |    323 |     r1r2s match {
 | 
|  |    324 |       case Nil => ZERO
 | 
|  |    325 |       case r :: Nil => r
 | 
|  |    326 |       case r1 :: r2 :: rs => 
 | 
|  |    327 |         ALTSS(r1 :: r2 :: rs)
 | 
|  |    328 |     }
 | 
|  |    329 |   case ALTSS(rs) => 
 | 
|  |    330 |     // println(rs)
 | 
|  |    331 |     (fltsplain(rs.map(simp2))).distinct match {
 | 
|  |    332 |       case Nil => ZERO
 | 
|  |    333 |       case r :: Nil => r
 | 
|  |    334 |       case r1 :: r2 :: rs =>
 | 
|  |    335 |         ALTSS(r1 :: r2 :: rs)
 | 
|  |    336 |     }
 | 
|  |    337 |   case r => r
 | 
|  |    338 | }
 | 
|  |    339 | 
 | 
|  |    340 | 
 | 
|  |    341 | def simp(r: Rexp) : Rexp = r match {
 | 
|  |    342 |   case SEQ(r1, r2) => 
 | 
|  |    343 |     (simp(r1), simp(r2)) match {
 | 
|  |    344 |       case (ZERO, _) => ZERO
 | 
|  |    345 |       case (_, ZERO) => ZERO
 | 
|  |    346 |       case (ONE, r2s) => r2s
 | 
|  |    347 |       case (r1s, ONE) => r1s
 | 
|  |    348 |       case (r1s, r2s) => SEQ(r1s, r2s)
 | 
|  |    349 |     }
 | 
|  |    350 |   case ALTS(r1, r2) => {
 | 
|  |    351 |     (simp(r1), simp(r2)) match {
 | 
|  |    352 |       case (ZERO, r2s) => r2s
 | 
|  |    353 |       case (r1s, ZERO) => r1s
 | 
|  |    354 |       case (r1s, r2s) =>
 | 
|  |    355 |         if(r1s == r2s) r1s else ALTS(r1s, r2s)
 | 
|  |    356 |     }
 | 
|  |    357 |   }
 | 
|  |    358 |   case r => r
 | 
|  |    359 | }
 | 
|  |    360 | 
 | 
|  |    361 | def fltsplain(rs: List[Rexp]) : List[Rexp] = rs match {
 | 
|  |    362 |   case Nil => Nil
 | 
|  |    363 |   case ZERO :: rs => fltsplain(rs)
 | 
|  |    364 |   case ALTS(r1, r2) :: rs => r1 :: r2 :: fltsplain(rs) 
 | 
|  |    365 |   case ALTSS(rs) :: rs1 => rs ::: fltsplain(rs1)
 | 
|  |    366 |   case r :: rs => r :: fltsplain(rs)
 | 
|  |    367 | }
 | 
|  |    368 | 
 | 
|  |    369 | 
 | 
|  |    370 | 
 | 
|  |    371 | 
 | 
|  |    372 | def matcher(s: String, r: Rexp) : Boolean = 
 | 
|  |    373 |   nullable(ders(s.toList, r))
 | 
|  |    374 | 
 | 
|  |    375 | def simp_matcher(s: String, r: Rexp) : Boolean = 
 | 
|  |    376 |   nullable(ders_simp(s.toList, r))
 | 
|  |    377 | 
 | 
| 431 |    378 | 
 | 
|  |    379 | // extracts a string from a value
 | 
|  |    380 | def flatten(v: Val) : String = v match {
 | 
|  |    381 |   case Empty => ""
 | 
|  |    382 |   case Chr(c) => c.toString
 | 
|  |    383 |   case Left(v) => flatten(v)
 | 
|  |    384 |   case Right(v) => flatten(v)
 | 
|  |    385 |   case Sequ(v1, v2) => flatten(v1) ++ flatten(v2)
 | 
|  |    386 |   case Stars(vs) => vs.map(flatten).mkString
 | 
|  |    387 |   case Ntime(vs) => vs.map(flatten).mkString
 | 
|  |    388 |   case Optionall(v) => flatten(v)
 | 
|  |    389 |   case Rec(_, v) => flatten(v)
 | 
|  |    390 | }
 | 
|  |    391 | 
 | 
|  |    392 | 
 | 
|  |    393 | // extracts an environment from a value;
 | 
|  |    394 | // used for tokenising a string
 | 
|  |    395 | def env(v: Val) : List[(String, String)] = v match {
 | 
|  |    396 |   case Empty => Nil
 | 
|  |    397 |   case Chr(c) => Nil
 | 
|  |    398 |   case Left(v) => env(v)
 | 
|  |    399 |   case Right(v) => env(v)
 | 
|  |    400 |   case Sequ(v1, v2) => env(v1) ::: env(v2)
 | 
|  |    401 |   case Stars(vs) => vs.flatMap(env)
 | 
|  |    402 |   case Ntime(vs) => vs.flatMap(env)
 | 
|  |    403 |   case Rec(x, v) => (x, flatten(v))::env(v)
 | 
|  |    404 |   case Optionall(v) => env(v)
 | 
|  |    405 |   case Nots(s) => ("Negative", s) :: Nil
 | 
|  |    406 | }
 | 
|  |    407 | 
 | 
|  |    408 | 
 | 
|  |    409 | // The injection and mkeps part of the lexer
 | 
|  |    410 | //===========================================
 | 
|  |    411 | 
 | 
|  |    412 | def mkeps(r: Rexp) : Val = r match {
 | 
|  |    413 |   case ONE => Empty
 | 
|  |    414 |   case ALTS(r1, r2) => 
 | 
|  |    415 |     if (nullable(r1)) Left(mkeps(r1)) else Right(mkeps(r2))
 | 
|  |    416 |   case SEQ(r1, r2) => Sequ(mkeps(r1), mkeps(r2))
 | 
|  |    417 |   case STAR(r) => Stars(Nil)
 | 
|  |    418 |   case RECD(x, r) => Rec(x, mkeps(r))
 | 
|  |    419 |   case NTIMES(n, r) => Ntime(List.fill(n)(mkeps(r)))
 | 
|  |    420 |   case OPTIONAL(r) => Optionall(Empty)
 | 
|  |    421 |   case NOT(rInner) => if(nullable(rInner)) throw new Exception("error")  
 | 
|  |    422 |                          else Nots("")//Nots(s.reverse.toString)
 | 
|  |    423 | //   case NOT(ZERO) => Empty
 | 
|  |    424 | //   case NOT(CHAR(c)) => Empty
 | 
|  |    425 | //   case NOT(SEQ(r1, r2)) => Sequ(mkeps(NOT(r1)), mkeps(NOT(r2)))
 | 
|  |    426 | //   case NOT(ALTS(r1, r2)) => if(!nullable(r1)) Left(mkeps(NOT(r1))) else Right(mkeps(NOT(r2)))
 | 
|  |    427 | //   case NOT(STAR(r)) => Stars(Nil) 
 | 
|  |    428 | 
 | 
|  |    429 | }
 | 
|  |    430 | 
 | 
|  |    431 | def inj(r: Rexp, c: Char, v: Val) : Val = (r, v) match {
 | 
|  |    432 |   case (STAR(r), Sequ(v1, Stars(vs))) => Stars(inj(r, c, v1)::vs)
 | 
|  |    433 |   case (SEQ(r1, r2), Sequ(v1, v2)) => Sequ(inj(r1, c, v1), v2)
 | 
|  |    434 |   case (SEQ(r1, r2), Left(Sequ(v1, v2))) => Sequ(inj(r1, c, v1), v2)
 | 
|  |    435 |   case (SEQ(r1, r2), Right(v2)) => Sequ(mkeps(r1), inj(r2, c, v2))
 | 
|  |    436 |   case (ALTS(r1, r2), Left(v1)) => Left(inj(r1, c, v1))
 | 
|  |    437 |   case (ALTS(r1, r2), Right(v2)) => Right(inj(r2, c, v2))
 | 
|  |    438 |   case (CHAR(d), Empty) => Chr(c) 
 | 
|  |    439 |   case (RECD(x, r1), _) => Rec(x, inj(r1, c, v))
 | 
|  |    440 |   case (NTIMES(n, r), Sequ(v1, Ntime(vs))) => Ntime(inj(r, c, v1)::vs)
 | 
|  |    441 |   case (OPTIONAL(r), v) => Optionall(inj(r, c, v))
 | 
|  |    442 |   case (NOT(r), Nots(s)) => Nots(c.toString ++ s)
 | 
|  |    443 |   case (ANYCHAR, Empty) => Chr(c)
 | 
|  |    444 | }
 | 
|  |    445 | 
 | 
|  |    446 | // some "rectification" functions for simplification
 | 
|  |    447 | 
 | 
|  |    448 | 
 | 
|  |    449 | 
 | 
|  |    450 | 
 | 
|  |    451 | // The Lexing Rules for the WHILE Language
 | 
|  |    452 | 
 | 
|  |    453 |   // bnullable function: tests whether the aregular 
 | 
|  |    454 |   // expression can recognise the empty string
 | 
|  |    455 | def bnullable (r: ARexp) : Boolean = r match {
 | 
|  |    456 |     case AZERO => false
 | 
|  |    457 |     case AONE(_) => true
 | 
|  |    458 |     case ACHAR(_,_) => false
 | 
|  |    459 |     case AALTS(_, rs) => rs.exists(bnullable)
 | 
|  |    460 |     case ASEQ(_, r1, r2) => bnullable(r1) && bnullable(r2)
 | 
|  |    461 |     case ASTAR(_, _) => true
 | 
|  |    462 |     case ANOT(_, rn) => !bnullable(rn)
 | 
|  |    463 |   }
 | 
|  |    464 | 
 | 
|  |    465 | def mkepsBC(r: ARexp) : Bits = r match {
 | 
|  |    466 |     case AONE(bs) => bs
 | 
|  |    467 |     case AALTS(bs, rs) => {
 | 
|  |    468 |       val n = rs.indexWhere(bnullable)
 | 
|  |    469 |       bs ++ mkepsBC(rs(n))
 | 
|  |    470 |     }
 | 
|  |    471 |     case ASEQ(bs, r1, r2) => bs ++ mkepsBC(r1) ++ mkepsBC(r2)
 | 
|  |    472 |     case ASTAR(bs, r) => bs ++ List(Z)
 | 
|  |    473 |     case ANOT(bs, rn) => bs
 | 
|  |    474 |   }
 | 
|  |    475 | 
 | 
|  |    476 | 
 | 
|  |    477 | def bder(c: Char, r: ARexp) : ARexp = r match {
 | 
|  |    478 |     case AZERO => AZERO
 | 
|  |    479 |     case AONE(_) => AZERO
 | 
|  |    480 |     case ACHAR(bs, f) => if (c == f) AONE(bs) else AZERO
 | 
|  |    481 |     case AALTS(bs, rs) => AALTS(bs, rs.map(bder(c, _)))
 | 
|  |    482 |     case ASEQ(bs, r1, r2) => 
 | 
|  |    483 |       if (bnullable(r1)) AALTS(bs, ASEQ(Nil, bder(c, r1), r2) :: fuse(mkepsBC(r1), bder(c, r2)) :: Nil )
 | 
|  |    484 |       else ASEQ(bs, bder(c, r1), r2)
 | 
|  |    485 |     case ASTAR(bs, r) => ASEQ(bs, fuse(List(S), bder(c, r)), ASTAR(Nil, r))
 | 
|  |    486 |     case ANOT(bs, rn) => ANOT(bs, bder(c, rn))
 | 
|  |    487 |     case AANYCHAR(bs) => AONE(bs)
 | 
|  |    488 |   } 
 | 
|  |    489 | 
 | 
|  |    490 | def fuse(bs: Bits, r: ARexp) : ARexp = r match {
 | 
|  |    491 |     case AZERO => AZERO
 | 
|  |    492 |     case AONE(cs) => AONE(bs ++ cs)
 | 
|  |    493 |     case ACHAR(cs, f) => ACHAR(bs ++ cs, f)
 | 
|  |    494 |     case AALTS(cs, rs) => AALTS(bs ++ cs, rs)
 | 
|  |    495 |     case ASEQ(cs, r1, r2) => ASEQ(bs ++ cs, r1, r2)
 | 
|  |    496 |     case ASTAR(cs, r) => ASTAR(bs ++ cs, r)
 | 
|  |    497 |     case ANOT(cs, r) => ANOT(bs ++ cs, r)
 | 
|  |    498 |   }
 | 
|  |    499 | 
 | 
|  |    500 | 
 | 
|  |    501 | def internalise(r: Rexp) : ARexp = r match {
 | 
|  |    502 |     case ZERO => AZERO
 | 
|  |    503 |     case ONE => AONE(Nil)
 | 
|  |    504 |     case CHAR(c) => ACHAR(Nil, c)
 | 
|  |    505 |     //case PRED(f) => APRED(Nil, f)
 | 
|  |    506 |     case ALTS(r1, r2) => 
 | 
|  |    507 |       AALTS(Nil, List(fuse(List(Z), internalise(r1)), fuse(List(S), internalise(r2))))
 | 
|  |    508 |     // case ALTS(r1::rs) => {
 | 
|  |    509 |     //   val AALTS(Nil, rs2) = internalise(ALTS(rs))
 | 
|  |    510 |     //   AALTS(Nil, fuse(List(Z), internalise(r1)) :: rs2.map(fuse(List(S), _)))
 | 
|  |    511 |     // }
 | 
|  |    512 |     case SEQ(r1, r2) => ASEQ(Nil, internalise(r1), internalise(r2))
 | 
|  |    513 |     case STAR(r) => ASTAR(Nil, internalise(r))
 | 
|  |    514 |     case RECD(x, r) => internalise(r)
 | 
|  |    515 |     case NOT(r) => ANOT(Nil, internalise(r))
 | 
|  |    516 |     case ANYCHAR => AANYCHAR(Nil)
 | 
|  |    517 |   }
 | 
|  |    518 | 
 | 
|  |    519 | 
 | 
|  |    520 | def bsimp(r: ARexp): ARexp = 
 | 
|  |    521 |   {
 | 
|  |    522 |     r match {
 | 
|  |    523 |       case ASEQ(bs1, r1, r2) => (bsimp(r1), bsimp(r2)) match {
 | 
|  |    524 |           case (AZERO, _) => AZERO
 | 
|  |    525 |           case (_, AZERO) => AZERO
 | 
|  |    526 |           case (AONE(bs2), r2s) => fuse(bs1 ++ bs2, r2s)
 | 
|  |    527 |           case (r1s, r2s) => ASEQ(bs1, r1s, r2s)
 | 
|  |    528 |       }
 | 
|  |    529 |       case AALTS(bs1, rs) => {
 | 
|  |    530 |             val rs_simp = rs.map(bsimp(_))
 | 
|  |    531 |             val flat_res = flats(rs_simp)
 | 
|  |    532 |             val dist_res = distinctBy(flat_res, erase)//strongDB(flat_res)//distinctBy(flat_res, erase)
 | 
|  |    533 |             dist_res match {
 | 
|  |    534 |               case Nil => AZERO
 | 
|  |    535 |               case s :: Nil => fuse(bs1, s)
 | 
|  |    536 |               case rs => AALTS(bs1, rs)  
 | 
|  |    537 |             }
 | 
|  |    538 |           
 | 
|  |    539 |       }
 | 
|  |    540 |       case r => r
 | 
|  |    541 |     }
 | 
| 492 |    542 | }
 | 
|  |    543 | def strongBsimp(r: ARexp): ARexp =
 | 
|  |    544 | {
 | 
|  |    545 |   r match {
 | 
|  |    546 |     case ASEQ(bs1, r1, r2) => (strongBsimp(r1), strongBsimp(r2)) match {
 | 
|  |    547 |         case (AZERO, _) => AZERO
 | 
|  |    548 |         case (_, AZERO) => AZERO
 | 
|  |    549 |         case (AONE(bs2), r2s) => fuse(bs1 ++ bs2, r2s)
 | 
|  |    550 |         case (r1s, r2s) => ASEQ(bs1, r1s, r2s)
 | 
| 431 |    551 |     }
 | 
| 492 |    552 |     case AALTS(bs1, rs) => {
 | 
|  |    553 |           val rs_simp = rs.map(strongBsimp(_))
 | 
|  |    554 |           val flat_res = flats(rs_simp)
 | 
|  |    555 |           val dist_res = distinctBy4(flat_res)//distinctBy(flat_res, erase)
 | 
|  |    556 |           dist_res match {
 | 
|  |    557 |             case Nil => AZERO
 | 
|  |    558 |             case s :: Nil => fuse(bs1, s)
 | 
|  |    559 |             case rs => AALTS(bs1, rs)  
 | 
|  |    560 |           }
 | 
|  |    561 |         
 | 
|  |    562 |     }
 | 
|  |    563 |     case r => r
 | 
| 431 |    564 |   }
 | 
| 492 |    565 | }
 | 
| 431 |    566 | 
 | 
| 494 |    567 | def strongBsimp5(r: ARexp): ARexp =
 | 
|  |    568 | {
 | 
|  |    569 |   // println("was this called?")
 | 
|  |    570 |   r match {
 | 
|  |    571 |     case ASEQ(bs1, r1, r2) => (strongBsimp5(r1), strongBsimp5(r2)) match {
 | 
|  |    572 |         case (AZERO, _) => AZERO
 | 
|  |    573 |         case (_, AZERO) => AZERO
 | 
|  |    574 |         case (AONE(bs2), r2s) => fuse(bs1 ++ bs2, r2s)
 | 
|  |    575 |         case (r1s, r2s) => ASEQ(bs1, r1s, r2s)
 | 
|  |    576 |     }
 | 
|  |    577 |     case AALTS(bs1, rs) => {
 | 
|  |    578 |         // println("alts case")
 | 
|  |    579 |           val rs_simp = rs.map(strongBsimp5(_))
 | 
|  |    580 |           val flat_res = flats(rs_simp)
 | 
| 500 |    581 |           var dist_res = distinctBy5(flat_res)//distinctBy(flat_res, erase)
 | 
| 518 |    582 |           // var dist2_res = distinctBy5(dist_res)
 | 
|  |    583 |           // while(dist_res != dist2_res){
 | 
|  |    584 |           //   dist_res = dist2_res
 | 
|  |    585 |           //   dist2_res = distinctBy5(dist_res)
 | 
|  |    586 |           // }
 | 
|  |    587 |           (dist_res) match {
 | 
|  |    588 |             case Nil => AZERO
 | 
|  |    589 |             case s :: Nil => fuse(bs1, s)
 | 
|  |    590 |             case rs => AALTS(bs1, rs)  
 | 
| 500 |    591 |           }
 | 
| 518 |    592 |     }
 | 
|  |    593 |     case r => r
 | 
|  |    594 |   }
 | 
|  |    595 | }
 | 
|  |    596 | 
 | 
|  |    597 | def strongBsimp6(r: ARexp): ARexp =
 | 
|  |    598 | {
 | 
|  |    599 |   // println("was this called?")
 | 
|  |    600 |   r match {
 | 
|  |    601 |     case ASEQ(bs1, r1, r2) => (strongBsimp6(r1), strongBsimp6(r2)) match {
 | 
|  |    602 |         case (AZERO, _) => AZERO
 | 
|  |    603 |         case (_, AZERO) => AZERO
 | 
|  |    604 |         case (AONE(bs2), r2s) => fuse(bs1 ++ bs2, r2s)
 | 
| 530 |    605 |         case (r1s, AONE(bs2)) => fuse(bs1, r1s) //asserted bs2 == Nil
 | 
| 532 |    606 |         //case (r1s, r2s) if(isOne(erase(r1s))) => fuse(bs1 ++ mkepsBC(r1s), r2s)
 | 
| 518 |    607 |         case (r1s, r2s) => ASEQ(bs1, r1s, r2s)
 | 
|  |    608 |     }
 | 
|  |    609 |     case AALTS(bs1, rs) => {
 | 
|  |    610 |         // println("alts case")
 | 
|  |    611 |           val rs_simp = rs.map(strongBsimp6(_))
 | 
|  |    612 |           val flat_res = flats(rs_simp)
 | 
|  |    613 |           var dist_res = distinctBy6(flat_res)//distinctBy(flat_res, erase)
 | 
|  |    614 |           (dist_res) match {
 | 
| 494 |    615 |             case Nil => AZERO
 | 
|  |    616 |             case s :: Nil => fuse(bs1, s)
 | 
|  |    617 |             case rs => AALTS(bs1, rs)  
 | 
|  |    618 |           }
 | 
|  |    619 |     }
 | 
| 543 |    620 |     //(erase(r0))) => AONE(bs)
 | 
| 494 |    621 |     case r => r
 | 
|  |    622 |   }
 | 
|  |    623 | }
 | 
|  |    624 | 
 | 
| 532 |    625 | def distinctWith(rs: List[ARexp], 
 | 
| 533 |    626 |                 pruneFunction: (ARexp, Set[Rexp]) => ARexp, 
 | 
| 532 |    627 |                 acc: Set[Rexp] = Set()) : List[ARexp] = 
 | 
|  |    628 |   rs match{
 | 
|  |    629 |     case Nil => Nil
 | 
|  |    630 |     case r :: rs => 
 | 
|  |    631 |       if(acc(erase(r)))
 | 
| 533 |    632 |         distinctWith(rs, pruneFunction, acc)
 | 
| 532 |    633 |       else {
 | 
| 533 |    634 |         val pruned_r = pruneFunction(r, acc)
 | 
|  |    635 |         pruned_r :: 
 | 
|  |    636 |         distinctWith(rs, 
 | 
|  |    637 |           pruneFunction, 
 | 
|  |    638 |           turnIntoTerms(erase(pruned_r)) ++: acc
 | 
|  |    639 |         )
 | 
| 532 |    640 |       }
 | 
|  |    641 |   }
 | 
|  |    642 | 
 | 
| 533 |    643 | // def distinctByPlus(rs: List[ARexp], acc: Set[Rexp] = Set()) : 
 | 
|  |    644 | // List[ARexp] =  rs match {
 | 
|  |    645 | //   case Nil => Nil
 | 
|  |    646 | //   case r :: rs =>
 | 
|  |    647 | //     if(acc.contains(erase(r)))
 | 
|  |    648 | //       distinctByPlus(rs, acc)
 | 
|  |    649 | //     else 
 | 
|  |    650 | //       prune(r, acc) :: distinctByPlus(rs, prune(r, acc) +: acc)
 | 
|  |    651 | // }
 | 
|  |    652 | 
 | 
| 532 |    653 | //r = r' ~ tail => returns r'
 | 
|  |    654 | def removeSeqTail(r: Rexp, tail: Rexp) : Rexp = r match {
 | 
|  |    655 |   case SEQ(r1, r2) => 
 | 
|  |    656 |     if(r2 == tail) 
 | 
|  |    657 |       r1
 | 
|  |    658 |     else
 | 
|  |    659 |       ZERO
 | 
|  |    660 |   case r => ZERO
 | 
|  |    661 | }
 | 
|  |    662 | 
 | 
|  |    663 | def prune7(r: ARexp, acc: Set[Rexp]) : ARexp = r match{
 | 
|  |    664 |   case AALTS(bs, rs) => rs.map(r => prune7(r, acc)).filter(_ != ZERO) match
 | 
|  |    665 |   {
 | 
|  |    666 |     case Nil => AZERO
 | 
|  |    667 |     case r::Nil => fuse(bs, r)
 | 
|  |    668 |     case rs1 => AALTS(bs, rs1)
 | 
|  |    669 |   }
 | 
|  |    670 |   case ASEQ(bs, r1, r2) => prune7(r1, acc.map(r => removeSeqTail(r, erase(r2)))) match {
 | 
|  |    671 |     case AZERO => AZERO
 | 
|  |    672 |     case r1p if(isOne(erase(r1p))) => fuse(bs ++ mkepsBC(r1p), r2)
 | 
|  |    673 |     case r1p => ASEQ(bs, r1p, r2)
 | 
|  |    674 |   }
 | 
|  |    675 |   case r => if(acc(erase(r))) AZERO else r
 | 
|  |    676 | }
 | 
|  |    677 | 
 | 
|  |    678 | def strongBsimp7(r: ARexp): ARexp =
 | 
|  |    679 | {
 | 
|  |    680 |   r match {
 | 
|  |    681 |     case ASEQ(bs1, r1, r2) => (strongBsimp7(r1), strongBsimp7(r2)) match {
 | 
|  |    682 |         case (AZERO, _) => AZERO
 | 
|  |    683 |         case (_, AZERO) => AZERO
 | 
|  |    684 |         case (AONE(bs2), r2s) => fuse(bs1 ++ bs2, r2s)
 | 
|  |    685 |         case (r1s, AONE(bs2)) => fuse(bs1, r1s) //asserted bs2 == Nil
 | 
|  |    686 |         //case (r1s, r2s) if(isOne(erase(r1s))) => fuse(bs1 ++ mkepsBC(r1s), r2s)
 | 
|  |    687 |         case (r1s, r2s) => ASEQ(bs1, r1s, r2s)
 | 
|  |    688 |     }
 | 
|  |    689 |     case AALTS(bs1, rs) => {
 | 
|  |    690 |         // println("alts case")
 | 
|  |    691 |           val rs_simp = rs.map(strongBsimp7(_))
 | 
|  |    692 |           val flat_res = flats(rs_simp)
 | 
|  |    693 |           var dist_res = distinctWith(flat_res, prune7)//distinctBy(flat_res, erase)
 | 
|  |    694 |           (dist_res) match {
 | 
|  |    695 |             case Nil => AZERO
 | 
|  |    696 |             case s :: Nil => fuse(bs1, s)
 | 
|  |    697 |             case rs => AALTS(bs1, rs)  
 | 
|  |    698 |           }
 | 
|  |    699 |     }
 | 
|  |    700 |     case ASTAR(bs, r0) if(atMostEmpty(erase(r0))) => AONE(bs)
 | 
|  |    701 |     case r => r
 | 
|  |    702 |   }
 | 
|  |    703 | }
 | 
|  |    704 | 
 | 
|  |    705 | 
 | 
| 492 |    706 | def bders (s: List[Char], r: ARexp) : ARexp = s match {
 | 
|  |    707 |   case Nil => r
 | 
|  |    708 |   case c::s => bders(s, bder(c, r))
 | 
|  |    709 | }
 | 
| 431 |    710 | 
 | 
| 492 |    711 | def flats(rs: List[ARexp]): List[ARexp] = rs match {
 | 
| 431 |    712 |     case Nil => Nil
 | 
| 492 |    713 |     case AZERO :: rs1 => flats(rs1)
 | 
|  |    714 |     case AALTS(bs, rs1) :: rs2 => rs1.map(fuse(bs, _)) ::: flats(rs2)
 | 
|  |    715 |     case r1 :: rs2 => r1 :: flats(rs2)
 | 
| 431 |    716 |   }
 | 
|  |    717 | 
 | 
| 492 |    718 | def distinctBy[B, C](xs: List[B], f: B => C, acc: List[C] = Nil): List[B] = xs match {
 | 
|  |    719 |   case Nil => Nil
 | 
|  |    720 |   case (x::xs) => {
 | 
|  |    721 |     val res = f(x)
 | 
|  |    722 |     if (acc.contains(res)) distinctBy(xs, f, acc)  
 | 
|  |    723 |     else x::distinctBy(xs, f, res::acc)
 | 
| 431 |    724 |   }
 | 
| 492 |    725 | } 
 | 
| 431 |    726 | 
 | 
|  |    727 | 
 | 
| 492 |    728 | def pruneRexp(r: ARexp, allowableTerms: List[Rexp]) : ARexp = {
 | 
|  |    729 |   r match {
 | 
|  |    730 |     case ASEQ(bs, r1, r2) => 
 | 
|  |    731 |       val termsTruncated = allowableTerms.collect(rt => rt match {
 | 
|  |    732 |         case SEQ(r1p, r2p) if(r2p == erase(r2)) => r1p//if(r2p == erase(r2)) 
 | 
|  |    733 |       })
 | 
|  |    734 |       val pruned : ARexp = pruneRexp(r1, termsTruncated)
 | 
|  |    735 |       pruned match {
 | 
|  |    736 |         case AZERO => AZERO
 | 
|  |    737 |         case AONE(bs1) => fuse(bs ++ bs1, r2)
 | 
|  |    738 |         case pruned1 => ASEQ(bs, pruned1, r2)
 | 
|  |    739 |       }
 | 
|  |    740 | 
 | 
|  |    741 |     case AALTS(bs, rs) => 
 | 
|  |    742 |       //allowableTerms.foreach(a => println(shortRexpOutput(a)))        
 | 
|  |    743 |       val rsp = rs.map(r => 
 | 
|  |    744 |                     pruneRexp(r, allowableTerms)
 | 
|  |    745 |                   )
 | 
|  |    746 |                   .filter(r => r != AZERO)
 | 
|  |    747 |       rsp match {
 | 
|  |    748 |         case Nil => AZERO
 | 
|  |    749 |         case r1::Nil => fuse(bs, r1)
 | 
|  |    750 |         case rs1 => AALTS(bs, rs1)
 | 
|  |    751 |       }
 | 
|  |    752 |     case r => 
 | 
|  |    753 |       if(allowableTerms.contains(erase(r))) r else AZERO //assert(r != AZERO)
 | 
|  |    754 |   }
 | 
|  |    755 | }
 | 
|  |    756 | 
 | 
|  |    757 | def oneSimp(r: Rexp) : Rexp = r match {
 | 
|  |    758 |   case SEQ(ONE, r) => r
 | 
|  |    759 |   case SEQ(r1, r2) => SEQ(oneSimp(r1), r2)
 | 
|  |    760 |   case r => r//assert r != 0 
 | 
| 432 |    761 |     
 | 
| 492 |    762 | }
 | 
| 431 |    763 | 
 | 
|  |    764 | 
 | 
| 492 |    765 | def distinctBy4(xs: List[ARexp], acc: List[Rexp] = Nil) : List[ARexp] = xs match {
 | 
|  |    766 |   case Nil => Nil
 | 
|  |    767 |   case x :: xs =>
 | 
|  |    768 |     //assert(acc.distinct == acc)
 | 
|  |    769 |     val erased = erase(x)
 | 
|  |    770 |     if(acc.contains(erased))
 | 
|  |    771 |       distinctBy4(xs, acc)
 | 
|  |    772 |     else{
 | 
|  |    773 |       val addToAcc =  breakIntoTerms(erased).filter(r => !acc.contains(oneSimp(r)))
 | 
|  |    774 |       //val xp = pruneRexp(x, addToAcc)
 | 
|  |    775 |       pruneRexp(x, addToAcc) match {
 | 
|  |    776 |         case AZERO => distinctBy4(xs, addToAcc.map(oneSimp(_)) ::: acc)
 | 
|  |    777 |         case xPrime => xPrime :: distinctBy4(xs, addToAcc.map(oneSimp(_)) ::: acc)
 | 
|  |    778 |       }
 | 
| 431 |    779 |     }
 | 
| 492 |    780 | }
 | 
|  |    781 | 
 | 
| 494 |    782 | // fun pAKC_aux :: "arexp list ⇒ arexp ⇒ rexp ⇒ arexp"
 | 
|  |    783 | //   where
 | 
| 564 |    784 | // "pAKC_aux rsa r ctx = (if (seqFold (SEQ (erase r) ( ctx) )) ⊆ (seqFold (erase (AALTs [] rsa))) then AZERO else
 | 
| 494 |    785 | //                           case r of (ASEQ bs r1 r2) ⇒ 
 | 
|  |    786 | //                             bsimp_ASEQ bs (pAKC_aux rsa r1 (SEQ  (erase r2) ( ctx) )) r2   |
 | 
|  |    787 | //                                     (AALTs bs rs) ⇒ 
 | 
|  |    788 | //                             bsimp_AALTs bs (flts (map (λr. pAKC_aux rsa r ( ctx) ) rs) )    |
 | 
|  |    789 | //                                     r             ⇒ r
 | 
|  |    790 | 
 | 
|  |    791 | // def canonicalSeq(rs: List[Rexp], acc: Rexp) = rs match {
 | 
|  |    792 | //   case r::Nil => SEQ(r, acc)
 | 
|  |    793 | //   case Nil => acc
 | 
|  |    794 | //   case r1::r2::Nil => SEQ(SEQ(r1, r2), acc)
 | 
|  |    795 | // }
 | 
|  |    796 | 
 | 
|  |    797 | 
 | 
|  |    798 | //the "fake" Language interpretation: just concatenates!
 | 
| 564 |    799 | def seqFold(acc: Rexp, rs: List[Rexp]) : Rexp = rs match {
 | 
| 494 |    800 |   case Nil => acc
 | 
|  |    801 |   case r :: rs1 => 
 | 
| 500 |    802 |     // if(acc == ONE) 
 | 
| 564 |    803 |     //   seqFold(r, rs1) 
 | 
| 500 |    804 |     // else
 | 
| 564 |    805 |       seqFold(SEQ(acc, r), rs1)
 | 
| 494 |    806 | }
 | 
|  |    807 | 
 | 
| 500 |    808 | def rprint(r: Rexp) : Unit = println(shortRexpOutput(r))
 | 
| 526 |    809 | def rsprint(rs: Iterable[Rexp]) = rs.foreach(r => println(shortRexpOutput(r)))
 | 
|  |    810 | 
 | 
| 500 |    811 | def aprint(a: ARexp) = println(shortRexpOutput(erase(a)))
 | 
|  |    812 | def asprint(as: List[ARexp]) = as.foreach(a => println(shortRexpOutput(erase(a))))
 | 
|  |    813 | 
 | 
|  |    814 | def pAKC(acc: List[Rexp], r: ARexp, ctx: List[Rexp]) : ARexp = {
 | 
|  |    815 |   // println("pakc")
 | 
|  |    816 |   // println(shortRexpOutput(erase(r)))
 | 
|  |    817 |   // println("acc")
 | 
|  |    818 |   // rsprint(acc)
 | 
|  |    819 |   // println("ctx---------")
 | 
|  |    820 |   // rsprint(ctx)
 | 
|  |    821 |   // println("ctx---------end")
 | 
| 564 |    822 |   // rsprint(breakIntoTerms(seqFold(erase(r), ctx)).map(oneSimp))
 | 
| 500 |    823 | 
 | 
| 564 |    824 |   if (breakIntoTerms(seqFold(erase(r), ctx)).map(oneSimp).forall(acc.contains)) {//acc.flatMap(breakIntoTerms
 | 
| 494 |    825 |     AZERO
 | 
|  |    826 |   }
 | 
|  |    827 |   else{
 | 
|  |    828 |     r match {
 | 
|  |    829 |       case ASEQ(bs, r1, r2) => 
 | 
| 500 |    830 |       (pAKC(acc, r1, erase(r2) :: ctx)) match{
 | 
| 494 |    831 |         case AZERO => 
 | 
|  |    832 |           AZERO
 | 
|  |    833 |         case AONE(bs1) => 
 | 
|  |    834 |           fuse(bs1, r2)
 | 
|  |    835 |         case r1p => 
 | 
|  |    836 |           ASEQ(bs, r1p, r2)
 | 
|  |    837 |       }
 | 
|  |    838 |       case AALTS(bs, rs0) => 
 | 
|  |    839 |         // println("before pruning")
 | 
|  |    840 |         // println(s"ctx is ")
 | 
|  |    841 |         // ctx.foreach(r => println(shortRexpOutput(r)))
 | 
|  |    842 |         // println(s"rs0 is ")
 | 
|  |    843 |         // rs0.foreach(r => println(shortRexpOutput(erase(r))))
 | 
| 500 |    844 |         // println(s"acc is ")
 | 
|  |    845 |         // acc.foreach(r => println(shortRexpOutput(r)))
 | 
|  |    846 |         rs0.map(r => pAKC(acc, r, ctx)).filter(_ != AZERO) match {
 | 
| 494 |    847 |           case Nil => 
 | 
|  |    848 |             // println("after pruning Nil")
 | 
|  |    849 |             AZERO
 | 
|  |    850 |           case r :: Nil => 
 | 
|  |    851 |             // println("after pruning singleton")
 | 
|  |    852 |             // println(shortRexpOutput(erase(r)))
 | 
|  |    853 |             r 
 | 
|  |    854 |           case rs0p => 
 | 
|  |    855 |           // println("after pruning non-singleton")
 | 
|  |    856 |             AALTS(bs, rs0p)
 | 
|  |    857 |         }
 | 
|  |    858 |       case r => r
 | 
|  |    859 |     }
 | 
|  |    860 |   }
 | 
|  |    861 | }
 | 
|  |    862 | 
 | 
|  |    863 | def distinctBy5(xs: List[ARexp], acc: List[Rexp] = Nil) : List[ARexp] = xs match {
 | 
|  |    864 |   case Nil => 
 | 
|  |    865 |     Nil
 | 
|  |    866 |   case x :: xs => {
 | 
|  |    867 |     val erased = erase(x)
 | 
|  |    868 |     if(acc.contains(erased)){
 | 
|  |    869 |       distinctBy5(xs, acc)
 | 
|  |    870 |     }
 | 
|  |    871 |     else{
 | 
|  |    872 |       val pruned = pAKC(acc, x, Nil)
 | 
|  |    873 |       val newTerms = breakIntoTerms(erase(pruned))
 | 
|  |    874 |       pruned match {
 | 
|  |    875 |         case AZERO => 
 | 
|  |    876 |           distinctBy5(xs, acc)
 | 
|  |    877 |         case xPrime => 
 | 
|  |    878 |           xPrime :: distinctBy5(xs, newTerms.map(oneSimp) ::: acc)//distinctBy5(xs, addToAcc.map(oneSimp(_)) ::: acc)
 | 
|  |    879 |       }
 | 
|  |    880 |     }
 | 
|  |    881 |   }
 | 
|  |    882 | }
 | 
|  |    883 | 
 | 
| 532 |    884 | def atMostEmpty(r: Rexp) : Boolean = r match {
 | 
|  |    885 |   case ZERO => true
 | 
|  |    886 |   case ONE => true
 | 
|  |    887 |   case STAR(r) => atMostEmpty(r)
 | 
|  |    888 |   case SEQ(r1, r2) => atMostEmpty(r1) && atMostEmpty(r2)
 | 
|  |    889 |   case ALTS(r1, r2) => atMostEmpty(r1) && atMostEmpty(r2)
 | 
|  |    890 |   case CHAR(_) => false
 | 
|  |    891 | }
 | 
|  |    892 | 
 | 
|  |    893 | def isOne(r: Rexp) : Boolean = r match {
 | 
|  |    894 |   case ONE => true
 | 
|  |    895 |   case SEQ(r1, r2) => isOne(r1) && isOne(r2)
 | 
|  |    896 |   case ALTS(r1, r2) => (isOne(r1) || isOne(r2)) && (atMostEmpty(r1) && atMostEmpty(r2))//rs.forall(atMostEmpty) && rs.exists(isOne)
 | 
|  |    897 |   case STAR(r0) => atMostEmpty(r0)
 | 
|  |    898 |   case CHAR(c) => false
 | 
|  |    899 |   case ZERO => false
 | 
|  |    900 | 
 | 
|  |    901 | }
 | 
|  |    902 | 
 | 
| 543 |    903 | def isOne1(r: Rexp) : Boolean = r match {
 | 
|  |    904 |   case ONE => true
 | 
|  |    905 |   case SEQ(r1, r2) => false//isOne1(r1) && isOne1(r2)
 | 
|  |    906 |   case ALTS(r1, r2) => false//(isOne1(r1) || isOne1(r2)) && (atMostEmpty(r1) && atMostEmpty(r2))//rs.forall(atMostEmpty) && rs.exists(isOne)
 | 
|  |    907 |   case STAR(r0) => false//atMostEmpty(r0)
 | 
|  |    908 |   case CHAR(c) => false
 | 
|  |    909 |   case ZERO => false
 | 
|  |    910 | 
 | 
|  |    911 | }
 | 
|  |    912 | 
 | 
| 533 |    913 | def turnIntoTerms(r: Rexp): List[Rexp] = r match {
 | 
| 543 |    914 |   case SEQ(r1, r2)  => if(isOne1(r1)) 
 | 
| 533 |    915 |                           turnIntoTerms(r2) 
 | 
| 526 |    916 |                        else 
 | 
| 543 |    917 |                           turnIntoTerms(r1).flatMap(r11 => furtherSEQ(r11, r2))
 | 
| 533 |    918 |   case ALTS(r1, r2) => turnIntoTerms(r1) ::: turnIntoTerms(r2)
 | 
| 526 |    919 |   case ZERO => Nil
 | 
|  |    920 |   case _ => r :: Nil
 | 
|  |    921 | }
 | 
|  |    922 | 
 | 
| 543 |    923 | def furtherSEQ(r11: Rexp, r2: Rexp) : List[Rexp] = {
 | 
|  |    924 |   if(r11 == ONE)
 | 
|  |    925 |     turnIntoTerms(r2)
 | 
|  |    926 |   else
 | 
|  |    927 |     SEQ(r11, r2) :: Nil
 | 
| 518 |    928 | }
 | 
|  |    929 | 
 | 
| 543 |    930 | def attachCtx(r: ARexp, ctx: List[Rexp]) : Set[Rexp] = {
 | 
| 564 |    931 |   turnIntoTerms((seqFold(erase(r), ctx)))
 | 
| 543 |    932 |     .toSet
 | 
| 541 |    933 | }
 | 
|  |    934 | 
 | 
| 543 |    935 | def attachCtxcc(r: Rexp, ctx: List[Rexp]) : Set[Rexp] =
 | 
| 564 |    936 |   turnIntoTerms(seqFold(r, ctx)).toSet
 | 
| 543 |    937 | 
 | 
|  |    938 | def ABIncludedByC[A, B, C](a: A, b: B, c: C, f: (A, B) => C, 
 | 
|  |    939 | subseteqPred: (C, C) => Boolean) : Boolean = {
 | 
| 530 |    940 |   subseteqPred(f(a, b), c)
 | 
| 518 |    941 | }
 | 
| 543 |    942 | def rs1_subseteq_rs2(rs1: Set[Rexp], rs2: Set[Rexp]) : Boolean = 
 | 
|  |    943 |   //rs1 \subseteq rs2
 | 
|  |    944 |   rs1.forall(rs2.contains)
 | 
|  |    945 | 
 | 
|  |    946 | def prune6(acc: Set[Rexp], r: ARexp, ctx: List[Rexp] = Nil) : ARexp = {
 | 
| 541 |    947 |   if (ABIncludedByC(a = r, b = ctx, c = acc, 
 | 
|  |    948 |                     f = attachCtx, subseteqPred = rs1_subseteq_rs2)) 
 | 
| 543 |    949 |   {
 | 
| 518 |    950 |     AZERO
 | 
|  |    951 |   }
 | 
|  |    952 |   else{
 | 
|  |    953 |     r match {
 | 
| 543 |    954 |       case ASEQ(bs, r1, r2) => (prune6(acc, r1, erase(r2) :: ctx)) match{
 | 
| 518 |    955 |         case AZERO => 
 | 
|  |    956 |           AZERO
 | 
| 543 |    957 |         case AONE(bs1) => //when r1 becomes 1, chances to further prune r2
 | 
|  |    958 |           prune6(acc, fuse(bs1, r2), ctx)
 | 
| 518 |    959 |         case r1p => 
 | 
|  |    960 |           ASEQ(bs, r1p, r2)
 | 
|  |    961 |       }
 | 
|  |    962 |       case AALTS(bs, rs0) => 
 | 
| 530 |    963 |         rs0.map(r => prune6(acc, r, ctx)).filter(_ != AZERO) match {
 | 
| 518 |    964 |           case Nil => 
 | 
|  |    965 |             AZERO
 | 
|  |    966 |           case r :: Nil => 
 | 
| 543 |    967 |             fuse(bs, r) 
 | 
| 518 |    968 |           case rs0p => 
 | 
|  |    969 |             AALTS(bs, rs0p)
 | 
|  |    970 |         }
 | 
|  |    971 |       case r => r
 | 
|  |    972 |     }
 | 
|  |    973 |   }
 | 
|  |    974 | }
 | 
|  |    975 | 
 | 
| 541 |    976 | def prune6cc(acc: Set[Rexp], r: Rexp, ctx: List[Rexp]) : Rexp = {
 | 
|  |    977 |   if (ABIncludedByC(a = r, b = ctx, c = acc, 
 | 
|  |    978 |                     f = attachCtxcc, subseteqPred = rs1_subseteq_rs2)) 
 | 
|  |    979 |   {//acc.flatMap(breakIntoTerms
 | 
|  |    980 |     ZERO
 | 
|  |    981 |   }
 | 
|  |    982 |   else{
 | 
|  |    983 |     r match {
 | 
|  |    984 |       case SEQ(r1, r2) => 
 | 
|  |    985 |       (prune6cc(acc, r1, r2 :: ctx)) match{
 | 
|  |    986 |         case ZERO => 
 | 
|  |    987 |           ZERO
 | 
|  |    988 |         case ONE => 
 | 
|  |    989 |           r2
 | 
|  |    990 |         case r1p => 
 | 
|  |    991 |           SEQ(r1p, r2)
 | 
|  |    992 |       }
 | 
|  |    993 |       case ALTS(r1, r2) => 
 | 
|  |    994 |         List(r1, r2).map(r => prune6cc(acc, r, ctx)).filter(_ != AZERO) match {
 | 
|  |    995 |           case Nil => 
 | 
|  |    996 |             ZERO
 | 
|  |    997 |           case r :: Nil => 
 | 
|  |    998 |             r 
 | 
|  |    999 |           case ra :: rb :: Nil => 
 | 
|  |   1000 |             ALTS(ra, rb)
 | 
|  |   1001 |         }
 | 
|  |   1002 |       case r => r
 | 
|  |   1003 |     }
 | 
|  |   1004 |   }
 | 
|  |   1005 | }
 | 
| 518 |   1006 | 
 | 
| 543 |   1007 | def distinctBy6(xs: List[ARexp], acc: Set[Rexp] = Set()) : 
 | 
|  |   1008 | List[ARexp] = xs match {
 | 
| 518 |   1009 |   case Nil => 
 | 
|  |   1010 |     Nil
 | 
|  |   1011 |   case x :: xs => {
 | 
|  |   1012 |     val erased = erase(x)
 | 
|  |   1013 |     if(acc.contains(erased)){
 | 
|  |   1014 |       distinctBy6(xs, acc)
 | 
|  |   1015 |     }
 | 
|  |   1016 |     else{
 | 
| 543 |   1017 |       val pruned = prune6(acc, x)
 | 
| 533 |   1018 |       val newTerms = turnIntoTerms(erase(pruned))
 | 
| 518 |   1019 |       pruned match {
 | 
|  |   1020 |         case AZERO => 
 | 
|  |   1021 |           distinctBy6(xs, acc)
 | 
|  |   1022 |         case xPrime => 
 | 
| 543 |   1023 |           xPrime :: distinctBy6(xs, newTerms ++: acc)//distinctBy5(xs, addToAcc.map(oneSimp(_)) ::: acc)
 | 
| 518 |   1024 |       }
 | 
|  |   1025 |     }
 | 
|  |   1026 |   }
 | 
|  |   1027 | }
 | 
| 431 |   1028 | 
 | 
| 541 |   1029 | def distinctByacc(xs: List[Rexp], acc: Set[Rexp] = Set()) : Set[Rexp] = xs match {
 | 
|  |   1030 |   case Nil => 
 | 
|  |   1031 |     acc
 | 
|  |   1032 |   case x :: xs => {
 | 
|  |   1033 |     if(acc.contains(x)){
 | 
|  |   1034 |       distinctByacc(xs, acc)
 | 
|  |   1035 |     }
 | 
|  |   1036 |     else{
 | 
|  |   1037 |       val pruned = prune6cc(acc, x, Nil)
 | 
|  |   1038 |       val newTerms = turnIntoTerms(pruned)
 | 
|  |   1039 |       pruned match {
 | 
|  |   1040 |         case ZERO => 
 | 
|  |   1041 |           distinctByacc(xs, acc)
 | 
|  |   1042 |         case xPrime => 
 | 
|  |   1043 |           distinctByacc(xs, newTerms.map(oneSimp) ++: acc)//distinctBy5(xs, addToAcc.map(oneSimp(_)) ::: acc)
 | 
|  |   1044 |       }
 | 
|  |   1045 |     }
 | 
|  |   1046 |   }
 | 
|  |   1047 | }
 | 
|  |   1048 | 
 | 
| 492 |   1049 | def breakIntoTerms(r: Rexp) : List[Rexp] = r match {
 | 
|  |   1050 |   case SEQ(r1, r2)  => breakIntoTerms(r1).map(r11 => SEQ(r11, r2))
 | 
|  |   1051 |   case ALTS(r1, r2) => breakIntoTerms(r1) ::: breakIntoTerms(r2)
 | 
| 494 |   1052 |   case ZERO => Nil
 | 
| 492 |   1053 |   case _ => r::Nil
 | 
|  |   1054 | }
 | 
| 431 |   1055 | 
 | 
| 543 |   1056 | def flatsIntoTerms(r: Rexp) : List[Rexp] = r match {
 | 
|  |   1057 |   //case SEQ(r1, r2)  => flatsIntoTerms(r1).map(r11 => SEQ(r11, r2))
 | 
|  |   1058 |   case ALTS(r1, r2) => flatsIntoTerms(r1) ::: flatsIntoTerms(r2)
 | 
|  |   1059 |   case ZERO => Nil
 | 
|  |   1060 |   case _ => r::Nil
 | 
|  |   1061 | }
 | 
| 431 |   1062 | 
 | 
|  |   1063 | 
 | 
| 492 |   1064 | def decode_aux(r: Rexp, bs: Bits) : (Val, Bits) = (r, bs) match {
 | 
|  |   1065 |   case (ONE, bs) => (Empty, bs)
 | 
|  |   1066 |   case (CHAR(f), bs) => (Chr(f), bs)
 | 
|  |   1067 |   case (ALTS(r1, r2), Z::bs1) => {
 | 
|  |   1068 |       val (v, bs2) = decode_aux(r1, bs1)
 | 
|  |   1069 |       (Left(v), bs2)
 | 
|  |   1070 |   }
 | 
|  |   1071 |   case (ALTS(r1, r2), S::bs1) => {
 | 
|  |   1072 |       val (v, bs2) = decode_aux(r2, bs1)
 | 
|  |   1073 |       (Right(v), bs2)
 | 
|  |   1074 |   }
 | 
|  |   1075 |   case (SEQ(r1, r2), bs) => {
 | 
|  |   1076 |     val (v1, bs1) = decode_aux(r1, bs)
 | 
|  |   1077 |     val (v2, bs2) = decode_aux(r2, bs1)
 | 
|  |   1078 |     (Sequ(v1, v2), bs2)
 | 
|  |   1079 |   }
 | 
|  |   1080 |   case (STAR(r1), S::bs) => {
 | 
|  |   1081 |     val (v, bs1) = decode_aux(r1, bs)
 | 
| 494 |   1082 |     //(v)
 | 
| 492 |   1083 |     val (Stars(vs), bs2) = decode_aux(STAR(r1), bs1)
 | 
|  |   1084 |     (Stars(v::vs), bs2)
 | 
|  |   1085 |   }
 | 
|  |   1086 |   case (STAR(_), Z::bs) => (Stars(Nil), bs)
 | 
|  |   1087 |   case (RECD(x, r1), bs) => {
 | 
|  |   1088 |     val (v, bs1) = decode_aux(r1, bs)
 | 
|  |   1089 |     (Rec(x, v), bs1)
 | 
|  |   1090 |   }
 | 
|  |   1091 |   case (NOT(r), bs) => (Nots(r.toString), bs)
 | 
| 431 |   1092 | }
 | 
|  |   1093 | 
 | 
| 492 |   1094 | def decode(r: Rexp, bs: Bits) = decode_aux(r, bs) match {
 | 
|  |   1095 |   case (v, Nil) => v
 | 
|  |   1096 |   case _ => throw new Exception("Not decodable")
 | 
|  |   1097 | }
 | 
|  |   1098 | 
 | 
|  |   1099 | 
 | 
|  |   1100 | 
 | 
| 431 |   1101 | def blexing_simp(r: Rexp, s: String) : Val = {
 | 
|  |   1102 |     val bit_code = blex_simp(internalise(r), s.toList)
 | 
|  |   1103 |     decode(r, bit_code)
 | 
| 492 |   1104 | }
 | 
|  |   1105 | def simpBlexer(r: Rexp, s: String) : Val = {
 | 
|  |   1106 |   Try(blexing_simp(r, s)).getOrElse(Failure)
 | 
|  |   1107 | }
 | 
| 431 |   1108 | 
 | 
| 492 |   1109 | def strong_blexing_simp(r: Rexp, s: String) : Val = {
 | 
|  |   1110 |   decode(r, strong_blex_simp(internalise(r), s.toList))
 | 
|  |   1111 | }
 | 
|  |   1112 | 
 | 
| 494 |   1113 | def strong_blexing_simp5(r: Rexp, s: String) : Val = {
 | 
|  |   1114 |   decode(r, strong_blex_simp5(internalise(r), s.toList))
 | 
|  |   1115 | }
 | 
|  |   1116 | 
 | 
|  |   1117 | 
 | 
| 492 |   1118 | def strongBlexer(r: Rexp, s: String) : Val = {
 | 
|  |   1119 |   Try(strong_blexing_simp(r, s)).getOrElse(Failure)
 | 
|  |   1120 | }
 | 
| 431 |   1121 | 
 | 
| 494 |   1122 | def strongBlexer5(r: Rexp, s: String): Val = {
 | 
|  |   1123 |   Try(strong_blexing_simp5(r, s)).getOrElse(Failure)
 | 
|  |   1124 | }
 | 
|  |   1125 | 
 | 
| 492 |   1126 | def strong_blex_simp(r: ARexp, s: List[Char]) : Bits = s match {
 | 
|  |   1127 |   case Nil => {
 | 
|  |   1128 |     if (bnullable(r)) {
 | 
|  |   1129 |       //println(asize(r))
 | 
|  |   1130 |       val bits = mkepsBC(r)
 | 
|  |   1131 |       bits
 | 
| 431 |   1132 |     }
 | 
| 492 |   1133 |     else 
 | 
|  |   1134 |       throw new Exception("Not matched")
 | 
| 431 |   1135 |   }
 | 
| 492 |   1136 |   case c::cs => {
 | 
|  |   1137 |     val der_res = bder(c,r)
 | 
|  |   1138 |     val simp_res = strongBsimp(der_res)  
 | 
|  |   1139 |     strong_blex_simp(simp_res, cs)      
 | 
|  |   1140 |   }
 | 
|  |   1141 | }
 | 
| 431 |   1142 | 
 | 
| 494 |   1143 | def strong_blex_simp5(r: ARexp, s: List[Char]) : Bits = s match {
 | 
|  |   1144 |   case Nil => {
 | 
|  |   1145 |     if (bnullable(r)) {
 | 
|  |   1146 |       val bits = mkepsBC(r)
 | 
|  |   1147 |       bits
 | 
|  |   1148 |     }
 | 
|  |   1149 |     else 
 | 
|  |   1150 |       throw new Exception("Not matched")
 | 
|  |   1151 |   }
 | 
|  |   1152 |   case c::cs => {
 | 
|  |   1153 |     val der_res = bder(c,r)
 | 
|  |   1154 |     val simp_res = strongBsimp5(der_res)  
 | 
|  |   1155 |     strong_blex_simp5(simp_res, cs)      
 | 
|  |   1156 |   }
 | 
|  |   1157 | }
 | 
| 431 |   1158 | 
 | 
|  |   1159 | 
 | 
|  |   1160 |   def bders_simp(s: List[Char], r: ARexp) : ARexp = s match {
 | 
|  |   1161 |     case Nil => r
 | 
| 435 |   1162 |     case c::s => 
 | 
| 494 |   1163 |       //println(erase(r))
 | 
| 435 |   1164 |       bders_simp(s, bsimp(bder(c, r)))
 | 
| 431 |   1165 |   }
 | 
|  |   1166 |   
 | 
| 494 |   1167 |   def bdersStrong5(s: List[Char], r: ARexp) : ARexp = s match {
 | 
|  |   1168 |     case Nil => r
 | 
|  |   1169 |     case c::s => bdersStrong5(s, strongBsimp5(bder(c, r)))
 | 
|  |   1170 |   }
 | 
| 518 |   1171 |   def bdersStrong6(s: List[Char], r: ARexp) : ARexp = s match {
 | 
|  |   1172 |     case Nil => r
 | 
|  |   1173 |     case c::s => bdersStrong6(s, strongBsimp6(bder(c, r)))
 | 
|  |   1174 |   }
 | 
| 532 |   1175 |   def bdersStrong7(s: List[Char], r: ARexp) : ARexp = s match {
 | 
|  |   1176 |     case Nil => r
 | 
|  |   1177 |     case c::s => bdersStrong7(s, strongBsimp7(bder(c, r)))
 | 
|  |   1178 |   }
 | 
| 431 |   1179 |   def bdersSimp(s: String, r: Rexp) : ARexp = bders_simp(s.toList, internalise(r))
 | 
|  |   1180 | 
 | 
| 492 |   1181 |   def bdersStrong(s: List[Char], r: ARexp) : ARexp = s match {
 | 
| 431 |   1182 |     case Nil => r 
 | 
| 492 |   1183 |     case c::s => bdersStrong(s, strongBsimp(bder(c, r)))
 | 
| 431 |   1184 |   }
 | 
|  |   1185 | 
 | 
| 492 |   1186 |   def bdersStrongRexp(s: String, r: Rexp) : ARexp = bdersStrong(s.toList, internalise(r))
 | 
| 431 |   1187 | 
 | 
|  |   1188 |   def erase(r:ARexp): Rexp = r match{
 | 
|  |   1189 |     case AZERO => ZERO
 | 
|  |   1190 |     case AONE(_) => ONE
 | 
|  |   1191 |     case ACHAR(bs, c) => CHAR(c)
 | 
|  |   1192 |     case AALTS(bs, Nil) => ZERO
 | 
|  |   1193 |     case AALTS(bs, a::Nil) => erase(a)
 | 
|  |   1194 |     case AALTS(bs, a::as) => ALTS(erase(a), erase(AALTS(bs, as)))
 | 
|  |   1195 |     case ASEQ(bs, r1, r2) => SEQ (erase(r1), erase(r2))
 | 
|  |   1196 |     case ASTAR(cs, r)=> STAR(erase(r))
 | 
|  |   1197 |     case ANOT(bs, r) => NOT(erase(r))
 | 
|  |   1198 |     case AANYCHAR(bs) => ANYCHAR
 | 
|  |   1199 |   }
 | 
|  |   1200 | 
 | 
|  |   1201 | 
 | 
| 492 |   1202 |   def allCharSeq(r: Rexp) : Boolean = r match {
 | 
|  |   1203 |     case CHAR(c) => true
 | 
|  |   1204 |     case SEQ(r1, r2) => allCharSeq(r1) && allCharSeq(r2)
 | 
|  |   1205 |     case _ => false
 | 
|  |   1206 |   }
 | 
|  |   1207 | 
 | 
|  |   1208 |   def flattenSeq(r: Rexp) : String = r match {
 | 
|  |   1209 |     case CHAR(c) => c.toString
 | 
|  |   1210 |     case SEQ(r1, r2) => flattenSeq(r1) ++ flattenSeq(r2)
 | 
|  |   1211 |     case _ => throw new Error("flatten unflattenable rexp")
 | 
|  |   1212 |   } 
 | 
| 431 |   1213 | 
 | 
| 492 |   1214 |   def shortRexpOutput(r: Rexp) : String = r match {
 | 
|  |   1215 |       case CHAR(c) => c.toString
 | 
|  |   1216 |       case ONE => "1"
 | 
|  |   1217 |       case ZERO => "0"
 | 
|  |   1218 |       case SEQ(r1, r2) if(allCharSeq(r)) => flattenSeq(r)//"[" ++ shortRexpOutput(r1) ++ "~" ++ shortRexpOutput(r2) ++ "]"
 | 
|  |   1219 |       case SEQ(r1, r2) => "[" ++ shortRexpOutput(r1) ++ "~" ++ shortRexpOutput(r2) ++ "]"
 | 
|  |   1220 |       case ALTS(r1, r2) => "(" ++ shortRexpOutput(r1) ++ "+" ++ shortRexpOutput(r2) ++ ")"
 | 
|  |   1221 |       case STAR(STAR(r)) => "(..)*"// ++ shortRexpOutput(r) ++ "]*"
 | 
|  |   1222 |       case STAR(r) => "STAR(" ++ shortRexpOutput(r) ++ ")"
 | 
|  |   1223 |       //case RTOP => "RTOP"
 | 
|  |   1224 |     }
 | 
|  |   1225 | 
 | 
| 431 |   1226 | 
 | 
| 492 |   1227 |   def blex_simp(r: ARexp, s: List[Char]) : Bits = s match {
 | 
|  |   1228 |       case Nil => {
 | 
|  |   1229 |         if (bnullable(r)) {
 | 
|  |   1230 |           val bits = mkepsBC(r)
 | 
|  |   1231 |           bits
 | 
|  |   1232 |         }
 | 
|  |   1233 |         else 
 | 
|  |   1234 |           throw new Exception("Not matched")
 | 
|  |   1235 |       }
 | 
|  |   1236 |       case c::cs => {
 | 
|  |   1237 |         val der_res = bder(c,r)
 | 
|  |   1238 |         val simp_res = bsimp(der_res)  
 | 
|  |   1239 |         blex_simp(simp_res, cs)      
 | 
|  |   1240 |       }
 | 
| 431 |   1241 |   }
 | 
|  |   1242 | 
 | 
|  |   1243 | 
 | 
| 492 |   1244 | 
 | 
|  |   1245 | 
 | 
|  |   1246 |     def size(r: Rexp) : Int = r match {
 | 
|  |   1247 |       case ZERO => 1
 | 
|  |   1248 |       case ONE => 1
 | 
|  |   1249 |       case CHAR(_) => 1
 | 
|  |   1250 |       case ANYCHAR => 1
 | 
|  |   1251 |       case NOT(r0) => 1 + size(r0)
 | 
|  |   1252 |       case SEQ(r1, r2) => 1 + size(r1) + size(r2)
 | 
|  |   1253 |       case ALTS(r1, r2) => 1 + List(r1, r2).map(size).sum
 | 
|  |   1254 |       case STAR(r) => 1 + size(r)
 | 
| 431 |   1255 |     }
 | 
|  |   1256 | 
 | 
| 492 |   1257 |     def asize(a: ARexp) = size(erase(a))
 | 
| 431 |   1258 | 
 | 
|  |   1259 | //pder related
 | 
|  |   1260 | type Mon = (Char, Rexp)
 | 
|  |   1261 | type Lin = Set[Mon]
 | 
|  |   1262 | 
 | 
|  |   1263 | def dot_prod(rs: Set[Rexp], r: Rexp): Set[Rexp] = r match {
 | 
|  |   1264 |     case ZERO => Set()
 | 
|  |   1265 |     case ONE => rs
 | 
|  |   1266 |     case r => rs.map((re) => if (re == ONE) r else SEQ(re, r)  )   
 | 
|  |   1267 |   }
 | 
|  |   1268 |   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
 | 
|  |   1269 |     case ZERO => Set()
 | 
| 543 |   1270 |     // case ONE => l
 | 
| 526 |   1271 |     case t => l.map( m => m._2 match 
 | 
|  |   1272 |       {
 | 
|  |   1273 |         case ZERO => m 
 | 
|  |   1274 |         case ONE => (m._1, t) 
 | 
|  |   1275 |         case p => (m._1, SEQ(p, t)) 
 | 
|  |   1276 |       }  
 | 
|  |   1277 |     
 | 
|  |   1278 |     )
 | 
| 431 |   1279 |   }
 | 
| 543 |   1280 |   def cir_prodList(l : List[Mon], t: Rexp): List[Mon] = t match {
 | 
|  |   1281 |     case ZERO => Nil
 | 
|  |   1282 |     case ONE => l
 | 
|  |   1283 |     case t => l.map(m => m._2 match
 | 
|  |   1284 |       {
 | 
|  |   1285 |         case ZERO => m
 | 
|  |   1286 |         case ONE => (m._1, t)
 | 
|  |   1287 |         case p => (m._1, SEQ(p, t))
 | 
|  |   1288 |       }
 | 
|  |   1289 |     )
 | 
|  |   1290 | 
 | 
|  |   1291 |   }
 | 
|  |   1292 |   def lfList(r: Rexp) : List[Mon] = r match {
 | 
|  |   1293 |     case ZERO => Nil
 | 
|  |   1294 |     case ONE => Nil
 | 
|  |   1295 |     case CHAR(f) => {
 | 
|  |   1296 |       (f, ONE) :: Nil
 | 
|  |   1297 |     }
 | 
|  |   1298 |     case ALTS(r1, r2) => {
 | 
|  |   1299 |       lfList(r1) ++ lfList(r2)
 | 
|  |   1300 |     }
 | 
|  |   1301 |     case STAR(r1) => cir_prodList(lfList(r1), STAR(r1))
 | 
|  |   1302 |     case SEQ(r1, r2) => {
 | 
|  |   1303 |       if(nullable(r1))
 | 
|  |   1304 |         cir_prodList(lfList(r1), r2) ++ lfList(r2)
 | 
|  |   1305 |       else
 | 
|  |   1306 |         cir_prodList(lfList(r1), r2)
 | 
|  |   1307 |     }
 | 
|  |   1308 |   }
 | 
|  |   1309 |   def lfprint(ls: Lin) = ls.foreach(m =>{
 | 
|  |   1310 |     println(m._1)
 | 
|  |   1311 |     rprint(m._2)
 | 
|  |   1312 |     })
 | 
| 431 |   1313 |   def lf(r: Rexp): Lin = r match {
 | 
|  |   1314 |     case ZERO => Set()
 | 
|  |   1315 |     case ONE => Set()
 | 
|  |   1316 |     case CHAR(f) => {
 | 
|  |   1317 |       //val Some(c) = alphabet.find(f) 
 | 
|  |   1318 |       Set((f, ONE))
 | 
|  |   1319 |     }
 | 
|  |   1320 |     case ALTS(r1, r2) => {
 | 
|  |   1321 |       lf(r1 ) ++ lf(r2)
 | 
|  |   1322 |     }
 | 
|  |   1323 |     case STAR(r1) => cir_prod(lf(r1), STAR(r1)) //may try infix notation later......
 | 
|  |   1324 |     case SEQ(r1, r2) =>{
 | 
|  |   1325 |       if (nullable(r1))
 | 
|  |   1326 |         cir_prod(lf(r1), r2) ++ lf(r2)
 | 
|  |   1327 |       else
 | 
|  |   1328 |         cir_prod(lf(r1), r2) 
 | 
|  |   1329 |     }
 | 
|  |   1330 |   }
 | 
|  |   1331 |   def lfs(r: Set[Rexp]): Lin = {
 | 
|  |   1332 |     r.foldLeft(Set[Mon]())((acc, r) => acc ++ lf(r))
 | 
|  |   1333 |   }
 | 
|  |   1334 | 
 | 
|  |   1335 |   def pder(x: Char, t: Rexp): Set[Rexp] = {
 | 
|  |   1336 |     val lft = lf(t)
 | 
|  |   1337 |     (lft.filter(mon => if(mon._1 == x) true else false)).map(mon => mon._2)
 | 
|  |   1338 |   }
 | 
|  |   1339 |   def pders_single(s: List[Char], t: Rexp) : Set[Rexp] = s match {
 | 
|  |   1340 |     case x::xs => pders(xs, pder(x, t))
 | 
|  |   1341 |     case Nil => Set(t)
 | 
|  |   1342 |   }
 | 
|  |   1343 |   def pders(s: List[Char], ts: Set[Rexp]) : Set[Rexp] = s match {
 | 
|  |   1344 |     case x::xs => pders(xs, ts.foldLeft(Set[Rexp]())((acc, t) => acc ++ pder(x, t)))
 | 
|  |   1345 |     case Nil => ts 
 | 
|  |   1346 |   }
 | 
| 526 |   1347 |   def pderss(ss: List[List[Char]], t: Rexp): Set[Rexp] = 
 | 
|  |   1348 |     ss.foldLeft( Set[Rexp]() )( (acc, s) => pders_single(s, t) ++ acc )
 | 
| 431 |   1349 |   def pdera(t: Rexp): Set[Rexp] = lf(t).map(mon => mon._2)
 | 
|  |   1350 |   //all implementation of partial derivatives that involve set union are potentially buggy
 | 
|  |   1351 |   //because they don't include the original regular term before they are pdered.
 | 
|  |   1352 |   //now only pderas is fixed.
 | 
| 526 |   1353 |   def pderas(t: Set[Rexp], d: Int): Set[Rexp] = 
 | 
|  |   1354 |     if(d > 0) 
 | 
|  |   1355 |       pderas(lfs(t).map(mon => mon._2), d - 1) ++ t 
 | 
|  |   1356 |     else 
 | 
|  |   1357 |       lfs(t).map(mon => mon._2) ++ t//repeated application of pderas over the newest set of pders.
 | 
| 431 |   1358 |   def pderUNIV(r: Rexp) : Set[Rexp] = pderas(Set(r), awidth(r) + 1)
 | 
|  |   1359 |   def awidth(r: Rexp) : Int = r match {
 | 
|  |   1360 |     case CHAR(c) => 1
 | 
|  |   1361 |     case SEQ(r1, r2) => awidth(r1) + awidth(r2)
 | 
|  |   1362 |     case ALTS(r1, r2) => awidth(r1) + awidth(r2)
 | 
|  |   1363 |     case ONE => 0
 | 
|  |   1364 |     case ZERO => 0
 | 
|  |   1365 |     case STAR(r) => awidth(r)
 | 
|  |   1366 |   }
 | 
|  |   1367 |   //def sigma_lf(l: Set[Mon]) : Rexp = ALTS(l.map(mon => SEQ(CHAR(mon._1),mon._2)).toList)
 | 
|  |   1368 |   //def sigma(rs: Set[Rexp]) : Rexp = ALTS(rs.toList)
 | 
|  |   1369 |   def o(r: Rexp) = if (nullable(r)) ONE else ZERO
 | 
|  |   1370 |   //def nlf(t: Rexp) : Rexp = ALTS(List( o(t), sigma_lf(lf(t)) ))
 | 
|  |   1371 |   def pdp(x: Char, r: Rexp) : Set[Rexp] = r match {
 | 
|  |   1372 |     case ZERO => Set[Rexp]()
 | 
|  |   1373 |     case ONE => Set[Rexp]()
 | 
|  |   1374 |     case CHAR(f) => if(x == f) Set(ONE) else Set[Rexp]()
 | 
|  |   1375 |     case ALTS(r1, r2) => pdp(x, r1) ++ pdp(x, r2)
 | 
|  |   1376 |     case STAR(r1) => pdp(x, r).map(a => SEQ(a, STAR(r1)))
 | 
|  |   1377 |     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))
 | 
|  |   1378 |   }
 | 
|  |   1379 |   def pdps(s: List[Char], ts: Set[Rexp]): Set[Rexp] = s match {
 | 
|  |   1380 |     case x::xs => pdps(xs, ts.foldLeft(Set[Rexp]())((acc, t) => acc ++ pder(x, t)))
 | 
|  |   1381 |     case Nil => ts   
 | 
|  |   1382 |   }
 | 
|  |   1383 |   def pdpss(ss: List[List[Char]], t: Rexp): Set[Rexp] = ss.foldLeft( Set[Rexp]())((acc, s) => pdps(s, Set(t)) ++ acc)
 | 
|  |   1384 | 
 | 
|  |   1385 | 
 | 
|  |   1386 | 
 | 
|  |   1387 | def starPrint(r: Rexp) : Unit = r match {
 | 
|  |   1388 |         
 | 
|  |   1389 |           case SEQ(head, rstar) =>
 | 
|  |   1390 |             println(shortRexpOutput(head) ++ "~STARREG")
 | 
|  |   1391 |           case STAR(rstar) =>
 | 
|  |   1392 |             println("STARREG")
 | 
|  |   1393 |           case ALTS(r1, r2) =>  
 | 
|  |   1394 |             println("(")
 | 
|  |   1395 |             starPrint(r1)
 | 
|  |   1396 |             println("+")
 | 
|  |   1397 |             starPrint(r2)
 | 
|  |   1398 |             println(")")
 | 
|  |   1399 |           case ZERO => println("0")
 | 
|  |   1400 |       }
 | 
|  |   1401 | 
 | 
|  |   1402 | // @arg(doc = "small tests")
 | 
| 516 |   1403 | def n_astar_list(d: Int) : Rexp = {
 | 
|  |   1404 |   if(d == 0) STAR("a") 
 | 
|  |   1405 |   else ALTS(STAR("a" * d), n_astar_list(d - 1))
 | 
|  |   1406 | }
 | 
|  |   1407 | def n_astar_alts(d: Int) : Rexp = d match {
 | 
|  |   1408 |   case 0 => n_astar_list(0)
 | 
|  |   1409 |   case d => STAR(n_astar_list(d))
 | 
| 533 |   1410 |   }
 | 
|  |   1411 | def n_astar_alts2(d: Int) : Rexp = d match {
 | 
|  |   1412 |   case 0 => n_astar_list(0)
 | 
|  |   1413 |   case d => STAR(STAR(n_astar_list(d)))
 | 
|  |   1414 |   }
 | 
| 516 |   1415 | def n_astar_aux(d: Int) = {
 | 
|  |   1416 |   if(d == 0) n_astar_alts(0)
 | 
|  |   1417 |   else ALTS(n_astar_alts(d), n_astar_alts(d - 1))
 | 
|  |   1418 | }
 | 
|  |   1419 | 
 | 
|  |   1420 | def n_astar(d: Int) : Rexp = STAR(n_astar_aux(d))
 | 
|  |   1421 | //val STARREG = n_astar(3)
 | 
|  |   1422 | // ( STAR("a") | 
 | 
|  |   1423 | //                  ("a" | "aa").% | 
 | 
|  |   1424 | //                 ( "a" | "aa" | "aaa").% 
 | 
|  |   1425 | //                 ).%
 | 
|  |   1426 |                 //( "a" | "aa" | "aaa" | "aaaa").% |
 | 
|  |   1427 |                 //( "a" | "aa" | "aaa" | "aaaa" | "aaaaa").% 
 | 
| 431 |   1428 | (((STAR("a") | ( STAR("aaa")) | STAR("aaaaa"| ( STAR("aaaaaaa")) | STAR("aaaaaaaaaaa"))).%).%).%
 | 
|  |   1429 | 
 | 
|  |   1430 | // @main
 | 
| 492 |   1431 | 
 | 
| 516 |   1432 | def lcm(list: Seq[Int]):Int=list.foldLeft(1:Int){
 | 
|  |   1433 |   (a, b) => b * a /
 | 
|  |   1434 |   Stream.iterate((a,b)){case (x,y) => (y, x%y)}.dropWhile(_._2 != 0).head._1.abs
 | 
|  |   1435 | }
 | 
| 492 |   1436 | 
 | 
| 431 |   1437 | def small() = {
 | 
| 516 |   1438 |   //val pderSTAR = pderUNIV(STARREG)
 | 
| 492 |   1439 | 
 | 
| 516 |   1440 |   //val refSize = pderSTAR.map(size(_)).sum
 | 
| 518 |   1441 |   for(n <- 5 to 5){
 | 
| 533 |   1442 |     val STARREG = n_astar_alts2(n)
 | 
| 516 |   1443 |     val iMax = (lcm((1 to n).toList))
 | 
| 533 |   1444 |     for(i <- 0 to iMax ){// 100, 400, 800, 840, 841, 900 
 | 
| 516 |   1445 |       val prog0 = "a" * i
 | 
|  |   1446 |       //println(s"test: $prog0")
 | 
| 518 |   1447 |       print(i)
 | 
| 516 |   1448 |       print(" ")
 | 
|  |   1449 |       // print(i)
 | 
|  |   1450 |       // print(" ")
 | 
|  |   1451 |       println(asize(bders_simp(prog0.toList, internalise(STARREG))))
 | 
| 533 |   1452 |       // println(asize(bdersStrong7(prog0.toList, internalise(STARREG))))
 | 
| 516 |   1453 |     }
 | 
| 539 |   1454 |     
 | 
| 492 |   1455 |   }
 | 
| 431 |   1456 | }
 | 
| 536 |   1457 | // small()
 | 
|  |   1458 | 
 | 
|  |   1459 | def aaa_star() = {
 | 
|  |   1460 |   val r = STAR(("a" | "aa"))
 | 
| 543 |   1461 |   for(i <- 0 to 20) {
 | 
| 536 |   1462 |     val prog = "a" * i
 | 
| 543 |   1463 |     print(i+" ")
 | 
| 536 |   1464 |     println(asize(bders_simp(prog.toList, internalise(r))))
 | 
| 543 |   1465 |     //println(size(ders_simp(prog.toList, r)))
 | 
| 536 |   1466 |   }
 | 
|  |   1467 | }
 | 
| 539 |   1468 | // aaa_star()
 | 
|  |   1469 | def naive_matcher() {
 | 
|  |   1470 |   val r = STAR(STAR("a") ~ STAR("a"))
 | 
| 536 |   1471 | 
 | 
| 539 |   1472 |   for(i <- 0 to 20) {
 | 
|  |   1473 |     val s = "a" * i 
 | 
|  |   1474 |     val t1 = System.nanoTime
 | 
|  |   1475 |     matcher(s, r)
 | 
|  |   1476 |     val duration = (System.nanoTime - t1) / 1e9d
 | 
|  |   1477 |     print(i)
 | 
|  |   1478 |     print(" ")
 | 
|  |   1479 |     // print(duration)
 | 
|  |   1480 |     // print(" ")
 | 
|  |   1481 |     print(asize(bders_simp(s.toList, internalise(r))))
 | 
|  |   1482 |     //print(size(ders_simp(s.toList, r)))
 | 
|  |   1483 |     println()
 | 
|  |   1484 |   }
 | 
|  |   1485 |   // for(i <- 1 to 40000000 by 500000) {
 | 
|  |   1486 |   //   val s = "a" * i
 | 
|  |   1487 |   //   val t1 = System.nanoTime
 | 
|  |   1488 |   //   val derssimp_result = ders_simp(s.toList, r)
 | 
|  |   1489 |   //   val duration = (System.nanoTime - t1) / 1e9d
 | 
|  |   1490 |   //   print(i)
 | 
|  |   1491 |   //   print(" ")
 | 
|  |   1492 |   //   print(duration)
 | 
|  |   1493 |   //   println()
 | 
|  |   1494 |   // }
 | 
|  |   1495 |   
 | 
|  |   1496 | }
 | 
| 541 |   1497 | // naive_matcher()
 | 
| 543 |   1498 | //single(SEQ(SEQ(STAR(CHAR('b')),STAR(STAR(SEQ(CHAR('a'),CHAR('b'))))),
 | 
|  |   1499 | //  SEQ(SEQ(CHAR('b'),STAR(ALTS(CHAR('a'),ONE))),ONE)))
 | 
|  |   1500 | 
 | 
|  |   1501 | 
 | 
|  |   1502 | //STAR(SEQ(ALTS(CHAR(b),STAR(CHAR(b))),ALTS(ALTS(ONE,CHAR(b)),SEQ(CHAR(c),ONE))))
 | 
| 492 |   1503 | def generator_test() {
 | 
|  |   1504 | 
 | 
| 543 |   1505 |   test(single(STAR(SEQ(ALTS(STAR(CHAR('b')),ALTS(CHAR('b'),CHAR('a'))),ALTS(ALTS(CHAR('b'),CHAR('c')),STAR(ZERO))))), 1) { (r: Rexp) => 
 | 
| 500 |   1506 |   // ALTS(SEQ(SEQ(ONE,CHAR('a')),STAR(CHAR('a'))),SEQ(ALTS(CHAR('c'),ONE),STAR(ZERO))))))), 1) { (r: Rexp) => 
 | 
| 543 |   1507 |     val ss = Set("ab")//stringsFromRexp(r)
 | 
| 541 |   1508 |     val boolList = ss.map(s => {
 | 
| 518 |   1509 |       //val bdStrong = bdersStrong(s.toList, internalise(r))
 | 
| 541 |   1510 |       val bdStrong6 = bdersStrong6(s.toList, internalise(r))
 | 
| 543 |   1511 |       val bdStrong6Set = breakIntoTerms(erase(bdStrong6))
 | 
| 533 |   1512 |       val pdersSet = pderUNIV(r)//.flatMap(r => turnIntoTerms(r))
 | 
|  |   1513 |       val pdersSetBroken = pdersSet.flatMap(r => turnIntoTerms(r))
 | 
| 543 |   1514 |       (rs1_subseteq_rs2(bdStrong6Set.toSet, distinctByacc(pdersSet.toList)) ||
 | 
|  |   1515 |       bdStrong6Set.size <= pdersSet.size || bdStrong6Set.size <= pdersSetBroken.size ||
 | 
|  |   1516 |       rs1_subseteq_rs2(bdStrong6Set.toSet, pdersSet union pdersSetBroken))//|| bdStrong6Set.size <= pdersSetBroken.size
 | 
|  |   1517 |       
 | 
| 493 |   1518 |     })
 | 
| 500 |   1519 |     //!boolList.exists(b => b == false)
 | 
| 493 |   1520 |     !boolList.exists(b => b == false)
 | 
| 492 |   1521 |   }
 | 
|  |   1522 | }
 | 
| 518 |   1523 | // small()
 | 
| 543 |   1524 | // generator_test()
 | 
|  |   1525 | 
 | 
|  |   1526 | 
 | 
| 532 |   1527 |   
 | 
|  |   1528 |   //STAR(STAR(STAR(SEQ(SEQ(ALTS(STAR(ALTS(ONE,CHAR('c'))),
 | 
|  |   1529 |           //  CHAR('c')),ALTS(CHAR('a'),ALTS(STAR(CHAR('b')),ALTS(CHAR('a'),ZERO)))),
 | 
|  |   1530 |           //  CHAR('c')))))//SEQ(STAR(CHAR('c')),STAR(SEQ(STAR(CHAR('c')),ONE)))//STAR(SEQ(ALTS(STAR(CHAR('c')),CHAR('c')),SEQ(ALTS(CHAR('c'),ONE),ONE)))
 | 
|  |   1531 | //counterexample1: STAR(SEQ(ALTS(STAR(ZERO),ALTS(CHAR(a),CHAR(b))),SEQ(ONE,ALTS(CHAR(a),CHAR(b)))))
 | 
|  |   1532 | //counterexample2: SEQ(ALTS(SEQ(CHAR(a),STAR(ONE)),STAR(ONE)),ALTS(CHAR(a),SEQ(ALTS(CHAR(c),CHAR(a)),CHAR(b))))
 | 
| 541 |   1533 | 
 | 
|  |   1534 | //new ce1 : STAR(SEQ(ALTS(ALTS(ONE,CHAR(a)),SEQ(ONE,CHAR(b))),ALTS(CHAR(a),ALTS(CHAR(b),CHAR(a)))))
 | 
|  |   1535 | //new ce2 : ALTS(CHAR(b),SEQ(ALTS(ZERO,ALTS(CHAR(b),CHAR(b))),ALTS(ALTS(CHAR(a),CHAR(b)),SEQ(CHAR(c),ONE))))
 | 
|  |   1536 | //new ce3 : SEQ(CHAR(b),ALTS(ALTS(ALTS(ONE,CHAR(a)),SEQ(CHAR(c),ONE)),SEQ(STAR(ZERO),SEQ(ONE,CHAR(b)))))
 | 
| 543 |   1537 | //circular double-nullable STAR(SEQ((STAR("b") | "a" | "b"), ("b" | ONE)))
 | 
| 518 |   1538 | def counterexample_check() {
 | 
| 543 |   1539 |   val r = STAR(SEQ(ALTS(STAR(CHAR('b')),ALTS(CHAR('b'),CHAR('a'))),
 | 
|  |   1540 |     ALTS(ALTS(CHAR('b'),CHAR('c')),STAR(ZERO))))//SEQ(SEQ(STAR(ALTS(CHAR('a'),CHAR('b'))),
 | 
| 541 |   1541 |     //ALTS(ALTS(CHAR('c'),CHAR('b')),STAR(ONE))),STAR(CHAR('b')))
 | 
| 543 |   1542 |   val s = "ab"
 | 
|  |   1543 |   val bdStrong5 = bdersStrong6(s.toList, internalise(r))
 | 
|  |   1544 |   val bdStrong5Set = breakIntoTerms(erase(bdStrong5))
 | 
| 533 |   1545 |   val pdersSet = pderUNIV(r)//.map(r => erase(bsimp(internalise(r))))//.flatMap(r => turnIntoTerms(r))//.map(oneSimp).flatMap(r => breakIntoTerms(r))
 | 
| 541 |   1546 |   val apdersSet = pdersSet.map(internalise)
 | 
| 518 |   1547 |   println("original regex ")
 | 
|  |   1548 |   rprint(r)
 | 
|  |   1549 |   println("after strong bsimp")
 | 
|  |   1550 |   aprint(bdStrong5)
 | 
|  |   1551 |   println("turned into a set %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%   ")
 | 
|  |   1552 |   rsprint(bdStrong5Set)
 | 
|  |   1553 |   println("after pderUNIV")
 | 
|  |   1554 |   rsprint(pdersSet.toList)
 | 
| 541 |   1555 |   println("pderUNIV distinctBy6")
 | 
|  |   1556 |   //asprint(distinctBy6(apdersSet.toList))
 | 
| 543 |   1557 |   rsprint((pdersSet.toList).flatMap(turnIntoTerms))
 | 
| 541 |   1558 |   // rsprint(turnIntoTerms(pdersSet.toList(3)))
 | 
|  |   1559 |   // println("NO 3 not into terms")
 | 
|  |   1560 |   // rprint((pdersSet.toList()))
 | 
|  |   1561 |   // println("after pderUNIV broken")
 | 
|  |   1562 |   // rsprint(pdersSet.flatMap(r => turnIntoTerms(r)).toList)
 | 
| 532 |   1563 | 
 | 
| 518 |   1564 | }
 | 
| 543 |   1565 | 
 | 
|  |   1566 | // counterexample_check()
 | 
|  |   1567 | 
 | 
|  |   1568 | def lf_display() {
 | 
|  |   1569 |   val r = STAR(SEQ(ALTS(STAR(CHAR('b')),ALTS(CHAR('b'),CHAR('a'))),
 | 
|  |   1570 |     ALTS(ALTS(CHAR('b'),CHAR('c')),STAR(ZERO))))//SEQ(SEQ(STAR(ALTS(CHAR('a'),CHAR('b'))),
 | 
|  |   1571 |     //ALTS(ALTS(CHAR('c'),CHAR('b')),STAR(ONE))),STAR(CHAR('b')))
 | 
|  |   1572 |   val s = "ab"
 | 
|  |   1573 |   val bdStrong5 = bdersStrong6(s.toList, internalise(r))
 | 
|  |   1574 |   val bdStrong5Set = breakIntoTerms(erase(bdStrong5))
 | 
|  |   1575 |   val pdersSet = pderUNIV(r)//.map(r => erase(bsimp(internalise(r))))//.flatMap(r => turnIntoTerms(r))//.map(oneSimp).flatMap(r => breakIntoTerms(r))
 | 
|  |   1576 |   val apdersSet = pdersSet.map(internalise)
 | 
|  |   1577 |   lfprint(lf(r))
 | 
|  |   1578 | 
 | 
|  |   1579 | }
 | 
|  |   1580 | 
 | 
|  |   1581 | lf_display()
 | 
| 541 |   1582 | 
 | 
|  |   1583 | def breakable(r: Rexp) : Boolean = r match {
 | 
|  |   1584 |   case SEQ(ALTS(_, _), _) => true
 | 
|  |   1585 |   case SEQ(r1, r2) => breakable(r1)
 | 
|  |   1586 |   case _ => false
 | 
|  |   1587 | }
 | 
| 532 |   1588 | 
 | 
| 526 |   1589 | def linform_test() {
 | 
| 530 |   1590 |   val r = STAR(SEQ(STAR(CHAR('c')),ONE))//SEQ(STAR(CHAR('c')),STAR(SEQ(STAR(CHAR('c')),ONE))) //
 | 
| 526 |   1591 |   val r_linforms = lf(r)
 | 
|  |   1592 |   println(r_linforms.size)
 | 
| 530 |   1593 |   val pderuniv = pderUNIV(r)
 | 
|  |   1594 |   println(pderuniv.size)
 | 
| 526 |   1595 | }
 | 
| 530 |   1596 | // linform_test()
 | 
| 516 |   1597 | // 1
 | 
| 518 |   1598 | def newStrong_test() {
 | 
|  |   1599 |   val r2 = (CHAR('b') | ONE)
 | 
|  |   1600 |   val r0 = CHAR('d')
 | 
|  |   1601 |   val r1 = (ONE | CHAR('c'))
 | 
|  |   1602 |   val expRexp = (SEQ(r2, r0) | SEQ(SEQ(r1, r2), r0))
 | 
|  |   1603 |   println(s"original regex is: ")
 | 
|  |   1604 |   rprint(expRexp)
 | 
|  |   1605 |   val expSimp5 = strongBsimp5(internalise(expRexp))
 | 
|  |   1606 |   val expSimp6 = strongBsimp6(internalise(expRexp))
 | 
|  |   1607 |   aprint(expSimp5)
 | 
|  |   1608 |   aprint(expSimp6)
 | 
|  |   1609 | }
 | 
|  |   1610 | // newStrong_test()
 | 
| 516 |   1611 | // SEQ(SEQ(SEQ(ONE,CHAR('b')),STAR(CHAR('b'))),
 | 
|  |   1612 | // SEQ(ALTS(ALTS(ZERO,STAR(CHAR('b'))),
 | 
|  |   1613 | // STAR(ALTS(CHAR('a'),SEQ(SEQ(STAR(ALTS(STAR(CHAR('c')),CHAR('a'))),
 | 
|  |   1614 | // SEQ(CHAR('a'),SEQ(ALTS(CHAR('b'),ZERO),SEQ(ONE,CHAR('b'))))),ONE)))),ONE))
 | 
| 493 |   1615 | 
 | 
|  |   1616 | 
 | 
|  |   1617 | // 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))
 | 
| 543 |   1618 | // 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))
 | 
|  |   1619 | 
 | 
|  |   1620 | 
 | 
|  |   1621 | 
 | 
|  |   1622 | 
 | 
|  |   1623 | 
 | 
|  |   1624 | def bders_bderssimp() {
 | 
|  |   1625 | 
 | 
|  |   1626 |   test(single(SEQ(ALTS(ONE,CHAR('c')),
 | 
|  |   1627 |   SEQ(SEQ(CHAR('a'),CHAR('a')),ALTS(ONE,CHAR('c'))))), 1) { (r: Rexp) => 
 | 
|  |   1628 |   // ALTS(SEQ(SEQ(ONE,CHAR('a')),STAR(CHAR('a'))),SEQ(ALTS(CHAR('c'),ONE),STAR(ZERO))))))), 1) { (r: Rexp) => 
 | 
|  |   1629 |     val ss = Set("aa")//stringsFromRexp(r)
 | 
|  |   1630 |     val boolList = ss.map(s => {
 | 
|  |   1631 |       //val bdStrong = bdersStrong(s.toList, internalise(r))
 | 
|  |   1632 |       val bds = bsimp(bders(s.toList, internalise(r)))
 | 
|  |   1633 |       val bdssimp = bsimp(bders_simp(s.toList, internalise(r)))
 | 
|  |   1634 |       //val pdersSet = pderUNIV(r)//.flatMap(r => turnIntoTerms(r))
 | 
|  |   1635 |       //val pdersSetBroken = pdersSet.flatMap(r => turnIntoTerms(r))
 | 
|  |   1636 |       //rs1_subseteq_rs2(bdStrong6Set.toSet, distinctByacc(pdersSet.toList))
 | 
|  |   1637 |       //bdStrong6Set.size <= pdersSet.size || bdStrong6Set.size <= pdersSetBroken.size ||
 | 
|  |   1638 |       //rs1_subseteq_rs2(bdStrong6Set.toSet, pdersSet union pdersSetBroken)//|| bdStrong6Set.size <= pdersSetBroken.size
 | 
|  |   1639 |       rprint(r)
 | 
|  |   1640 |       println(bds)
 | 
|  |   1641 |       println(bdssimp)
 | 
|  |   1642 |       bds == bdssimp
 | 
|  |   1643 |     })
 | 
|  |   1644 |     //!boolList.exists(b => b == false)
 | 
|  |   1645 |     !boolList.exists(b => b == false)
 | 
|  |   1646 |   }
 | 
|  |   1647 | }
 | 
|  |   1648 | //  bders_bderssimp()
 | 
|  |   1649 |   
 | 
|  |   1650 | 
 | 
|  |   1651 | 
 |