turing.scala
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Mon, 18 Feb 2013 13:33:00 +0000
changeset 182 0329bc0bceec
parent 180 8f443f2ed1f6
child 183 4cf023ee2f4c
permissions -rw-r--r--
tuned


//some list functions
def tl[A](xs: List[A]) : List[A] = xs match {
  case Nil => Nil
  case x::xs => xs
}

def hd[A](xs: List[A]) : A = xs.head

def nth_of[A](xs: List[A], n: Int) = 
  if (n <= xs.length) Some(xs(n)) else None 

def replicate[A](x: A, n: Int) : List[A] = n match {
  case 0 => List(x)
  case n => x::replicate(x, n - 1)
}


//some map functions
def dget(m: Map[Int, Int], n: Int) = m.getOrElse(n, 0)

//some string functions
def join[A](xs: List[A]) = xs.foldLeft("")(_.toString + _)



abstract class Cell
case object Bk extends Cell { override def toString = "0" }
case object Oc extends Cell { override def toString = "1" }

abstract class Action
case object WBk extends Action
case object WOc extends Action
case object R extends Action
case object L extends Action
case object Nop extends Action

type State = Int
type Prog = List[(Action, State)]

//tapes
case class Tape(l: List[Cell], r: List[Cell]) {
  def update(a: Action) = a match {
    case WBk => Tape(l, Bk::tl(r))
    case WOc => Tape(l, Oc::tl(r))
    case L => if (l == Nil) Tape(Nil, Bk::r) else Tape (tl(l), hd(l)::r)
    case R => if (r == Nil) Tape(Bk::l, Nil) else Tape(hd(r)::l, tl(r))
    case Nop => Tape(l, r)
  }

  def read = r match {
    case Nil => Bk
    case x::_ => x
  }

  override def toString = join(l.reverse) + ">" + join(r)
}

// standard tapes
class STape(ns: Int*) extends 
  Tape(Nil, ns.map(replicate[Cell](Oc, _)).reduceLeft(_ ::: List(Bk) ::: _))
  

// configurations
case class Config(s: State, tp: Tape) {
  def is_final = s == 0

  override def toString = tp.toString
}


// Turing Machines
case class TM(p: Prog) {

  def shift(n: Int) =
    TM(p.map{case (a, s) => (a, if (s == 0) 0 else s + n)})
  
  def fetch(s: State, a: Cell) = (s, a) match {
    case (0, _) => (Nop, 0)
      case (s, Bk) => nth_of(p, 2 * (s - 1)) match {
        case None => (Nop, 0)
        case Some(i) => i
      }
    case (s, Oc) => nth_of(p, 2 * (s - 1) + 1) match {
      case None => (Nop, 0)
      case Some(i) => i
    }
  } 

  def step(cf: Config) : Config = fetch (cf.s, cf.tp.read) match {
    case (a, s_next) => Config(s_next, cf.tp.update(a))
  }

  def steps(cf: Config, n: Int) : Config = n match {
    case 0 => cf
    case n => steps(step(cf), n - 1)
  } 

  def run(cf: Config) : Config = {
    if (cf.is_final) cf else run(step(cf))
  }
 
  def run(tp: Tape) : Tape = run(Config(1, tp)).tp
}



// examples
val TMCopy = TM(List((WBk, 5), (R, 2), (R, 3), (R, 2), (WOc, 3), 
                     (L, 4), (L, 4), (L, 5), (R, 11), (R, 6), 
                     (R, 7), (WBk, 6), (R, 7), (R, 8), (WOc, 9), 
                     (R, 8), (L, 10), (L, 9), (L, 10), (L, 5), 
                     (L, 0), (R, 12), (WOc, 13), (L, 14), (R, 12), 
                     (R, 12), (L, 15), (WBk, 14), (R, 0), (L, 15)))

def TMFindnth(n: Int) : TM = n match {
  case 0 => TM(Nil)
  case n => {
    val tm = TMFindnth(n - 1)
    TM(tm.p ::: List((WOc, 2 * n - 1), (R, 2 * n), (R, 2 * n + 1), (R, 2 * n)))
  }
}

def TMMopup(n: Int) = {
  def TMMopup1(n: Int) : TM = n match {
    case 0 => TM(Nil)
    case n => {
      val tm = TMMopup1(n - 1) 
      TM(tm.p ::: List((R, 2 * n + 1), (WBk, 2 * n), (R, 2 * n - 1), (WOc, 2 * n)))
    }
  }

  val TMMopup2 = TM(List((R, 2), (R, 1), (L, 5), (WBk, 3), (R, 4), (WBk, 3),
                         (R, 2), (WBk, 3), (L, 5), (L, 6), (R, 0), (L, 6)))

  val tm1 = TMMopup1(n)
  val tm2 = TMMopup2.shift(2 * n)
  TM(tm1.p ::: tm2.p)
}

println("TMCopy:    " + (TMCopy.run(new STape(3))))
println("TMfindnth: " + (TMFindnth(3).run(new STape(1,2,3,4,5))))
println("TMMopup: "   + (TMMopup(3).run(new STape(1,2,3,4,5))))


// Abacus
abstract class AInst 
case class Inc(n: Int) extends AInst
case class Dec(n: Int, l: Int) extends AInst
case class Goto(l: Int) extends AInst

def ashift(ai: AInst, offset: Int, jump: Int) = ai match {
  case Inc(n) => Inc(n)
  case Dec(n, l) => if (l == jump) Dec(n, l) else Dec(n, l + offset)
  case Goto(l) => if (l == jump) Goto(l) else Goto(l + offset)
}

def aadjust(ai: AInst, old_jump: Int, jump: Int) = ai match {
  case Inc(n) => Inc(n)
  case Dec(n, l) => if (l == old_jump) Dec(n, jump) else Dec(n, l)
  case Goto(l) => if (l == old_jump) Goto(jump) else Goto(l)
}



type AProg = List[AInst]
type Regs = Map[Int, Int]

case class AConfig(s: Int, regs: Regs)  

case class Abacus(p: AProg) {

  def shift(offset: Int, jump: Int) = Abacus(p.map(ashift(_, offset, jump)))
  def adjust(old_jump: Int, jump: Int) = Abacus(p.map(aadjust(_, old_jump, jump)))
 
  def step(cf: AConfig) : AConfig = (nth_of(p, cf.s), cf.s) match {
    case (None, _) => cf
    case (Some(Inc(n)), s) => AConfig(s + 1, cf.regs + (n -> (dget(cf.regs, n) + 1)))
    case (Some(Dec(n, l)), s) => {
      if (dget(cf.regs, n) == 0) AConfig(l, cf.regs)
      else AConfig(s + 1, cf.regs + (n -> (dget(cf.regs, n) - 1)))
    }
    case (Some(Goto(l)), _) => AConfig(l, cf.regs)
  }  

  def steps(cf: AConfig, n: Int) : AConfig = n match {
    case 0 => cf
    case n => steps(step(cf), n - 1)
  } 

  def steps(regs: Regs, n: Int) : AConfig = steps(AConfig(0, regs), n)

  def run(cf: AConfig) : AConfig = {
    if (cf.s >= p.length || cf.s < 0) cf else run(step(cf))
  }
 
  def run(regs: Regs): Regs = run(AConfig(0, regs)).regs
}

// examples
def Copy(in: Int, out: Int) = 
  Abacus(List(Dec(in, -1), Inc(out), Goto(0))) 

def Plus(m: Int, n: Int, tmp: Int, jump: Int) =
  Abacus(List(Dec(m, 4), Inc(n), Inc(tmp),
              Goto(0), Dec(tmp, jump), Inc(m), Goto(4))) 

def Mult(in1: Int, in2: Int, out: Int, tmp: Int, jump: Int) = {
  val tm = Plus(in2, out, tmp, -1).shift(1, -1).adjust(-1, 0)
  Abacus(Dec(in1, jump) :: tm.p)
}

def Expo(in1: Int, in2: Int, out: Int, tmp: Int, jump: Int) = {
  val tm = Plus(in2, out, tmp, -1).shift(1, -1).adjust(-1, 0)
  Abacus(Dec(in1, jump) :: tm.p)
}


println("Copy: " + (Copy(0, 1).run(Map(0 -> 3, 1 -> 0))))
println("Plus: " + (Plus(0, 1, 2, -1).run(Map(0 -> 3, 1 -> 4, 2 -> 0))))
println("Mult: " + (Mult(0, 1, 2, 3, -1).run(Map(0 -> 3, 1 -> 5, 2 -> 0, 3 -> 0))))



//Recursive Functions
abstract class Rec {
  def eval(ns: List[Int]) : Int
}
case object Z extends Rec {
  override def eval(ns: List[Int]) = ns match {
    case n::Nil => 0
    case _ => throw new IllegalArgumentException("Z: args")
  }
} 
case object S extends Rec {
  override def eval(ns: List[Int]) = ns match {
    case n::Nil => n + 1
    case _ => throw new IllegalArgumentException("S: args")
  }
} 
case class Id(n: Int, m: Int) extends Rec {
  override def eval(ns: List[Int]) = 
    if (ns.length == n && m < n) ns(m)
    else throw new IllegalArgumentException("Id: args")
}
case class Cn(n: Int, f: Rec, gs: List[Rec]) extends Rec {
  override def eval(ns: List[Int]) = 
    if (ns.length == n) f.eval(gs.map(_.eval(ns)))
    else throw new IllegalArgumentException("Cn: args")
}
case class Pr(n: Int, f: Rec, g: Rec) extends Rec {
  override def eval(ns: List[Int]) = 
    if (ns.length == n - 1) {
      if (ns.last == 0) f.eval(ns.init)
      else {
        val r = Pr(n, f, g).eval(ns.init ::: List(ns.last - 1))
        g.eval(ns.init ::: List(ns.last - 1, r))
      }
    }
    else throw new IllegalArgumentException("Cn: args")
}
case class Mn(n: Int, f: Rec) extends Rec {
  override def eval(ns: List[Int]) = 0
    
}