re1.scala
author Christian Urban <urbanc@in.tum.de>
Wed, 31 Oct 2012 21:46:27 +0000
changeset 54 485f38b530ab
parent 49 d2c6852ca8da
child 59 b64e876832cc
permissions -rw-r--r--
updated

    
abstract class Rexp

case object NULL extends Rexp
case object EMPTY extends Rexp
case class CHAR(c: Char) extends Rexp
case class ALT(r1: Rexp, r2: Rexp) extends Rexp 
case class SEQ(r1: Rexp, r2: Rexp) extends Rexp 
case class STAR(r: Rexp) extends Rexp 

// some convenience for typing in regular expressions
def charlist2rexp(s : List[Char]) : Rexp = s match {
  case Nil => EMPTY
  case c::Nil => CHAR(c)
  case c::s => SEQ(CHAR(c), charlist2rexp(s))
}
implicit def string2rexp(s : String) : Rexp = charlist2rexp(s.toList)


// nullable function: tests whether the regular 
// expression can recognise the empty string
def nullable (r: Rexp) : Boolean = r match {
  case NULL => false
  case EMPTY => true
  case CHAR(_) => false
  case ALT(r1, r2) => nullable(r1) || nullable(r2)
  case SEQ(r1, r2) => nullable(r1) && nullable(r2)
  case STAR(_) => true
}

// derivative of a regular expression w.r.t. a character
def der (c: Char, r: Rexp) : Rexp = r match {
  case NULL => NULL
  case EMPTY => NULL
  case CHAR(d) => if (c == d) EMPTY else NULL
  case ALT(r1, r2) => ALT(der(c, r1), der(c, r2))
  case SEQ(r1, r2) => 
    if (nullable(r1)) ALT(SEQ(der(c, r1), r2), der(c, r2))
    else SEQ(der(c, r1), r2)
  case STAR(r) => SEQ(der(c, r), STAR(r))
}

// derivative w.r.t. a string (iterates der)
def ders (s: List[Char], r: Rexp) : Rexp = s match {
  case Nil => r
  case c::s => ders(s, der(c, r))
}

// main matcher function
def matcher(r: Rexp, s: String) : Boolean = nullable(ders(s.toList, r))


//one or zero
def OPT(r: Rexp) = ALT(r, EMPTY)

//n-times
def NTIMES(r: Rexp, n: Int) : Rexp = n match {
  case 0 => NULL
  case 1 => r
  case n => SEQ(r, NTIMES(r, n - 1))
}

def RTEST(n: Int) = SEQ(NTIMES(OPT("a"), n), NTIMES("a", n))

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)
}

for (i <- 1 to 29) {
  println(i + ": " + "%.5f".format(time_needed(1, matcher(RTEST(i), "a" * i))))
}