updated
authorChristian Urban <urbanc@in.tum.de>
Wed, 13 Mar 2019 10:36:29 +0000
changeset 314 20a57552d722
parent 313 3b8e3a156200
child 315 ab7fe342e004
updated
exps/bit.scala
exps/both.scala
thys/BitCoded.thy
thys/RegLangs.thy
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/exps/bit.scala	Wed Mar 13 10:36:29 2019 +0000
@@ -0,0 +1,829 @@
+
+import scala.language.implicitConversions    
+import scala.language.reflectiveCalls
+import scala.annotation.tailrec   
+import scala.util.Try
+
+// for escaping strings
+def escape(raw: String) : String = {
+  import scala.reflect.runtime.universe._
+  Literal(Constant(raw)).toString
+}
+
+def esc2(r: (String, String)) = (escape(r._1), escape(r._2))
+
+def distinctBy[B, C](xs: List[B], f: B => C, acc: List[C] = Nil): List[B] = xs match {
+  case Nil => Nil
+  case (x::xs) => {
+    val res = f(x)
+    if (acc.contains(res)) distinctBy(xs, f, acc)  
+    else x::distinctBy(xs, f, res::acc)
+  }
+} 
+
+abstract class Bit
+case object Z extends Bit
+case object S extends Bit
+case class C(c: Char) extends Bit
+
+type Bits = List[Bit]
+
+// usual regular expressions
+abstract class Rexp 
+case object ZERO extends Rexp
+case object ONE extends Rexp
+case class PRED(f: Char => Boolean, s: String = "_") extends Rexp 
+case class ALTS(rs: List[Rexp]) extends Rexp 
+case class SEQ(r1: Rexp, r2: Rexp) extends Rexp 
+case class STAR(r: Rexp) extends Rexp 
+case class RECD(x: String, r: Rexp) extends Rexp
+
+
+// abbreviations
+def CHAR(c: Char) = PRED(_ == c, c.toString)
+def ALT(r1: Rexp, r2: Rexp) = ALTS(List(r1, r2))
+def PLUS(r: Rexp) = SEQ(r, STAR(r))
+val ANYCHAR = PRED(_ => true, ".")
+
+// annotated regular expressions
+abstract class ARexp 
+case object AZERO extends ARexp
+case class AONE(bs: Bits) extends ARexp
+case class APRED(bs: Bits, f: Char => Boolean, s: String = "_") extends ARexp
+case class AALTS(bs: Bits, rs: List[ARexp]) extends ARexp 
+case class ASEQ(bs: Bits, r1: ARexp, r2: ARexp) extends ARexp 
+case class ASTAR(bs: Bits, r: ARexp) extends ARexp 
+
+// abbreviations
+def AALT(bs: Bits, r1: ARexp, r2: ARexp) = AALTS(bs, List(r1, r2))
+
+// values
+abstract class Val
+case object Empty extends Val
+case class Chr(c: Char) extends Val
+case class Sequ(v1: Val, v2: Val) extends Val
+case class Left(v: Val) extends Val
+case class Right(v: Val) extends Val
+case class Stars(vs: List[Val]) extends Val
+case class Rec(x: String, v: Val) extends Val
+
+
+   
+// some convenience for typing in regular expressions
+def charlist2rexp(s : List[Char]): Rexp = s match {
+  case Nil => ONE
+  case c::Nil => CHAR(c)
+  case c::s => SEQ(CHAR(c), charlist2rexp(s))
+}
+implicit def string2rexp(s : String) : Rexp = charlist2rexp(s.toList)
+
+implicit def RexpOps(r: Rexp) = new {
+  def | (s: Rexp) = ALT(r, s)
+  def % = STAR(r)
+  def ~ (s: Rexp) = SEQ(r, s)
+}
+
+implicit def stringOps(s: String) = new {
+  def | (r: Rexp) = ALT(s, r)
+  def | (r: String) = ALT(s, r)
+  def % = STAR(s)
+  def ~ (r: Rexp) = SEQ(s, r)
+  def ~ (r: String) = SEQ(s, r)
+  def $ (r: Rexp) = RECD(s, r)
+}
+
+
+// string of a regular expressions - for testing purposes
+def string(r: Rexp): String = r match {
+  case ZERO => "0"
+  case ONE => "1"
+  case PRED(_, s) => s
+  case ALTS(rs) => rs.map(string).mkString("[", "|", "]")
+  case SEQ(r1, r2) => s"(${string(r1)} ~ ${string(r2)})"
+  case STAR(r) => s"{${string(r)}}*"
+  case RECD(x, r) => s"(${x}! ${string(r)})"
+}
+
+// string of an annotated regular expressions - for testing purposes
+
+def astring(a: ARexp): String = a match {
+  case AZERO => "0"
+  case AONE(_) => "1"
+  case APRED(_, _, s) => s
+  case AALTS(_, rs) => rs.map(astring).mkString("[", "|", "]")
+  case ASEQ(_, r1, r2) => s"(${astring(r1)} ~ ${astring(r2)})"
+  case ASTAR(_, r) => s"{${astring(r)}}*"
+}
+ 
+
+//--------------------------------------------------------------
+// START OF NON-BITCODE PART
+//
+
+// nullable function: tests whether the regular 
+// expression can recognise the empty string
+def nullable(r: Rexp) : Boolean = r match {
+  case ZERO => false
+  case ONE => true
+  case PRED(_, _) => false
+  case ALTS(rs) => rs.exists(nullable)
+  case SEQ(r1, r2) => nullable(r1) && nullable(r2)
+  case STAR(_) => true
+  case RECD(_, r) => nullable(r)
+}
+
+// derivative of a regular expression w.r.t. a character
+def der(c: Char, r: Rexp) : Rexp = r match {
+  case ZERO => ZERO
+  case ONE => ZERO
+  case PRED(f, _) => if (f(c)) ONE else ZERO
+  case ALTS(List(r1, r2)) => ALTS(List(der(c, r1), der(c, r2)))
+  case SEQ(r1, r2) => 
+    if (nullable(r1)) ALTS(List(SEQ(der(c, r1), r2), der(c, r2)))
+    else SEQ(der(c, r1), r2)
+  case STAR(r) => SEQ(der(c, r), STAR(r))
+  case RECD(_, r1) => der(c, r1)
+}
+
+
+def flatten(v: Val) : String = v match {
+  case Empty => ""
+  case Chr(c) => c.toString
+  case Left(v) => flatten(v)
+  case Right(v) => flatten(v)
+  case Sequ(v1, v2) => flatten(v1) + flatten(v2)
+  case Stars(vs) => vs.map(flatten).mkString
+  case Rec(_, v) => flatten(v)
+}
+
+// extracts an environment from a value
+def env(v: Val) : List[(String, String)] = v match {
+  case Empty => Nil
+  case Chr(c) => Nil
+  case Left(v) => env(v)
+  case Right(v) => env(v)
+  case Sequ(v1, v2) => env(v1) ::: env(v2)
+  case Stars(vs) => vs.flatMap(env)
+  case Rec(x, v) => (x, flatten(v))::env(v)
+}
+
+
+// injection part
+def mkeps(r: Rexp) : Val = r match {
+  case ONE => Empty
+  case ALTS(List(r1, r2)) => 
+    if (nullable(r1)) Left(mkeps(r1)) else Right(mkeps(r2))
+  case SEQ(r1, r2) => Sequ(mkeps(r1), mkeps(r2))
+  case STAR(r) => Stars(Nil)
+  case RECD(x, r) => Rec(x, mkeps(r))
+}
+
+def inj(r: Rexp, c: Char, v: Val) : Val = (r, v) match {
+  case (STAR(r), Sequ(v1, Stars(vs))) => Stars(inj(r, c, v1)::vs)
+  case (SEQ(r1, r2), Sequ(v1, v2)) => Sequ(inj(r1, c, v1), v2)
+  case (SEQ(r1, r2), Left(Sequ(v1, v2))) => Sequ(inj(r1, c, v1), v2)
+  case (SEQ(r1, r2), Right(v2)) => Sequ(mkeps(r1), inj(r2, c, v2))
+  case (ALTS(List(r1, r2)), Left(v1)) => Left(inj(r1, c, v1))
+  case (ALTS(List(r1, r2)), Right(v2)) => Right(inj(r2, c, v2))
+  case (PRED(_, _), Empty) => Chr(c) 
+  case (RECD(x, r1), _) => Rec(x, inj(r1, c, v))
+}
+
+// lexing without simplification
+def lex(r: Rexp, s: List[Char]) : Val = s match {
+  case Nil => if (nullable(r)) mkeps(r) else throw new Exception("Not matched")
+  case c::cs => inj(r, c, lex(der(c, r), cs))
+}
+
+def lexing(r: Rexp, s: String) : Val = lex(r, s.toList)
+
+//println(lexing(("ab" | "ab") ~ ("b" | ONE), "ab"))
+
+// some "rectification" functions for simplification
+def F_ID(v: Val): Val = v
+def F_RIGHT(f: Val => Val) = (v:Val) => Right(f(v))
+def F_LEFT(f: Val => Val) = (v:Val) => Left(f(v))
+def F_ALT(f1: Val => Val, f2: Val => Val) = (v:Val) => v match {
+  case Right(v) => Right(f2(v))
+  case Left(v) => Left(f1(v))
+}
+def F_SEQ(f1: Val => Val, f2: Val => Val) = (v:Val) => v match {
+  case Sequ(v1, v2) => Sequ(f1(v1), f2(v2))
+}
+def F_SEQ_Empty1(f1: Val => Val, f2: Val => Val) = 
+  (v:Val) => Sequ(f1(Empty), f2(v))
+def F_SEQ_Empty2(f1: Val => Val, f2: Val => Val) = 
+  (v:Val) => Sequ(f1(v), f2(Empty))
+def F_RECD(f: Val => Val) = (v:Val) => v match {
+  case Rec(x, v) => Rec(x, f(v))
+}
+def F_ERROR(v: Val): Val = throw new Exception("error")
+
+// simplification of regular expressions returning also an
+// rectification function; no simplification under STAR 
+def simp(r: Rexp): (Rexp, Val => Val) = r match {
+  case ALTS(List(r1, r2)) => {
+    val (r1s, f1s) = simp(r1)
+    val (r2s, f2s) = simp(r2)
+    (r1s, r2s) match {
+      case (ZERO, _) => (r2s, F_RIGHT(f2s))
+      case (_, ZERO) => (r1s, F_LEFT(f1s))
+      case _ => if (r1s == r2s) (r1s, F_LEFT(f1s))
+                else (ALTS(List(r1s, r2s)), F_ALT(f1s, f2s)) 
+    }
+  }
+  case SEQ(r1, r2) => {
+    val (r1s, f1s) = simp(r1)
+    val (r2s, f2s) = simp(r2)
+    (r1s, r2s) match {
+      case (ZERO, _) => (ZERO, F_ERROR)
+      //case (_, ZERO) => (ZERO, F_ERROR)
+      case (ONE, _) => (r2s, F_SEQ_Empty1(f1s, f2s))
+      //case (_, ONE) => (r1s, F_SEQ_Empty2(f1s, f2s))
+      case _ => (SEQ(r1s,r2s), F_SEQ(f1s, f2s))
+    }
+  }
+  case RECD(x, r1) => {
+    val (r1s, f1s) = simp(r1)
+    (RECD(x, r1s), F_RECD(f1s))
+  }
+  case r => (r, F_ID)
+}
+
+def ders_simp(s: List[Char], r: Rexp) : Rexp = s match {
+  case Nil => r
+  case c::s => ders_simp(s, simp(der(c, r))._1)
+}
+
+
+def lex_simp(r: Rexp, s: List[Char]) : Val = s match {
+  case Nil => if (nullable(r)) mkeps(r) else throw new Exception("Not matched")
+  case c::cs => {
+    val (r_simp, f_simp) = simp(der(c, r))
+    inj(r, c, f_simp(lex_simp(r_simp, cs)))
+  }
+}
+
+def lexing_simp(r: Rexp, s: String) : Val = lex_simp(r, s.toList)
+
+//println(lexing_simp(("a" | "ab") ~ ("b" | ""), "ab"))
+
+
+def tokenise_simp(r: Rexp, s: String) = 
+  env(lexing_simp(r, s)).map(esc2)
+
+//--------------------------------------------------------------------
+// Partial Derivatives
+
+
+def pder(c: Char, r: Rexp): Set[Rexp] = r match {
+  case ZERO => Set()
+  case ONE => Set()
+  case PRED(f, _) => if (f(c)) Set(ONE) else Set()
+  case ALTS(rs) => rs.toSet.flatMap(pder(c, _))
+  case SEQ(r1, r2) =>
+    (for (pr1 <- pder(c, r1)) yield SEQ(pr1, r2)) ++
+    (if (nullable(r1)) pder(c, r2) else Set())
+  case STAR(r1) =>
+    for (pr1 <- pder(c, r1)) yield SEQ(pr1, STAR(r1))
+  case RECD(_, r1) => pder(c, r1)
+}
+
+def pders(cs: List[Char], r: Rexp): Set[Rexp] = cs match {
+  case Nil => Set(r)
+  case c::cs => pder(c, r).flatMap(pders(cs, _))
+}
+
+def pders_simp(cs: List[Char], r: Rexp): Set[Rexp] = cs match {
+  case Nil => Set(r)
+  case c::cs => pder(c, r).flatMap(pders_simp(cs, _)).map(simp(_)._1)
+}
+
+def psize(rs: Set[Rexp])  = 
+  rs.map(size).sum
+
+//--------------------------------------------------------------------
+// BITCODED PART
+
+
+def fuse(bs: Bits, r: ARexp) : ARexp = r match {
+  case AZERO => AZERO
+  case AONE(cs) => AONE(bs ++ cs)
+  case APRED(cs, f, s) => APRED(bs ++ cs, f, s)
+  case AALTS(cs, rs) => AALTS(bs ++ cs, rs)
+  case ASEQ(cs, r1, r2) => ASEQ(bs ++ cs, r1, r2)
+  case ASTAR(cs, r) => ASTAR(bs ++ cs, r)
+}
+
+// translation into ARexps
+def internalise(r: Rexp) : ARexp = r match {
+  case ZERO => AZERO
+  case ONE => AONE(Nil)
+  case PRED(f, s) => APRED(Nil, f, s)
+  case ALTS(List(r1, r2)) => 
+    AALTS(Nil, List(fuse(List(Z), internalise(r1)), fuse(List(S), internalise(r2))))
+  case ALTS(r1::rs) => {
+     val AALTS(Nil, rs2) = internalise(ALTS(rs))
+     AALTS(Nil, fuse(List(Z), internalise(r1)) :: rs2.map(fuse(List(S), _)))
+  }
+  case SEQ(r1, r2) => ASEQ(Nil, internalise(r1), internalise(r2))
+  case STAR(r) => ASTAR(Nil, internalise(r))
+  case RECD(x, r) => internalise(r)
+}
+
+internalise(("a" | "ab") ~ ("b" | ""))
+
+// decoding of values from bit sequences
+def decode_aux(r: Rexp, bs: Bits) : (Val, Bits) = (r, bs) match {
+  case (ONE, bs) => (Empty, bs)
+  case (PRED(f, _), C(c)::bs) => (Chr(c), bs)
+  case (ALTS(r::Nil), bs) => decode_aux(r, bs)
+  case (ALTS(rs), bs) => bs match {
+    case Z::bs1 => {
+      val (v, bs2) = decode_aux(rs.head, bs1)
+      (Left(v), bs2)
+    }
+    case S::bs1 => {
+      val (v, bs2) = decode_aux(ALTS(rs.tail), bs1)
+      (Right(v), bs2)			
+    }
+  }
+  case (SEQ(r1, r2), bs) => {
+    val (v1, bs1) = decode_aux(r1, bs)
+    val (v2, bs2) = decode_aux(r2, bs1)
+    (Sequ(v1, v2), bs2)
+  }
+  case (STAR(r1), S::bs) => {
+    val (v, bs1) = decode_aux(r1, bs)
+    val (Stars(vs), bs2) = decode_aux(STAR(r1), bs1)
+    (Stars(v::vs), bs2)
+  }
+  case (STAR(_), Z::bs) => (Stars(Nil), bs)
+  case (RECD(x, r1), bs) => {
+    val (v, bs1) = decode_aux(r1, bs)
+    (Rec(x, v), bs1)
+  }
+}
+
+def decode(r: Rexp, bs: Bits) = decode_aux(r, bs) match {
+  case (v, Nil) => v
+  case _ => throw new Exception("Not decodable")
+}
+
+
+//erase function: extracts a Rexp from Arexp
+def erase(r: ARexp) : Rexp = r match{
+  case AZERO => ZERO
+  case AONE(_) => ONE
+  case APRED(bs, f, s) => PRED(f, s)
+  case AALTS(bs, rs) => ALTS(rs.map(erase(_)))
+  case ASEQ(bs, r1, r2) => SEQ (erase(r1), erase(r2))
+  case ASTAR(cs, r)=> STAR(erase(r))
+}
+
+
+// bnullable function: tests whether the aregular 
+// expression can recognise the empty string
+def bnullable (r: ARexp) : Boolean = r match {
+  case AZERO => false
+  case AONE(_) => true
+  case APRED(_,_,_) => false
+  case AALTS(_, rs) => rs.exists(bnullable)
+  case ASEQ(_, r1, r2) => bnullable(r1) && bnullable(r2)
+  case ASTAR(_, _) => true
+}
+
+def bmkeps(r: ARexp) : Bits = r match {
+  case AONE(bs) => bs
+  case AALTS(bs, rs) => {
+    val n = rs.indexWhere(bnullable)
+    bs ++ bmkeps(rs(n))
+  }
+  case ASEQ(bs, r1, r2) => bs ++ bmkeps(r1) ++ bmkeps(r2)
+  case ASTAR(bs, r) => bs ++ List(Z)
+}
+
+// derivative of a regular expression w.r.t. a character
+def bder(c: Char, r: ARexp) : ARexp = r match {
+  case AZERO => AZERO
+  case AONE(_) => AZERO
+  case APRED(bs, f, _) => if (f(c)) AONE(bs:::List(C(c))) else AZERO
+  case AALTS(bs, rs) => AALTS(bs, rs.map(bder(c, _)))
+  case ASEQ(bs, r1, r2) => 
+    if (bnullable(r1)) AALT(bs, ASEQ(Nil, bder(c, r1), r2), fuse(bmkeps(r1), bder(c, r2)))
+    else ASEQ(bs, bder(c, r1), r2)
+  case ASTAR(bs, r) => ASEQ(bs, fuse(List(S), bder(c, r)), ASTAR(Nil, r))
+}
+
+
+// derivative w.r.t. a string (iterates bder)
+@tailrec
+def bders (s: List[Char], r: ARexp) : ARexp = s match {
+  case Nil => r
+  case c::s => bders(s, bder(c, r))
+}
+
+def flats(rs: List[ARexp]): List[ARexp] = rs match {
+    case Nil => Nil
+    case AZERO :: rs1 => flats(rs1)
+    case AALTS(bs, rs1) :: rs2 => rs1.map(fuse(bs, _)) ::: flats(rs2)
+    case r1 :: rs2 => r1 :: flats(rs2)
+}
+
+def bsimp(r: ARexp): ARexp = r match {
+  case ASEQ(bs1, r1, r2) => (bsimp(r1), bsimp(r2)) match {
+      case (AZERO, _) => AZERO
+      case (_, AZERO) => AZERO
+      case (AONE(bs2), r2s) => fuse(bs1 ++ bs2, r2s)
+      case (r1s, r2s) => ASEQ(bs1, r1s, r2s)
+  }
+  case AALTS(bs1, rs) => distinctBy(flats(rs.map(bsimp)), erase) match {
+    case Nil => AZERO
+    case r :: Nil => fuse(bs1, r)
+    case rs => AALTS(bs1, rs)  
+  }
+  //case ASTAR(bs1, r1) => ASTAR(bs1, bsimp(r1))
+  case r => r
+}
+
+def bders_simp (s: List[Char], r: ARexp) : ARexp = s match {
+  case Nil => r
+  case c::s => bders_simp(s, bsimp(bder(c, r)))
+}
+
+def blex_simp(r: ARexp, s: List[Char]) : Bits = s match {
+  case Nil => if (bnullable(r)) bmkeps(r)
+	      else throw new Exception("Not matched")
+  case c::cs => blex_simp(bsimp(bder(c, r)), cs)
+}
+
+
+def blexing_simp(r: Rexp, s: String) : Val = 
+ decode(r, blex_simp(internalise(r), s.toList))
+
+
+def btokenise_simp(r: Rexp, s: String) = 
+  env(blexing_simp(r, s)).map(esc2)
+
+
+
+// INCLUDING SIMPLIFICATION UNDER STARS
+
+def bsimp_full(r: ARexp): ARexp = r match {
+  case ASEQ(bs1, r1, r2) => (bsimp_full(r1), bsimp_full(r2)) match {
+      case (AZERO, _) => AZERO
+      case (_, AZERO) => AZERO
+      case (AONE(bs2), r2s) => fuse(bs1 ++ bs2, r2s)
+      case (r1s, r2s) => ASEQ(bs1, r1s, r2s)
+  }
+  case AALTS(bs1, rs) => distinctBy(flats(rs.map(bsimp_full)), erase) match {
+    case Nil => AZERO
+    case r :: Nil => fuse(bs1, r)
+    case rs => AALTS(bs1, rs)  
+  }
+  case ASTAR(bs1, r1) => ASTAR(bs1, bsimp_full(r1))
+  case r => r
+}
+
+def bders_simp_full(s: List[Char], r: ARexp) : ARexp = s match {
+  case Nil => r
+  case c::s => bders_simp_full(s, bsimp_full(bder(c, r)))
+}
+
+def blex_simp_full(r: ARexp, s: List[Char]) : Bits = s match {
+  case Nil => if (bnullable(r)) bmkeps(r)
+	      else throw new Exception("Not matched")
+  case c::cs => blex_simp_full(bsimp_full(bder(c, r)), cs)
+}
+
+
+def blexing_simp_full(r: Rexp, s: String) : Val = 
+ decode(r, blex_simp_full(internalise(r), s.toList))
+
+
+def btokenise_simp_full(r: Rexp, s: String) = env(blexing_simp_full(r, s)).map(esc2)
+
+// bders2 for strings in the ALTS case
+
+def bders2_simp(s: List[Char], a: ARexp) : ARexp = {
+  //println(s"s = ${s.length}   a = ${asize(a)}")
+  //Console.readLine
+  (s, a) match {
+  case (Nil, r) => r
+  case (s, AZERO) => AZERO
+  case (s, AONE(_)) => AZERO
+  case (s, APRED(bs, f, _)) => 
+    if (f(s.head) && s.tail == Nil) AONE(bs:::List(C(s.head))) else AZERO
+  case (s, AALTS(bs, rs)) => bsimp(AALTS(bs, rs.map(bders2_simp(s, _)))) 
+  case (c::s, r) => bders2_simp(s, bsimp(bder(c, r)))
+}}
+
+
+
+def blexing2_simp(r: Rexp, s: String) : Val = {
+  val bder = bders2_simp(s.toList, internalise(r))
+  if (bnullable(bder)) decode(r, bmkeps(bder)) else 
+  throw new Exception("Not matched")
+
+}
+
+def btokenise2_simp(r: Rexp, s: String) = 
+  env(blexing2_simp(r, s)).map(esc2)
+
+
+// Parser for regexes
+
+case class Parser(s: String) {
+  var i = 0
+  
+  def peek() = s(i)
+  def eat(c: Char) = 
+    if (c == s(i)) i = i + 1 else throw new Exception("Expected " + c + " got " + s(i))
+  def next() = { i = i + 1; s(i - 1) }
+  def more() = s.length - i > 0
+
+  def Regex() : Rexp = {
+    val t = Term();
+    if (more() && peek() == '|') {
+      eat ('|') ; 
+      ALT(t, Regex()) 
+    } 
+    else t
+  }
+
+  def Term() : Rexp = {
+    var f : Rexp = 
+      if (more() && peek() != ')' && peek() != '|') Factor() else ONE;
+    while (more() && peek() != ')' && peek() != '|') {
+      f = SEQ(f, Factor()) ;
+    }
+    f
+  }
+
+  def Factor() : Rexp = {
+    var b = Base();
+    while (more() && peek() == '*') {
+      eat('*') ;
+      b = STAR(b) ;
+    }
+    while (more() && peek() == '?') {
+      eat('?') ;
+      b = ALT(b, ONE) ;
+    }
+    while (more() && peek() == '+') {
+      eat('+') ;
+      b = SEQ(b, STAR(b)) ;
+    }
+    b
+  }
+
+  def Base() : Rexp = {
+    peek() match {
+      case '(' => { eat('(') ; val r = Regex(); eat(')') ; r }   // if groups should be groups RECD("",r) }
+      case '.' => { eat('.'); ANYCHAR }
+      case _ => CHAR(next())
+    }
+  }
+}
+
+// two simple examples for the regex parser
+
+println("two simple examples for the regex parser")
+
+println(string(Parser("a|(bc)*").Regex()))
+println(string(Parser("(a|b)*(babab(a|b)*bab|bba(a|b)*bab)(a|b)*").Regex()))
+
+
+
+//System.exit(0)
+
+//   Testing
+//============
+
+def time[T](code: => T) = {
+  val start = System.nanoTime()
+  val result = code
+  val end = System.nanoTime()
+  ((end - start)/1.0e9).toString
+  //result
+}
+
+def timeR[T](code: => T) = {
+  val start = System.nanoTime()
+  for (i <- 1 to 10) code
+  val result = code
+  val end = System.nanoTime()
+  (result, (end - start))
+}
+
+//size: of a Aregx for testing purposes 
+def size(r: Rexp) : Int = r match {
+  case ZERO => 1
+  case ONE => 1
+  case PRED(_,_) => 1
+  case SEQ(r1, r2) => 1 + size(r1) + size(r2)
+  case ALTS(rs) => 1 + rs.map(size).sum
+  case STAR(r) => 1 + size(r)
+  case RECD(_, r) => size(r)
+}
+
+def asize(a: ARexp) = size(erase(a))
+
+
+// Lexing Rules for a Small While Language
+
+//symbols
+val SYM = PRED("abcdefghijklmnopqrstuvwxyzABCDEFGHIJKLMNOPQRSTUVWXYZ".contains(_), "SYM")
+//digits
+val DIGIT = PRED("0123456789".contains(_), "NUM")
+//identifiers
+val ID = SYM ~ (SYM | DIGIT).% 
+//numbers
+val NUM = STAR(DIGIT)
+//keywords
+val KEYWORD : Rexp = "skip" | "while" | "do" | "if" | "then" | "else" | "read" | "write" | "true" | "false"
+//semicolons
+val SEMI: Rexp = ";"
+//operators
+val OP: Rexp = ":=" | "==" | "-" | "+" | "*" | "!=" | "<" | ">" | "<=" | ">=" | "%" | "/"
+//whitespaces
+val WHITESPACE = PLUS(" " | "\n" | "\t")
+//parentheses
+val RPAREN: Rexp = ")"
+val LPAREN: Rexp = "("
+val BEGIN: Rexp = "{"
+val END: Rexp = "}"
+//strings...but probably needs not
+val STRING: Rexp = "\"" ~ SYM.% ~ "\""
+
+
+
+val WHILE_REGS = (("k" $ KEYWORD) | 
+                  ("i" $ ID) | 
+                  ("o" $ OP) | 
+                  ("n" $ NUM) | 
+                  ("s" $ SEMI) | 
+                  ("str" $ STRING) |
+                  ("p" $ (LPAREN | RPAREN)) | 
+                  ("b" $ (BEGIN | END)) | 
+                  ("w" $ WHITESPACE)).%
+
+
+// Some Small Tests
+//==================
+
+println("Small tests")
+
+val q = STAR(STAR("bb" | "ab"))
+val qs = "bbb"
+
+println("Size Bit  " + asize(bders_simp(qs.toList, internalise(q))))
+println("Size Bitf " + asize(bders_simp_full(qs.toList, internalise(q))))
+println("Size Bit2 " + asize(bders2_simp(qs.toList, internalise(q))))
+println("Size Old  " + size(ders_simp(qs.toList, q)))
+println("Size Pder " + psize(pders_simp(qs.toList, q)))
+
+
+val re1 = STAR("a" | "aa")
+println(astring(bders_simp("".toList, internalise(re1))))
+println(astring(bders_simp("a".toList, internalise(re1))))
+println(astring(bders_simp("aa".toList, internalise(re1))))
+println(astring(bders_simp("aaa".toList, internalise(re1))))
+println(astring(bders_simp("aaaaaa".toList, internalise(re1))))
+println(astring(bders_simp("aaaaaaaaa".toList, internalise(re1))))
+println(astring(bders_simp("aaaaaaaaaaaa".toList, internalise(re1))))
+println(astring(bders_simp("aaaaaaaaaaaaaaaaaaaaaaaaa".toList, internalise(re1))))
+println(astring(bders_simp("aaaaaabaaaabbbbbaaaaaaaaaaaaaaa".toList, internalise(re1))))
+
+
+for (i <- 0 to 100 by 5) {
+  //print("Old: " + time(tokenise_simp(re1, "a" * i)))
+  print(" Bit: " + time(btokenise_simp(re1, "a" * i)))
+  print(" Bit full simp: " + time(btokenise_simp_full(re1, "a" * i)))
+  println(" Bit2: " + time(btokenise2_simp(re1, "a" * i)))
+}
+
+Console.readLine
+
+
+// Bigger Tests
+//==============
+
+
+println("Big tests")
+
+val fib_prog = """
+write "Fib";
+read n;
+minus1 := 0;
+minus2 := 1;
+while n > 0 do {
+  temp := minus2;
+  minus2 := minus1 + minus2;
+  minus1 := temp;
+  n := n - 1
+};
+write "Result";
+write minus2
+"""
+
+
+println("fib prog tests :")
+println(tokenise_simp(WHILE_REGS, fib_prog))
+println(btokenise_simp(WHILE_REGS, fib_prog))
+println("equal? " + (tokenise_simp(WHILE_REGS, fib_prog) == btokenise_simp(WHILE_REGS, fib_prog)))
+
+for (i <- 1 to 20) {
+  print("Old: " + time(tokenise_simp(WHILE_REGS, fib_prog * i)))
+  print(" Bit: " + time(btokenise_simp(WHILE_REGS, fib_prog * i)))
+  println(" Bit full simp: " + time(btokenise_simp_full(WHILE_REGS, fib_prog * i)))
+  //println(" Bit2: " + time(btokenise2_simp(WHILE_REGS, fib_prog * i)))
+}
+
+
+println("Original " + size(WHILE_REGS))
+println("Size Bit  " + asize(bders_simp((fib_prog * 1).toList, internalise(WHILE_REGS))))
+println("Size Bitf " + asize(bders_simp_full((fib_prog * 1).toList, internalise(WHILE_REGS))))
+println("Size Bit2 " + asize(bders2_simp((fib_prog * 1).toList, internalise(WHILE_REGS))))
+println("Size Old  " + size(ders_simp((fib_prog * 1).toList, WHILE_REGS)))
+println("Size Pder " + psize(pders_simp((fib_prog * 1).toList, WHILE_REGS)))
+
+System.exit(0)
+
+println("Internal sizes test OK or strange")
+
+def perc(p1: Double, p2: Double) : String =
+  f"${(((p1 - p2) / p2) * 100.0) }%5.0f" + "%"
+
+def ders_test(n: Int, s: List[Char], r: Rexp, a: ARexp) : (Rexp, ARexp) = s match {
+  case Nil => (r, a)
+  case c::s => {
+    // derivative 
+    val (rd1, tr1) = timeR(der(c, r))
+    val (ad1, ta1) = timeR(bder(c, a))
+    val trs1 = f"${tr1}%.5f"
+    val tas1 = f"${ta1}%.5f"
+    if (tr1 < ta1) println(s"Time strange der  (step) ${n} ${perc(ta1, tr1)} sizes  der ${size(rd1)} ${asize(ad1)}")
+    //simplification
+    val (rd, tr) = timeR(simp(rd1)._1)
+    val (ad, ta) = timeR(bsimp(ad1))
+    val trs = f"${tr}%.5f"
+    val tas = f"${ta}%.5f"
+    //full simplification
+    val (adf, taf) = timeR(bsimp_full(ad1))
+    if (tr < ta) println(s"Time strange simp (step) ${n} ${perc(ta, tr)} sizes simp ${size(rd)} ${asize(ad)}")
+    if (n == 1749 || n == 1734) {
+      println{s"Aregex before bder (size: ${asize(a)})\n ${string(erase(a))}"}
+      println{s"Aregex after bder (size: ${asize(ad1)})\n ${string(erase(ad1))}"}
+      println{s"Aregex after bsimp (size: ${asize(ad)})\n ${string(erase(ad))}"}
+      println{s"Aregex after bsimp_full (size: ${asize(adf)})\n ${string(erase(adf))}"}
+    }
+    ders_test(n + 1, s, rd, ad)
+  }
+}
+
+val prg = (fib_prog * 10).toList
+ders_test(0, prg, WHILE_REGS, internalise(WHILE_REGS))
+
+
+//testing the two lexings produce the same value
+//enumerates strings of length n over alphabet cs
+def strs(n: Int, cs: String) : Set[String] = {
+  if (n == 0) Set("")
+  else {
+    val ss = strs(n - 1, cs)
+    ss ++
+    (for (s <- ss; c <- cs.toList) yield c + s)
+  }
+}
+def enum(n: Int, s: String) : Stream[Rexp] = n match {
+  case 0 => ZERO #:: ONE #:: s.toStream.map(CHAR)
+  case n => {  
+    val rs = enum(n - 1, s)
+    rs #:::
+    (for (r1 <- rs; r2 <- rs) yield ALT(r1, r2)) #:::
+    (for (r1 <- rs; r2 <- rs) yield SEQ(r1, r2)) #:::
+    (for (r1 <- rs) yield STAR(r1))
+  }
+}
+
+//tests blexing and lexing
+def tests(ss: Set[String])(r: Rexp) = {
+  //println(s"Testing ${r}")
+  for (s <- ss.par) yield {
+    val res1 = Try(Some(lexing_simp(r, s))).getOrElse(None)
+    val res2 = Try(Some(blexing_simp(r, s))).getOrElse(None)
+    if (res1 != res2) 
+      { println(s"Disagree on ${r} and ${s}")
+	println(s"   ${res1} !=  ${res2}")
+	Some((r, s)) } else None
+  }
+}
+
+
+println("Partial searching: ")
+enum(2, "abc").map(tests(strs(3, "abc"))).toSet
+
+
+
+
--- a/exps/both.scala	Sat Feb 23 21:52:06 2019 +0000
+++ b/exps/both.scala	Wed Mar 13 10:36:29 2019 +0000
@@ -4,6 +4,7 @@
 import scala.annotation.tailrec   
 import scala.util.Try
 
+// for escaping strings
 def escape(raw: String) : String = {
   import scala.reflect.runtime.universe._
   Literal(Constant(raw)).toString
@@ -121,7 +122,7 @@
 
 // nullable function: tests whether the regular 
 // expression can recognise the empty string
-def nullable (r: Rexp) : Boolean = r match {
+def nullable(r: Rexp) : Boolean = r match {
   case ZERO => false
   case ONE => true
   case PRED(_, _) => false
@@ -132,7 +133,7 @@
 }
 
 // derivative of a regular expression w.r.t. a character
-def der (c: Char, r: Rexp) : Rexp = r match {
+def der(c: Char, r: Rexp) : Rexp = r match {
   case ZERO => ZERO
   case ONE => ZERO
   case PRED(f, _) => if (f(c)) ONE else ZERO
@@ -236,9 +237,9 @@
     val (r2s, f2s) = simp(r2)
     (r1s, r2s) match {
       case (ZERO, _) => (ZERO, F_ERROR)
-      case (_, ZERO) => (ZERO, F_ERROR)
+      //case (_, ZERO) => (ZERO, F_ERROR)
       case (ONE, _) => (r2s, F_SEQ_Empty1(f1s, f2s))
-      case (_, ONE) => (r1s, F_SEQ_Empty2(f1s, f2s))
+      //case (_, ONE) => (r1s, F_SEQ_Empty2(f1s, f2s))
       case _ => (SEQ(r1s,r2s), F_SEQ(f1s, f2s))
     }
   }
@@ -304,6 +305,20 @@
 //--------------------------------------------------------------------
 // BITCODED PART
 
+def retrieve(r: ARexp, v: Val) : List[Boolean] = (r, v) match {
+  case (AONE(bs), Empty) => bs
+  case (ACHAR(bs, c), Chr(d)) => bs
+  case (AALTS(bs, r::Nil), v) => bs ++ retrieve(r, v)
+  case (AALTS(bs, r::rs), Left(v)) => bs ++ retrieve(r, v)
+  case (AALTS(bs, r::rs), Right(v)) => bs ++ retrieve(AALTS(Nil, rs), v)
+  case (ASEQ(bs, r1, r2), Sequ(v1, v2)) => 
+    bs ++ retrieve(r1, v1) ++ retrieve(r2, v2)
+  case (ASTAR(bs, r), Stars(Nil)) => bs ++ List(S)
+  case (ASTAR(bs, r), Stars(v::vs)) => 
+     bs ++ List(Z) ++ retrieve(r, v) ++ retrieve(ASTAR(Nil, r), Stars(vs))
+}
+
+
 
 def fuse(bs: Bits, r: ARexp) : ARexp = r match {
   case AZERO => AZERO
@@ -464,6 +479,41 @@
 def btokenise_simp(r: Rexp, s: String) = 
   env(blexing_simp(r, s)).map(esc2)
 
+// Quick example
+
+val r : Rexp = ZERO | "a" 
+
+lexing(r, "a")
+
+val a0 = internalise(r)
+val a1 = bder('a', a0)
+val a1s = bsimp(bder('a', a0))
+
+val a2 = bmkeps(a1)
+val a2s = bmkeps(a1s)
+
+val v  = decode(r, a2)
+val vs  = decode(r, a2s)
+
+
+
+val Rr : Rexp = ONE ~ "a" 
+
+lexing(Rr, "a")
+
+val Ra0 = internalise(Rr)
+astring(Ra0)
+val Ra1 = bder('a', Ra0)
+astring(Ra1)
+val Ra1s = bsimp(bder('a', Ra0))
+astring(Ra1s)
+
+val Ra2 = bmkeps(Ra1)
+val Ra2s = bmkeps(Ra1s)
+
+val Rv  = decode(Rr, Ra2)
+val Rvs  = decode(Rr, Ra2s)
+
 
 
 // INCLUDING SIMPLIFICATION UNDER STARS
@@ -530,6 +580,11 @@
   env(blexing2_simp(r, s)).map(esc2)
 
 
+
+
+
+
+//============================================
 // Parser for regexes
 
 case class Parser(s: String) {
@@ -594,7 +649,7 @@
 
 
 
-System.exit(0)
+//System.exit(0)
 
 //   Testing
 //============
--- a/thys/BitCoded.thy	Sat Feb 23 21:52:06 2019 +0000
+++ b/thys/BitCoded.thy	Wed Mar 13 10:36:29 2019 +0000
@@ -87,7 +87,7 @@
 
 section {* Annotated Regular Expressions *}
 
-datatype arexp =
+datatype arexp = 
   AZERO
 | AONE "bit list"
 | ACHAR "bit list" char
@@ -107,6 +107,13 @@
 | "fuse bs (ASEQ cs r1 r2) = ASEQ (bs @ cs) r1 r2"
 | "fuse bs (ASTAR cs r) = ASTAR (bs @ cs) r"
 
+lemma fuse_append:
+  shows "fuse (bs1 @ bs2) r = fuse bs1 (fuse bs2 r)"
+  apply(induct r)
+  apply(auto)
+  done
+
+
 fun intern :: "rexp \<Rightarrow> arexp" where
   "intern ZERO = AZERO"
 | "intern ONE = AONE []"
@@ -451,5 +458,901 @@
 qed
 
 
+fun distinctBy :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b set \<Rightarrow> 'a list"
+  where
+  "distinctBy [] f acc = []"
+| "distinctBy (x#xs) f acc = 
+     (if (f x) \<in> acc then distinctBy xs f acc 
+      else x # (distinctBy xs f ({f x} \<union> acc)))"
+
+fun flts :: "arexp list \<Rightarrow> arexp list"
+  where 
+  "flts [] = []"
+| "flts (AZERO # rs) = flts rs"
+| "flts ((AALTs bs  rs1) # rs) = (map (fuse bs) rs1) @ flts rs"
+| "flts (r1 # rs) = r1 # flts rs"
+
+(*
+lemma flts_map:
+  assumes "\<forall>r \<in> set rs. f r = r"
+  shows "map f (flts rs) = flts (map f rs)"
+  using assms
+  apply(induct rs rule: flts.induct)
+        apply(simp_all)
+       apply(case_tac rs)
+  apply(simp)
+*)
+
+fun bsimp_ASEQ :: "bit list \<Rightarrow> arexp \<Rightarrow> arexp \<Rightarrow> arexp"
+  where
+  "bsimp_ASEQ _ AZERO _ = AZERO"
+| "bsimp_ASEQ _ _ AZERO = AZERO"
+| "bsimp_ASEQ bs1 (AONE bs2) r2 = fuse (bs1 @ bs2) r2"
+| "bsimp_ASEQ bs1 r1 r2 = ASEQ  bs1 r1 r2"
+
+
+fun bsimp_AALTs :: "bit list \<Rightarrow> arexp list \<Rightarrow> arexp"
+  where
+  "bsimp_AALTs _ [] = AZERO"
+| "bsimp_AALTs bs1 [r] = fuse bs1 r"
+| "bsimp_AALTs bs1 rs = AALTs bs1 rs"
+
+
+fun bsimp :: "arexp \<Rightarrow> arexp" 
+  where
+  "bsimp (ASEQ bs1 r1 r2) = bsimp_ASEQ bs1 (bsimp r1) (bsimp r2)"
+| "bsimp (AALTs bs1 rs) = bsimp_AALTs bs1 (flts (map bsimp rs))"
+| "bsimp r = r"
+
+fun 
+  bders_simp :: "arexp \<Rightarrow> string \<Rightarrow> arexp"
+where
+  "bders_simp r [] = r"
+| "bders_simp r (c # s) = bders_simp (bsimp (bder c r)) s"
+
+definition blexer_simp where
+ "blexer_simp r s \<equiv> if bnullable (bders_simp (intern r) s) then 
+                decode (bmkeps (bders_simp (intern r) s)) r else None"
+
+
+lemma bders_simp_append:
+  shows "bders_simp r (s1 @ s2) = bders_simp (bders_simp r s1) s2"
+  apply(induct s1 arbitrary: r s2)
+   apply(simp)
+  apply(simp)
+  done
+
+
+lemma L_bsimp_ASEQ:
+  "L (SEQ (erase r1) (erase r2)) = L (erase (bsimp_ASEQ bs r1 r2))"
+  apply(induct bs r1 r2 rule: bsimp_ASEQ.induct)
+  apply(simp_all)
+  by (metis erase_fuse fuse.simps(4))
+
+lemma L_bsimp_AALTs:
+  "L (erase (AALTs bs rs)) = L (erase (bsimp_AALTs bs rs))"
+  apply(induct bs rs rule: bsimp_AALTs.induct)
+  apply(simp_all add: erase_fuse)
+  done
+
+lemma L_erase_AALTs:
+  shows "L (erase (AALTs bs rs)) = \<Union> (L ` erase ` (set rs))"
+  apply(induct rs)
+   apply(simp)
+  apply(simp)
+  apply(case_tac rs)
+   apply(simp)
+  apply(simp)
+  done
+
+lemma L_erase_flts:
+  shows "\<Union> (L ` erase ` (set (flts rs))) = \<Union> (L ` erase ` (set rs))"
+  apply(induct rs rule: flts.induct)
+        apply(simp_all)
+  apply(auto)
+  using L_erase_AALTs erase_fuse apply auto[1]
+  by (simp add: L_erase_AALTs erase_fuse)
+
+
+lemma L_bsimp_erase:
+  shows "L (erase r) = L (erase (bsimp r))"
+  apply(induct r)
+  apply(simp)
+  apply(simp)
+  apply(simp)
+  apply(auto simp add: Sequ_def)[1]
+  apply(subst L_bsimp_ASEQ[symmetric])
+  apply(auto simp add: Sequ_def)[1]
+  apply(subst (asm)  L_bsimp_ASEQ[symmetric])
+  apply(auto simp add: Sequ_def)[1]
+   apply(simp)
+   apply(subst L_bsimp_AALTs[symmetric])
+   defer
+   apply(simp)
+  apply(subst (2)L_erase_AALTs)
+  apply(subst L_erase_flts)
+  apply(auto)
+   apply (simp add: L_erase_AALTs)
+  using L_erase_AALTs by blast
+
+
+lemma bsimp_ASEQ1:
+  assumes "r1 \<noteq> AZERO" "r2 \<noteq> AZERO" "\<forall>bs. r1 \<noteq> AONE bs"
+  shows "bsimp_ASEQ bs r1 r2 = ASEQ bs r1 r2"
+  using assms
+  apply(induct bs r1 r2 rule: bsimp_ASEQ.induct)
+  apply(auto)
+  done
+
+lemma bsimp_ASEQ2:
+  shows "bsimp_ASEQ bs (AONE bs1) r2 = fuse (bs @ bs1) r2"
+  apply(induct r2)
+  apply(auto)
+  done
+
+
+lemma L_bders_simp:
+  shows "L (erase (bders_simp r s)) = L (erase (bders r s))"
+  apply(induct s arbitrary: r rule: rev_induct)
+   apply(simp)
+  apply(simp)
+  apply(simp add: ders_append)
+  apply(simp add: bders_simp_append)
+  apply(simp add: L_bsimp_erase[symmetric])
+  by (simp add: der_correctness)
+
+lemma b1:
+  "bsimp_ASEQ bs1 (AONE bs) r =  fuse (bs1 @ bs) r" 
+  apply(induct r)
+       apply(auto)
+  done
+
+lemma b2:
+  assumes "bnullable r"
+  shows "bmkeps (fuse bs r) = bs @ bmkeps r"
+  by (simp add: assms bmkeps_retrieve bnullable_correctness erase_fuse mkeps_nullable retrieve_fuse2)
+
+lemma b3:
+  shows "bnullable r = bnullable (bsimp r)"
+  using L_bsimp_erase bnullable_correctness nullable_correctness by auto
+
+
+lemma b4:
+  shows "bnullable (bders_simp r s) = bnullable (bders r s)"
+  by (metis L_bders_simp bnullable_correctness lexer.simps(1) lexer_correct_None option.distinct(1))
+  
+
+lemma q1:
+  assumes "\<forall>r \<in> set rs. bmkeps(bsimp r) = bmkeps r"
+  shows "map (\<lambda>r. bmkeps(bsimp r)) rs = map bmkeps rs"
+  using assms
+  apply(induct rs)
+  apply(simp)
+  apply(simp)
+  done
+
+lemma q3:
+  assumes "\<exists>r \<in> set rs. bnullable r"
+  shows "bmkeps (AALTs bs rs) = bmkeps (bsimp_AALTs bs rs)"
+  using assms
+  apply(induct bs rs rule: bsimp_AALTs.induct)
+    apply(simp)
+   apply(simp)
+  apply (simp add: b2)
+  apply(simp)
+  done
+
+lemma qq1:
+  assumes "\<exists>r \<in> set rs. bnullable r"
+  shows "bmkeps (AALTs bs (rs @ rs1)) = bmkeps (AALTs bs rs)"
+  using assms
+  apply(induct rs arbitrary: rs1 bs)
+  apply(simp)
+  apply(simp)
+  by (metis Nil_is_append_conv bmkeps.simps(4) neq_Nil_conv r0 split_list_last)
+
+lemma qq2:
+  assumes "\<forall>r \<in> set rs. \<not> bnullable r" "\<exists>r \<in> set rs1. bnullable r"
+  shows "bmkeps (AALTs bs (rs @ rs1)) = bmkeps (AALTs bs rs1)"
+  using assms
+  apply(induct rs arbitrary: rs1 bs)
+  apply(simp)
+  apply(simp)
+  by (metis append_assoc in_set_conv_decomp r1 r2)
+  
+lemma qq3:
+  shows "bnullable (AALTs bs rs) = (\<exists>r \<in> set rs. bnullable r)"
+  apply(induct rs arbitrary: bs)
+  apply(simp)
+  apply(simp)
+  done
+
+lemma q3a:
+  assumes "\<exists>r \<in> set rs. bnullable r"
+  shows "bmkeps (AALTs bs (map (fuse bs1) rs)) = bmkeps (AALTs (bs@bs1) rs)"
+  using assms
+  apply(induct rs arbitrary: bs bs1)
+   apply(simp)
+  apply(simp)
+  apply(auto)
+   apply (metis append_assoc b2 bnullable_correctness erase_fuse r0)
+  apply(case_tac "bnullable a")
+   apply (metis append.assoc b2 bnullable_correctness erase_fuse r0)
+  apply(case_tac rs)
+  apply(simp)
+  apply(simp)
+  apply(auto)[1]
+   apply (metis bnullable_correctness erase_fuse)+
+  done
+
+lemma qq4:
+  assumes "\<exists>x\<in>set list. bnullable x"
+  shows "\<exists>x\<in>set (flts list). bnullable x"
+  using assms
+  apply(induct list rule: flts.induct)
+        apply(auto)
+  by (metis UnCI bnullable_correctness erase_fuse imageI)
+  
+
+lemma qs3:
+  assumes "\<exists>r \<in> set rs. bnullable r"
+  shows "bmkeps (AALTs bs rs) = bmkeps (AALTs bs (flts rs))"
+  using assms
+  apply(induct rs arbitrary: bs taking: size rule: measure_induct)
+  apply(case_tac x)
+  apply(simp)
+  apply(simp)
+  apply(case_tac a)
+       apply(simp)
+       apply (simp add: r1)
+      apply(simp)
+      apply (simp add: r0)
+     apply(simp)
+     apply(case_tac "flts list")
+      apply(simp)
+  apply (metis L_erase_AALTs L_erase_flts L_flat_Prf1 L_flat_Prf2 Prf_elims(1) bnullable_correctness erase.simps(4) mkeps_nullable r2)
+     apply(simp)
+     apply (simp add: r1)
+    prefer 3
+    apply(simp)
+    apply (simp add: r0)
+   prefer 2
+   apply(simp)
+  apply(case_tac "\<exists>x\<in>set x52. bnullable x")
+  apply(case_tac "list")
+    apply(simp)
+    apply (metis b2 fuse.simps(4) q3a r2)
+   apply(erule disjE)
+    apply(subst qq1)
+     apply(auto)[1]
+     apply (metis bnullable_correctness erase_fuse)
+    apply(simp)
+     apply (metis b2 fuse.simps(4) q3a r2)
+    apply(simp)
+    apply(auto)[1]
+     apply(subst qq1)
+      apply (metis bnullable_correctness erase_fuse image_eqI set_map)
+     apply (metis b2 fuse.simps(4) q3a r2)
+  apply(subst qq1)
+      apply (metis bnullable_correctness erase_fuse image_eqI set_map)
+    apply (metis b2 fuse.simps(4) q3a r2)
+   apply(simp)
+   apply(subst qq2)
+     apply (metis bnullable_correctness erase_fuse imageE set_map)
+  prefer 2
+  apply(case_tac "list")
+     apply(simp)
+    apply(simp)
+   apply (simp add: qq4)
+  apply(simp)
+  apply(auto)
+   apply(case_tac list)
+    apply(simp)
+   apply(simp)
+   apply (simp add: r0)
+  apply(case_tac "bnullable (ASEQ x41 x42 x43)")
+   apply(case_tac list)
+    apply(simp)
+   apply(simp)
+   apply (simp add: r0)
+  apply(simp)
+  using qq4 r1 r2 by auto
+
+lemma  k0:
+  shows "flts (r # rs1) = flts [r] @ flts rs1"
+  apply(induct r arbitrary: rs1)
+   apply(auto)
+  done
+
+lemma k1:
+  assumes "\<And>x2aa. \<lbrakk>x2aa \<in> set x2a; bnullable x2aa\<rbrakk> \<Longrightarrow> bmkeps x2aa = bmkeps (bsimp x2aa)"
+          "\<exists>x\<in>set x2a. bnullable x"
+        shows "bmkeps (AALTs x1 (flts x2a)) = bmkeps (AALTs x1 (flts (map bsimp x2a)))"
+  using assms
+  apply(induct x2a)
+  apply fastforce
+  apply(simp)
+  apply(subst k0)
+  apply(subst (2) k0)
+  apply(auto)[1]
+  apply (metis b3 k0 list.set_intros(1) qs3 r0)
+  by (smt b3 imageI insert_iff k0 list.set(2) qq3 qs3 r0 r1 set_map)
+  
+  
+  
+lemma bmkeps_simp:
+  assumes "bnullable r"
+  shows "bmkeps r = bmkeps (bsimp r)"
+  using  assms
+  apply(induct r)
+       apply(simp)
+      apply(simp)
+     apply(simp)
+    apply(simp)
+    prefer 3
+  apply(simp)
+   apply(case_tac "bsimp r1 = AZERO")
+    apply(simp)
+    apply(auto)[1]
+  apply (metis L_bsimp_erase L_flat_Prf1 L_flat_Prf2 Prf_elims(1) bnullable_correctness erase.simps(1) mkeps_nullable)
+ apply(case_tac "bsimp r2 = AZERO")
+    apply(simp)  
+    apply(auto)[1]
+  apply (metis L_bsimp_erase L_flat_Prf1 L_flat_Prf2 Prf_elims(1) bnullable_correctness erase.simps(1) mkeps_nullable)
+  apply(case_tac "\<exists>bs. bsimp r1 = AONE bs")
+    apply(auto)[1]
+    apply(subst b1)
+    apply(subst b2)
+  apply(simp add: b3[symmetric])
+    apply(simp)
+   apply(subgoal_tac "bsimp_ASEQ x1 (bsimp r1) (bsimp r2) = ASEQ x1 (bsimp r1) (bsimp r2)")
+    prefer 2
+  apply (smt b3 bnullable.elims(2) bsimp_ASEQ.simps(17) bsimp_ASEQ.simps(19) bsimp_ASEQ.simps(20) bsimp_ASEQ.simps(21) bsimp_ASEQ.simps(22) bsimp_ASEQ.simps(24) bsimp_ASEQ.simps(25) bsimp_ASEQ.simps(26) bsimp_ASEQ.simps(27) bsimp_ASEQ.simps(29) bsimp_ASEQ.simps(30) bsimp_ASEQ.simps(31))
+   apply(simp)
+  apply(simp)
+  apply(subst q3[symmetric])
+   apply simp
+  using b3 qq4 apply auto[1]
+  apply(subst qs3)
+   apply simp
+  using k1 by blast
+
+thm bmkeps_retrieve bmkeps_simp bder_retrieve
+
+lemma bmkeps_bder_AALTs:
+  assumes "\<exists>r \<in> set rs. bnullable (bder c r)" 
+  shows "bmkeps (bder c (bsimp_AALTs bs rs)) = bmkeps (bsimp_AALTs bs (map (bder c) rs))"
+  using assms
+  apply(induct rs)
+   apply(simp)
+  apply(simp)
+  apply(auto)
+  apply(case_tac rs)
+    apply(simp)
+  apply (metis (full_types) Prf_injval bder_retrieve bmkeps_retrieve bnullable_correctness erase_bder erase_fuse mkeps_nullable retrieve_fuse2)
+   apply(simp)
+  apply(case_tac  rs)
+   apply(simp_all)
+  done
+
+
+
+
+lemma MAIN_decode:
+  assumes "\<Turnstile> v : ders s r"
+  shows "Some (flex r id s v) = decode (retrieve (bders_simp (intern r) s) v) r"
+  using assms
+proof (induct s arbitrary: v rule: rev_induct)
+  case Nil
+  have "\<Turnstile> v : ders [] r" by fact
+  then have "\<Turnstile> v : r" by simp
+  then have "Some v = decode (retrieve (intern r) v) r"
+    using decode_code retrieve_code by auto
+  then show "Some (flex r id [] v) = decode (retrieve (bders_simp (intern r) []) v) r"
+    by simp
+next
+  case (snoc c s v)
+  have IH: "\<And>v. \<Turnstile> v : ders s r \<Longrightarrow> 
+     Some (flex r id s v) = decode (retrieve (bders_simp (intern r) s) v) r" by fact
+  have asm: "\<Turnstile> v : ders (s @ [c]) r" by fact
+  then have asm2: "\<Turnstile> injval (ders s r) c v : ders s r" 
+    by(simp add: Prf_injval ders_append)
+  have "Some (flex r id (s @ [c]) v) = Some (flex r id s (injval (ders s r) c v))"
+    by (simp add: flex_append)
+  also have "... = decode (retrieve (bders_simp (intern r) s) (injval (ders s r) c v)) r"
+    using asm2 IH by simp
+  also have "... = decode (retrieve (bder c (bders_simp (intern r) s)) v) r"
+    using asm bder_retrieve ders_append 
+    apply -
+    apply(drule_tac x="v" in meta_spec)
+    apply(drule_tac x="c" in meta_spec)
+    apply(drule_tac x="bders_simp (intern r) s" in meta_spec)
+    apply(drule_tac  meta_mp)
+     apply(simp add: ders_append)
+     defer
+    apply(simp)
+    oops
+
+function (sequential) bretrieve :: "arexp \<Rightarrow> bit list \<Rightarrow> (bit list) * (bit list)" where
+  "bretrieve (AZERO) bs1 = ([], bs1)"
+| "bretrieve (AONE bs) bs1 = (bs, bs1)"
+| "bretrieve (ACHAR bs c) bs1 = (bs, bs1)"
+| "bretrieve (AALTs bs rs) [] = (bs, [])"
+| "bretrieve (AALTs bs [r]) bs1 = 
+     (let (bs2, bs3) = bretrieve r bs1 in (bs @ bs2, bs3))"
+| "bretrieve (AALTs bs (r#rs)) (Z#bs1) = 
+     (let (bs2, bs3) = bretrieve r bs1 in (bs @ [Z] @ bs2, bs3))"
+| "bretrieve (AALTs bs (r#rs)) (S#bs1) = 
+     (let (bs2, bs3) = bretrieve (AALTs [] rs) bs1 in (bs @ [S] @ bs2, bs3))"
+| "bretrieve (ASEQ bs r1 r2) bs1 = 
+     (let (bs2, bs3) = bretrieve r1 bs1 in 
+      let (bs4, bs5) = bretrieve r2 bs3 in (bs @ bs2 @ bs4, bs5))"
+| "bretrieve (ASTAR bs r) [] = (bs, [])"
+| "bretrieve (ASTAR bs r) (S#bs1) = (bs @ [S], bs1)"
+| "bretrieve (ASTAR bs r) (Z#bs1) = 
+     (let (bs2, bs3) = bretrieve r bs1 in 
+      let (bs4, bs5) = bretrieve (ASTAR [] r) bs3 in (bs @ bs2 @ [Z] @ bs4, bs5))"
+  by (pat_completeness) (auto)
+
+termination
+  sorry
+
+thm Prf.intros
+
+
+lemma retrieve_XXX:
+  assumes "\<Turnstile> v : erase r" 
+  shows "\<exists>v'. \<Turnstile> v' : erase (bsimp r)  \<and> retrieve (bsimp r) v' = retrieve r v"
+  using assms
+  apply(induct r arbitrary: v)
+       apply(simp)
+  using Prf_elims(1) apply blast
+      apply(simp)
+  using Prf_elims(4) apply fastforce
+    apply(simp)
+  apply blast
+  apply simp
+   apply(case_tac  "r1 = AZERO")
+     apply(simp)
+  apply (meson Prf_elims(1) Prf_elims(2))
+  apply(case_tac  "r2 = AZERO")
+     apply(simp)
+     apply (meson Prf_elims(1) Prf_elims(2))
+  apply(erule Prf_elims)
+  apply(simp)
+   apply(case_tac "\<exists>bs. bsimp r1 = AONE bs")
+     apply(clarify)
+     apply(simp)
+    apply(subst bsimp_ASEQ2)
+     defer
+     apply(subst bsimp_ASEQ1)
+  using L_bsimp_erase L_flat_Prf1 L_flat_Prf2 apply fastforce
+  using L_bsimp_erase L_flat_Prf1 L_flat_Prf2 apply fastforce
+      apply(simp)
+    apply(simp)
+  apply(drule_tac  x="v1" in meta_spec)
+  apply(drule_tac  x="v2" in meta_spec)
+     apply(simp)
+  apply(clarify)
+     apply(rule_tac x="Seq v' v'a" in exI)
+     apply(simp)
+  apply (metis Prf.intros(1) Prf_elims(1) bsimp_ASEQ1 erase.simps(1) retrieve.simps(6))
+    prefer 3
+  apply(drule_tac  x="v1" in meta_spec)
+  apply(drule_tac  x="v2" in meta_spec)
+     apply(simp)
+    apply(clarify)
+    apply(rule_tac x="v'a" in exI)
+    apply(subst bsimp_ASEQ2)
+    apply (metis Prf_elims(4) append_assoc erase_fuse retrieve.simps(1) retrieve_fuse2)
+   prefer 2  
+   apply(auto)
+  apply(case_tac "x2a")
+   apply(simp)
+  using Prf_elims(1) apply blast
+  apply(simp)
+  apply(case_tac "list")
+   apply(simp)
+  sorry
+
+
+lemma TEST:
+  assumes "\<Turnstile> v : ders s r"
+  shows "retrieve (bders (intern r) s) v = retrieve (bsimp (bders (intern r) s)) v"
+  using assms
+  apply(induct s arbitrary: r v rule: rev_induct)
+   apply(simp)
+   defer
+   apply(simp add: ders_append)
+   apply(frule Prf_injval)
+  apply(drule_tac x="r" in meta_spec)
+   apply(drule_tac x="injval (ders xs r) x v" in meta_spec)
+   apply(simp)
+   apply(simp add: bders_append)
+   apply(subst bder_retrieve)
+    apply(simp)
+   apply(simp)
+  thm bder_retrieve
+  thm bmkeps_retrieve
+
+
+lemma bmkeps_simp2:
+  assumes "bnullable (bder c r)"
+  shows "bmkeps (bder c (bsimp r)) = bmkeps (bder c r)"
+  using  assms
+  apply(induct r)
+       apply(simp)
+      apply(simp)
+     apply(simp)
+  prefer 3
+    apply(simp)
+   apply(simp)
+   apply(auto)[1]
+     prefer 2
+     apply(case_tac "r1 = AZERO")
+      apply(simp)
+   apply(case_tac "r2 = AZERO")
+      apply(simp)
+     apply(case_tac "\<exists>bs. (bsimp r1) = AONE bs")
+      apply(clarify)
+      apply(simp)
+      apply(subst bsimp_ASEQ2)
+  
+   apply(simp add: bmkeps_simp)
+  apply(simp add: bders_append)
+  apply(drule_tac x="bder a r" in meta_spec)
+       apply(simp)
+      apply(simp)
+     apply(simp)
+    prefer 3
+    apply(simp)
+   prefer 2
+   apply(simp)
+   apply(case_tac x2a)
+    apply(simp)
+  apply(simp add: )
+  apply(subst k0)
+   apply(auto)[1]
+    apply(case_tac list)
+     apply(simp)
+  
+
+   apply(case_tac  "r1=AZERO")
+    apply(simp)
+   apply(case_tac  "r2=AZERO")
+    apply(simp)
+    apply(auto)[1]
+   apply(case_tac  "\<exists>bs. r1=AONE bs")
+  apply(simp)
+    apply(auto)[1]
+     apply(subst bsimp_ASEQ2)
+
+  
+   prefer 2
+   apply(simp)
+  apply(subst  bmkeps_bder_AALTs)
+   apply(case_tac x2a)
+    apply(simp)
+   apply(simp)
+   apply(auto)[1]
+    apply(subst  bmkeps_bder_AALTs)
+  
+   apply(case_tac a)
+  apply(simp_all)
+   apply(auto)[1]
+    apply(case_tac list)
+         apply(simp)
+        apply(simp)
+  
+     prefer 2
+  apply(simp)
+
+
+lemma bbs0:
+  shows "blexer_simp r [] = blexer r []"
+  apply(simp add: blexer_def blexer_simp_def)
+  done
+
+lemma bbs1:
+  shows "blexer_simp r [c] = blexer r [c]"
+  apply(simp add: blexer_def blexer_simp_def)
+  apply(auto)
+    defer
+  using b3 apply auto[1]
+  using b3 apply auto[1]  
+  apply(subst bmkeps_simp[symmetric])
+   apply(simp)
+  apply(simp)
+  done
+  
+lemma bbs1:
+  shows "blexer_simp r [c1, c2] = blexer r [c1, c2]"
+  apply(simp add: blexer_def blexer_simp_def)
+  apply(auto)
+    defer
+  apply (metis L_bsimp_erase bnullable_correctness der_correctness erase_bder lexer.simps(1) lexer_correct_None option.distinct(1))
+  apply (metis L_bsimp_erase bnullable_correctness der_correctness erase_bder lexer.simps(1) lexer_correct_None option.distinct(1))
+  apply(subst bmkeps_simp[symmetric])
+  using b3 apply auto[1]
+  apply(subst bmkeps_retrieve)
+  using b3 bnullable_correctness apply blast
+  apply(subst bder_retrieve)
+  using b3 bnullable_correctness mkeps_nullable apply fastforce
+  apply(subst bmkeps_retrieve)
+  using bnullable_correctness apply blast
+  apply(subst bder_retrieve)
+  using bnullable_correctness mkeps_nullable apply force
+  
+  using bder_retrieve bmkeps_simp bmkeps_retrieve
+
+  
+
+lemma bsimp_retrieve_bder:
+  assumes "\<Turnstile> v : der c (erase r)"
+  shows "retrieve (bder c r) v = retrieve (bsimp (bder c r)) v"
+  thm bder_retrieve bmkeps_simp
+  apply(induct  r arbitrary: c v)
+       apply(simp)
+      apply(simp)
+     apply(simp)
+    apply(auto)[1]
+     apply(case_tac "bsimp (bder c r1) = AZERO")
+      apply(simp)
+  
+    prefer 3
+    apply(simp)
+  apply(auto elim!: Prf_elims)[1]
+    apply(case_tac "(bsimp (fuse [Z] (bder c r))) = AZERO")
+     apply(simp)
+     apply (metis L_bsimp_erase L_flat_Prf1 L_flat_Prf2 Prf_elims(1) erase.simps(1) erase_bder erase_fuse)
+    apply(case_tac "\<exists>bs. bsimp (fuse [Z] (bder c r)) = AONE bs")
+     apply(clarify)
+     apply(subgoal_tac "L (der c (erase r)) = {[]}")
+      prefer 2
+      apply (metis L.simps(2) L_bsimp_erase erase.simps(2) erase_bder erase_fuse)
+     apply(simp)
+  
+  
+  
+    apply(subst bsimp_ASEQ1)
+       apply(simp)
+      apply(simp)
+  apply(auto)[1]
+  
+     prefer 2
+  
+
+lemma oo:
+  shows "(case (blexer (der c r) s) of None \<Rightarrow> None | Some v \<Rightarrow> Some (injval r c v)) = blexer r (c # s)"
+  apply(simp add: blexer_correctness)
+  done
+
+lemma oo2a:
+  assumes "\<forall>r. bmkeps (bders_simp r s) = bmkeps (bders r s)" "c # s \<in> L r"
+          "bnullable (bders_simp (bsimp (bder c (intern r))) s)"
+  shows "(case (blexer_simp (der c r) s) of None \<Rightarrow> None | Some v \<Rightarrow> Some (injval r c v)) = blexer_simp r (c # s)"
+  using assms
+  apply(simp add: blexer_simp_def)
+  apply(auto  split: option.split)
+    apply (metis blexer_correctness blexer_def lexer.simps(2) lexer_correct_None option.simps(4))
+   prefer 2
+  apply (metis L_bders_simp L_bsimp_erase Posix1(1) Posix_mkeps bnullable_correctness ders_correctness erase_bder erase_bders erase_intern lexer.simps(1) lexer_correct_None)
+  apply(subst bmkeps_retrieve)
+  using L_bders_simp bnullable_correctness nullable_correctness apply blast
+  
+  thm bder_retrieve
+  
+  
+  apply(subst bder_retrieve[symmetric])
+  
+  
+
+  apply(drule_tac x="bsimp (bder c (intern r))" in  spec)
+  apply(drule sym)
+   apply(simp)
+   apply(subst blexer_simp_def)
+  apply(case_tac "bnullable (bders_simp (intern (der c r)) s)")
+   apply(simp)
+  apply(auto split: option.split)
+  apply(subst oo)
+  apply(simp)
+  done
+
+
+
+lemma oo3:
+  assumes "\<forall>r. bders r s = bders_simp r s"
+  shows "blexer_simp r (s @ [c]) = blexer r (s @ [c])"
+  using assms
+  apply(simp (no_asm) add: blexer_simp_def)
+  apply(auto)
+   prefer 2
+  apply (metis L_bders_simp blexer_def bnullable_correctness lexer.simps(1) lexer_correct_None option.distinct(1))
+  apply(simp add: bders_simp_append)
+  apply(subst bmkeps_simp[symmetric])
+  using b3 apply auto[1]
+  apply(simp add: blexer_def)
+  apply(auto)[1]
+   prefer 2
+  apply (metis (mono_tags, lifting) L_bders_simp Posix_mkeps append.right_neutral bders_simp.simps(1) bders_simp.simps(2) bders_simp_append bnullable_correctness lexer.simps(1) lexer_correct_None lexer_correctness(1) option.distinct(1))
+  apply(simp add: bders_append)     
+  done
+
+lemma oo4:
+  assumes "\<forall>r. bmkeps (bders r s) = bmkeps (bders_simp r s)" "bnullable (bder c (bders r s))"
+  shows "bmkeps (bders_simp r (s @ [c])) = bmkeps (bders r (s @ [c]))"
+  using assms
+  apply(simp add: bders_simp_append)
+  apply(subst bmkeps_simp[symmetric])
+  apply (metis L_bders_simp bnullable_correctness der_correctness erase_bder lexer.simps(1) lexer_correct_None option.distinct(1))
+  apply(simp add: bders_append)
+  done
+
+lemma oo4:
+  shows "blexer_simp r s = blexer r s"
+  apply(induct s arbitrary: r rule: rev_induct)
+   apply(simp only: blexer_simp_def blexer_def)
+   apply(simp)
+  apply(simp only: blexer_simp_def blexer_def)
+  apply(subgoal_tac "bnullable (bders_simp (intern r) (xs @ [x])) = bnullable (bders (intern r) (xs @ [x]))")
+   prefer 2
+   apply (simp add: b4)
+  apply(simp)
+  apply(rule impI)
+  apply(simp add: bders_simp_append)
+  apply(subst bmkeps_simp[symmetric])
+  using b3 apply auto[1]
+  apply(subst bmkeps_retrieve)
+  using b3 bnullable_correctness apply blast
+  apply(subst bder_retrieve)
+  using b3 bnullable_correctness mkeps_nullable apply fastforce
+  apply(simp add: bders_append)
+  apply(subst bmkeps_retrieve)
+  using bnullable_correctness apply blast
+  apply(subst bder_retrieve)
+  using bnullable_correctness mkeps_nullable apply fastforce
+  apply(subst erase_bder)
+  apply(case_tac "xs \<in> L")
+  apply(subst (asm) (2) bmkeps_retrieve)
+  
+  
+  thm bmkeps_retrieve bmkeps_retrieve
+  apply(subst bmkeps_retrieve[symmetric])
+  
+   apply (simp add: bnullable_correctness)
+  apply(simp add: bders_simp_append)
+   
+  
+  apply(induct s arbitrary: r rule: rev_induct) 
+   apply(simp add: blexer_def blexer_simp_def)
+  apply(rule oo3)
+  apply(simp (no_asm) add: blexer_simp_def)
+  apply(auto)
+   prefer 2
+  apply (metis L_bders_simp blexer_def bnullable_correctness lexer.simps(1) lexer_correct_None option.distinct(1))
+  apply(simp add: bders_simp_append)
+  apply(subst bmkeps_simp[symmetric])
+  using b3 apply auto[1]
+  apply(simp add: blexer_def)
+  apply(auto)[1]
+   prefer 2
+  apply (m etis (mono_tags, lifting) L_bders_simp Posix_mkeps append.right_neutral bders_simp.simps(1) bders_simp.simps(2) bders_simp_append bnullable_correctness lexer.simps(1) lexer_correct_None lexer_correctness(1) option.distinct(1))
+  apply(simp add: bders_append)     
+  oops
+
+
+lemma bnullable_simp:
+  assumes "s \<in> L (erase r)"
+  shows "bmkeps (bders r s) = bmkeps (bders_simp r s)"
+  using assms
+  apply(induct s arbitrary: r rule: rev_induct)
+   apply(simp)
+  apply(simp add: bders_append bders_simp_append)
+  apply(subst bmkeps_simp[symmetric])
+  apply (metis L_bders_simp b3 bders_simp.simps(1) bders_simp.simps(2) bders_simp_append blexer_correctness blexer_def bnullable_correctness erase_bders erase_intern lexer.simps(1) lexer_correct_None lexer_correct_Some lexer_correctness(1))
+  apply(subst bmkeps_retrieve)
+  apply (metis bders.simps(1) bders.simps(2) bders_append blexer_correctness blexer_def bnullable_correctness erase_bders erase_intern lexer_correct_Some option.distinct(1))
+  apply(subst bmkeps_retrieve)
+  apply (metis L_bders_simp L_bsimp_erase Posix_mkeps bders_simp.simps(1) bders_simp.simps(2) bders_simp_append blexer_correctness blexer_def bnullable_correctness erase_bders erase_intern lexer.simps(1) lexer_correct_None lexer_correctness(2))
+  apply(subst bder_retrieve)
+  apply (metis bders.simps(1) bders.simps(2) bders_append blexer_correctness blexer_def bnullable_correctness erase_bder erase_bders erase_intern lexer_correct_Some mkeps_nullable option.distinct(1))
+  apply(subst bder_retrieve)
+  apply (metis L_bders_simp L_bsimp_erase Posix_mkeps bders_simp.simps(1) bders_simp.simps(2) bders_simp_append blexer_correctness blexer_def bnullable_correctness erase_bder erase_bders erase_intern lexer.simps(1) lexer_correct_None lexer_correctness(2) mkeps_nullable)
+    
+  apply(drule_tac x="bder a r" in meta_spec)
+  apply(drule_tac  meta_mp)
+  apply (me tis erase_bder lexer.simps(2) lexer_correct_None option.simps(4))
+  apply(simp)
+  oops
+
+
+lemma
+  shows "blexer r s = blexer_simp r s"
+  apply(induct s arbitrary: r rule: rev_induct)
+   apply(simp add: blexer_def blexer_simp_def)
+  apply(case_tac "xs @ [x] \<in> L r")
+   defer
+   apply(subgoal_tac "blexer (ders xs r) [x] = None")
+    prefer 2
+    apply(subst blexer_correctness)
+    apply(simp (no_asm) add: lexer_correct_None)
+    apply(simp add: nullable_correctness)
+    apply(simp add: der_correctness ders_correctness)
+  apply(simp add: Der_def Ders_def)
+apply(subgoal_tac "blexer r (xs @ [x]) = None")
+    prefer 2
+    apply (simp add: blexer_correctness)
+  using lexer_correct_None apply auto[1]
+  apply(simp)
+   apply(subgoal_tac "blexer_simp (ders xs r) [x] = None")
+    prefer 2
+  apply (metis L_bders_simp Posix_injval Posix_mkeps bders.simps(2) blexer_correctness blexer_simp_def bnullable_correctness ders.simps(1) erase_bder erase_bders erase_intern lexer_correct_None lexer_correctness(2))
+   apply(subgoal_tac "[] \<notin> L (erase (bders_simp (intern r) (xs @ [x])))")
+  prefer 2
+  apply(metis L_bders_simp Posix_injval bders.simps(2) blexer_correctness ders.simps(1) ders_append erase_bder erase_bders erase_intern lexer_correct_None lexer_correctness(2))
+  using blexer_simp_def bnullable_correctness lexer_correct_None apply auto[1]
+(* case xs @ [x] \<in> L r *)
+  apply(subgoal_tac "\<exists>v. blexer (ders xs r) [x] = Some v \<and> [x] \<in> (ders xs r) \<rightarrow> v")
+   prefer 2  
+  using blexer_correctness lexer_correct_Some apply auto[1]
+    apply (simp add: Posix_injval Posix_mkeps)
+  apply (metis ders.simps(1) ders.simps(2) ders_append lexer_correct_None lexer_flex)    
+  apply(clarify)
+  apply(subgoal_tac "blexer_simp (ders xs r) [x] = Some v")
+   prefer 2
+   apply(simp add: blexer_simp_def)
+   apply(auto)[1]
+    apply (metis bders.simps(1) bders.simps(2) blexer_def bmkeps_simp option.simps(3))
+  using b3 blexer_def apply fastforce
+  apply(subgoal_tac "blexer_simp (ders xs r) [x] = blexer_simp r (xs @ [x])")
+   prefer 2
+   apply(simp add: blexer_simp_def)
+  
+   apply(simp)
+  
+  
+  
+   apply(simp)
+  apply(subst blexer_simp_def)
+  apply(simp)
+  apply(auto)
+  apply(drule_tac x="der a r" in meta_spec)
+  apply(subst blexer_def)
+   apply(subgoal_tac "bnullable (bders (intern r) (a # s))")
+    prefer 2
+    apply (metis Posix_injval blexer_correctness blexer_def lexer_correctness(2))
+  apply(simp)
+  
+
+
+lemma
+  shows "blexer r s = blexer_simp r s"
+  apply(induct s arbitrary: r)
+   apply(simp add: blexer_def blexer_simp_def)
+  apply(case_tac "s \<in> L (der a r)")
+   defer
+   apply(subgoal_tac "blexer (der a r) s = None")
+    prefer 2
+    apply (simp add: blexer_correctness lexer_correct_None)
+apply(subgoal_tac "blexer r (a # s) = None")
+    prefer 2
+    apply (simp add: blexer_correctness)
+   apply(simp)
+  
+   apply(subst blexer_simp_def)
+   apply(simp)
+  apply(drule_tac  x="der a r" in  meta_spec)
+   apply(subgoal_tac "s \<notin> L(erase (bder a (intern  r)))")
+  prefer 2
+    apply simp
+   
+   apply(simp only:)
+   apply(subst blexer_simp_def)
+   apply(subgoal_tac "\<not> bnullable (bders_simp (intern r) (a # s))")
+    apply(simp)
+   apply(subst bnullable_correctness[symmetric])
+  apply(simp)
+  
+
 
 end
\ No newline at end of file
--- a/thys/RegLangs.thy	Sat Feb 23 21:52:06 2019 +0000
+++ b/thys/RegLangs.thy	Wed Mar 13 10:36:29 2019 +0000
@@ -193,6 +193,8 @@
   shows "ders (s1 @ s2) r = ders s2 (ders s1 r)"
   by (induct s1 arbitrary: s2 r) (auto)
 
-
+lemma ders_snoc:
+  shows "ders (s @ [c]) r = der c (ders s r)"
+  by (simp add: ders_append)
 
 end
\ No newline at end of file