// A version with simplification of derivatives;// this keeps the regular expressions small, which// is good for the run-time// // call the test cases with X = {1,2}//// amm re3.sc testX//// or //// amm re3.sc allabstract class Rexpcase object ZERO extends Rexpcase object ONE extends Rexpcase class CHAR(c: Char) extends Rexpcase class ALT(r1: Rexp, r2: Rexp) extends Rexp case class SEQ(r1: Rexp, r2: Rexp) extends Rexp case class STAR(r: Rexp) extends Rexp case class NTIMES(r: Rexp, n: Int) extends Rexp // the nullable function: tests whether the regular // expression can recognise the empty stringdef nullable (r: Rexp) : Boolean = r match { case ZERO => false case ONE => true case CHAR(_) => false case ALT(r1, r2) => nullable(r1) || nullable(r2) case SEQ(r1, r2) => nullable(r1) && nullable(r2) case STAR(_) => true case NTIMES(r, i) => if (i == 0) true else nullable(r)}// the derivative of a regular expression w.r.t. a characterdef der(c: Char, r: Rexp) : Rexp = r match { case ZERO => ZERO case ONE => ZERO case CHAR(d) => if (c == d) ONE else ZERO case ALT(r1, r2) => ALT(der(c, r1), der(c, r2)) case SEQ(r1, r2) => if (nullable(r1)) ALT(SEQ(der(c, r1), r2), der(c, r2)) else SEQ(der(c, r1), r2) case STAR(r1) => SEQ(der(c, r1), STAR(r1)) case NTIMES(r, i) => if (i == 0) ZERO else SEQ(der(c, r), NTIMES(r, i - 1))}def simp(r: Rexp) : Rexp = r match { case ALT(r1, r2) => (simp(r1), simp(r2)) match { case (ZERO, r2s) => r2s case (r1s, ZERO) => r1s case (r1s, r2s) => if (r1s == r2s) r1s else ALT (r1s, r2s) } case SEQ(r1, r2) => (simp(r1), simp(r2)) match { case (ZERO, _) => ZERO case (_, ZERO) => ZERO case (ONE, r2s) => r2s case (r1s, ONE) => r1s case (r1s, r2s) => SEQ(r1s, r2s) } case r => r}// the derivative w.r.t. a string (iterates der)def ders(s: List[Char], r: Rexp) : Rexp = s match { case Nil => r case c::s => ders(s, simp(der(c, r)))}// the main matcher functiondef matcher(r: Rexp, s: String) : Boolean = nullable(ders(s.toList, r))// one or zerodef OPT(r: Rexp) = ALT(r, ONE)// Test Cases// evil regular expressions: (a?){n} a{n} and (a*)* bdef EVIL1(n: Int) = SEQ(NTIMES(OPT(CHAR('a')), n), NTIMES(CHAR('a'), n))val EVIL2 = SEQ(STAR(STAR(CHAR('a'))), CHAR('b'))def time_needed[T](i: Int, code: => T) = { val start = System.nanoTime() for (j <- 1 to i) code val end = System.nanoTime() (end - start)/(i * 1.0e9)}@arg(doc = "Test (a?{n}) (a{n})")@maindef test1() = { println("Test (a?{n}) (a{n})") for (i <- 0 to 9000 by 1000) { println(f"$i: ${time_needed(3, matcher(EVIL1(i), "a" * i))}%.5f") }} @arg(doc = "Test (a*)* b")@maindef test2() = { println("Test (a*)* b") for (i <- 0 to 6000000 by 500000) { println(f"$i: ${time_needed(3, matcher(EVIL2, "a" * i))}%.5f") }}// size of a regular expressions - for testing purposes def size(r: Rexp) : Int = r match { case ZERO => 1 case ONE => 1 case CHAR(_) => 1 case ALT(r1, r2) => 1 + size(r1) + size(r2) case SEQ(r1, r2) => 1 + size(r1) + size(r2) case STAR(r) => 1 + size(r) case NTIMES(r, _) => 1 + size(r)}// now the size of the derivatives grows // much, much slowersize(ders("".toList, EVIL2)) // 5size(ders("a".toList, EVIL2)) // 8size(ders("aa".toList, EVIL2)) // 8size(ders("aaa".toList, EVIL2)) // 8size(ders("aaaa".toList, EVIL2)) // 8size(ders("aaaaa".toList, EVIL2)) // 8@arg(doc = "All tests.")@maindef all() = { test1(); test2() } // PS://// If you want to dig deeper into the topic, you can have // a look at these examples which still explode. They// need an even more aggressive simplification.// test: (a + aa)*val EVIL3 = STAR(ALT(CHAR('a'), SEQ(CHAR('a'), CHAR('a'))))@arg(doc = "Test EVIL3")@maindef test3() = { println("Test (a + aa)*") for (i <- 0 to 30 by 5) { println(f"$i: ${time_needed(1, matcher(EVIL3, "a" * i))}%.5f") }}// test: (1 + a + aa)*val EVIL4 = STAR(ALT(ONE, ALT(CHAR('a'), SEQ(CHAR('a'), CHAR('a')))))@arg(doc = "Test EVIL4")@maindef test4() = { println("Test (1 + a + aa)*") for (i <- 0 to 30 by 5) { println(f"$i: ${time_needed(1, matcher(EVIL4, "a" * i))}%.5f") }}@arg(doc = "Tests that show not all is hunky-dory, but a solution leads too far afield.")@maindef fail() = { test3(); test4() } // runs with amm2 and amm3def pp(r: Rexp): String = r match { case SEQ(CHAR(a1), SEQ(r1, r2)) => s"${a1}${pp(r1)}${pp(r2)}" case SEQ(ONE, SEQ(r1, r2)) => s"1${pp(r1)}${pp(r2)}" case SEQ(ZERO, SEQ(r1, r2)) => s"0${pp(r1)}${pp(r2)}" case SEQ(CHAR(a1), CHAR(a2)) => s"${a1}${a2}" case SEQ(ONE, CHAR(a2)) => s"1${a2}" case SEQ(ZERO, CHAR(a2)) => s"0${a2}" case ZERO => "0" case ONE => "1" case CHAR(a) => a.toString case ALT(r1, r2) => s"(${pp(r1)} + ${pp(r2)})" case SEQ(r1, r2) => s"(${pp(r1)} o ${pp(r2)})" case STAR(r1) => s"(${pp(r1)})*"}val REG = STAR(ALT(CHAR('a'), SEQ(CHAR('a'), CHAR('a'))))print(pp(ders("".toList, REG)))print(pp(ders("a".toList, REG)))print(pp(ders("aa".toList, REG)))print(pp(ders("aaa".toList, REG)))size(ders("".toList, REG)) // 6size(ders("a".toList, REG)) // 12size(ders("aa".toList, REG)) // 27size(ders("aaa".toList, REG)) // 55size(ders("aaaa".toList, REG)) // 8size(ders("aaaaa".toList, REG)) // 169size(ders("aaaaaa".toList, REG)) // 283size(ders(("a" * 7).toList, REG)) // 468size(ders(("a" * 8).toList, REG)) // 767size(ders(("a" * 9).toList, REG)) // 1251size(ders(("a" * 10).toList, REG))// 2034size(ders(("a" * 11).toList, REG))// 3301for (i <- (0 to 40)) { println(s"$i:" + size(ders(("a" * i).toList, REG)))}