RegexExplosionPlot/evilForBsimp.sc
changeset 666 6da4516ea87d
parent 545 333013923c5a
--- a/RegexExplosionPlot/evilForBsimp.sc	Fri Jul 14 00:32:41 2023 +0100
+++ b/RegexExplosionPlot/evilForBsimp.sc	Mon Jul 24 11:09:48 2023 +0100
@@ -360,6 +360,18 @@
   case (ANYCHAR, Empty) => Chr(c)
 }
 
+
+def print_lexer(r: Rexp, s: List[Char]): Val = s match{
+  case Nil => {println("mkeps!"); println(r); val v = mkeps(r); println(v); v}
+  case c::cs => { 
+    val derc = der(c, r);
+    println("derivative is: "+ derc);
+    val v = print_lexer(derc, cs); println("Injection!");println("r: "+r); println("c: "+c); println("v:"+v); inj(r, c, v ) }
+}
+val r = STAR("a"|"aa") ~ ("c")
+val v = print_lexer(r, "aac".toList)
+println(v)
+
 // some "rectification" functions for simplification
 
 
@@ -734,14 +746,24 @@
 }
 
 
+def blexing(r: Rexp, s: String) : Option[Val] = {
+    val bit_code = blex(internalise(r), s.toList)
+    bit_code match {
+      case Some(bits) => Some(decode(r, bits))
+      case None => None
+    }
+}
 
-def blexing_simp(r: Rexp, s: String) : Val = {
+def blexing_simp(r: Rexp, s: String) : Option[Val] = {
     val bit_code = blex_simp(internalise(r), s.toList)
-    decode(r, bit_code)
+    bit_code match {
+      case Some(bits) => Some(decode(r, bits))
+      case None => None
+    }
 }
-def simpBlexer(r: Rexp, s: String) : Val = {
-  Try(blexing_simp(r, s)).getOrElse(Failure)
-}
+//def simpBlexer(r: Rexp, s: String) : Val = {
+//  Try(blexing_simp(r, s)).getOrElse(Failure)
+//}
 
 def strong_blexing_simp(r: Rexp, s: String) : Val = {
   decode(r, strong_blex_simp(internalise(r), s.toList))
@@ -854,15 +876,32 @@
       //case RTOP => "RTOP"
     }
 
-
-  def blex_simp(r: ARexp, s: List[Char]) : Bits = s match {
+  def blex(r: ARexp, s: List[Char]) : Option[Bits]= s match {
       case Nil => {
         if (bnullable(r)) {
           val bits = mkepsBC(r)
-          bits
+          Some(bits)
         }
         else 
-          throw new Exception("Not matched")
+          None
+          //throw new Exception("Not matched")
+      }
+      case c::cs => {
+        val der_res = bder(c,r)
+        //val simp_res = bsimp(der_res)  
+        blex(der_res, cs)      
+      }
+  }
+
+  def blex_simp(r: ARexp, s: List[Char]) : Option[Bits]= s match {
+      case Nil => {
+        if (bnullable(r)) {
+          val bits = mkepsBC(r)
+          Some(bits)
+        }
+        else 
+          None
+          //throw new Exception("Not matched")
       }
       case c::cs => {
         val der_res = bder(c,r)
@@ -1125,7 +1164,7 @@
 
 
 }
-small()
+//small()
 // generator_test()
 
 // 1
@@ -1137,4 +1176,19 @@
 
 
 // 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))
-// 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))
\ No newline at end of file
+// 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))
+
+
+//def time[R](block: => R): R = {
+//    val t0 = System.nanoTime()
+//    val result = block    // call-by-name
+//    val t1 = System.nanoTime()
+//    println(" " + (t1 - t0)/1e9 + "")
+//    result
+//}
+//val astarstarB  = STAR(STAR("a"))~("b")
+//val data_points_x_axis = List.range(0, 60000, 3000)
+//for (i <- data_points_x_axis) {print(i); time { blexing(astarstarB, ("a"*i)) } }
+////println(data_points_x_axis.zip( List(0,1,3,4)))
+//for(i <- List.range(0, 40000, 3000)) print(i + " ")
+