progs/dfa.scala
author Christian Urban <urbanc@in.tum.de>
Fri, 17 Mar 2017 12:15:58 +0000
changeset 479 52aa298211f6
parent 145 920f675b4ed1
child 481 acd8780bfc8b
permissions -rw-r--r--
updated

// convert normal string to hex bytes string
def string2hex(str: String): String = {
  str.toList.map(_.toInt.toHexString).mkString
}
    
// convert hex bytes string to normal string


val byteArray: Array[Byte] = List.range(1, 255).map(_.toByte).toArray
val bos = new BufferedOutputStream(new FileOutputStream("test.txt"))
bos.write(byteArray)
bos.close() 


def string2int(hex: String) = {
  hex.sliding(2, 2).map(Integer.parseInt(_, 16)).toArray
}

def test(l: List[Int]) = {
  l.map(_.toChar).mkString
}


import java.io._
val writer = new PrintWriter(new File("test.txt" ))

//writer.write(string2int("cafebabe"))
writer.write(test(List.range(1, 255)))
writer.close()

import scalax.io._
import scalax.io.Resource

FileOutputStream fos = new FileOutputStream("test");
fos.write(bytesArray);
fos.close();


string2int("cafebabe")

202.toHexString

"cafebabe".sliding(1, 1).toList.map(Integer.parseInt(_, 16)).map(_.toByte).map(_.toChar)

hex2string("ca")
string2hex("ca")
hex2string("cafebabe")

    
val appkey = "9GLV//lv/kYFW2o3/bihxwnMcmo="
        
        // string to hex
        val appkey_hex = string2hex(appkey)
        // 39474c562f2f6c762f6b594657326f332f62696878776e4d636d6f3d
        println(appkey_hex)
        
        // hex to string
        val appkey_string_again = hex2string(appkey_hex)
        // 9GLV//lv/kYFW2o3/bihxwnMcmo=
        println(appkey_string_again)
    }


List("ca", "fe", "ba", "be").map(_.toByte)

"ca".toByte


// DFAs
import scala.util._

abstract class State
type States = Set[State]

case class IntState(i: Int) extends State

object NewState {
  var counter = 0
  
  def apply() : IntState = {
    counter += 1;
    new IntState(counter - 1)
  }
}


// DFA class
case class DFA(states: States, 
               start: State, 
               delta: (State, Char) => State, 
               fins: States) {
  
  def deltas(q: State, s: List[Char]) : State = s match {
    case Nil => q
    case c::cs => deltas(delta(q, c), cs) 
  }
  
  // wether a string is accepted by the automaton or not
  def accepts(s: String) : Boolean = 
    Try(fins contains 
      (deltas(start, s.toList))) getOrElse false
}


// example DFA from the lectures
val Q0 = NewState()
val Q1 = NewState()
val Q2 = NewState()
val Q3 = NewState()
val Q4 = NewState()


val delta : (State, Char) => State = {
  case (Q0, 'a') => Q1
  case (Q0, 'b') => Q2
  case (Q1, 'a') => Q4
  case (Q1, 'b') => Q2
  case (Q2, 'a') => Q3
  case (Q2, 'b') => Q2
  case (Q3, 'a') => Q4
  case (Q3, 'b') => Q0
  case (Q4, 'a') => Q4
  case (Q4, 'b') => Q4
}

val DFA1 = DFA(Set(Q0, Q1, Q2, Q3, Q4), Q0, delta, Set(Q4))

println(DFA1.accepts("aaa"))
println(DFA1.accepts("bb"))
println(DFA1.accepts("aaac"))