diff -r 04700724250f -r df6e8cb22995 thys/Turing2.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys/Turing2.thy Tue May 21 13:49:31 2013 +0100 @@ -0,0 +1,69 @@ +(* Title: thys/Turing.thy + Author: Jian Xu, Xingyuan Zhang, and Christian Urban +*) + +header {* Turing Machines *} + +theory Turing2 +imports Main +begin + +section {* Basic definitions of Turing machine *} + +datatype action = W0 | W1 | L | R | Nop + +datatype cell = Bk | Oc + +type_synonym tape = "cell list \ cell list" + +type_synonym state = nat + +type_synonym instr = "action \ state" + +type_synonym tprog = "(instr \ instr) list" + +type_synonym config = "state \ tape" + +fun nth_of where + "nth_of xs i = (if i \ length xs then None else Some (xs ! i))" + +fun + fetch :: "tprog \ state \ cell \ instr" +where + "fetch p 0 b = (Nop, 0)" +| "fetch p (Suc s) b = + (case nth_of p s of + Some i \ (case b of Bk \ fst i | Oc \ snd i) + | None \ (Nop, 0))" + +fun + update :: "action \ tape \ tape" +where + "update W0 (l, r) = (l, Bk # (tl r))" +| "update W1 (l, r) = (l, Oc # (tl r))" +| "update L (l, r) = (if l = [] then ([], Bk # r) else (tl l, (hd l) # r))" +| "update R (l, r) = (if r = [] then (Bk # l, []) else ((hd r) # l, tl r))" +| "update Nop (l, r) = (l, r)" + +abbreviation + "read r == if (r = []) then Bk else hd r" + +fun step :: "config \ tprog \ config" + where + "step (s, l, r) p = + (let (a, s') = fetch p s (read r) in (s', update a (l, r)))" + +fun steps :: "config \ tprog \ nat \ config" + where + "steps c p 0 = c" | + "steps c p (Suc n) = steps (step c p) p n" + +(* well-formedness of Turing machine programs *) + +fun + tm_wf :: "tprog \ bool" +where + "tm_wf p = (length p \ 1 \ (\((_, s1), (_, s2)) \ set p. s1 \ length p \ s2 \ length p))" + +end +