thys/Turing2.thy
changeset 257 df6e8cb22995
equal deleted inserted replaced
256:04700724250f 257:df6e8cb22995
       
     1 (* Title: thys/Turing.thy
       
     2    Author: Jian Xu, Xingyuan Zhang, and Christian Urban
       
     3 *)
       
     4 
       
     5 header {* Turing Machines *}
       
     6 
       
     7 theory Turing2
       
     8 imports Main
       
     9 begin
       
    10 
       
    11 section {* Basic definitions of Turing machine *}
       
    12 
       
    13 datatype action = W0 | W1 | L | R | Nop
       
    14 
       
    15 datatype cell = Bk | Oc
       
    16 
       
    17 type_synonym tape = "cell list \<times> cell list"
       
    18 
       
    19 type_synonym state = nat
       
    20 
       
    21 type_synonym instr = "action \<times> state"
       
    22 
       
    23 type_synonym tprog = "(instr \<times> instr) list"
       
    24 
       
    25 type_synonym config = "state \<times> tape"
       
    26 
       
    27 fun nth_of where
       
    28   "nth_of xs i = (if i \<ge> length xs then None else Some (xs ! i))"
       
    29 
       
    30 fun 
       
    31   fetch :: "tprog \<Rightarrow> state \<Rightarrow> cell \<Rightarrow> instr"
       
    32 where
       
    33   "fetch p 0 b = (Nop, 0)"
       
    34 | "fetch p (Suc s) b = 
       
    35      (case nth_of p s of
       
    36         Some i \<Rightarrow> (case b of Bk \<Rightarrow> fst i | Oc \<Rightarrow> snd i)
       
    37       | None \<Rightarrow> (Nop, 0))"
       
    38 
       
    39 fun 
       
    40   update :: "action \<Rightarrow> tape \<Rightarrow> tape"
       
    41 where 
       
    42   "update W0 (l, r) = (l, Bk # (tl r))" 
       
    43 | "update W1 (l, r) = (l, Oc # (tl r))"
       
    44 | "update L  (l, r) = (if l = [] then ([], Bk # r) else (tl l, (hd l) # r))" 
       
    45 | "update R  (l, r) = (if r = [] then (Bk # l, []) else ((hd r) # l, tl r))" 
       
    46 | "update Nop (l, r) = (l, r)"
       
    47 
       
    48 abbreviation 
       
    49   "read r == if (r = []) then Bk else hd r"
       
    50 
       
    51 fun step :: "config \<Rightarrow> tprog \<Rightarrow> config"
       
    52   where 
       
    53   "step (s, l, r) p = 
       
    54      (let (a, s') = fetch p s (read r) in (s', update a (l, r)))"
       
    55 
       
    56 fun steps :: "config \<Rightarrow> tprog \<Rightarrow> nat \<Rightarrow> config"
       
    57   where
       
    58   "steps c p 0 = c" |
       
    59   "steps c p (Suc n) = steps (step c p) p n"
       
    60 
       
    61 (* well-formedness of Turing machine programs *)
       
    62 
       
    63 fun 
       
    64   tm_wf :: "tprog \<Rightarrow> bool"
       
    65 where
       
    66   "tm_wf p = (length p \<ge> 1 \<and> (\<forall>((_, s1), (_, s2)) \<in> set p. s1 \<le> length p \<and> s2 \<le> length p))"
       
    67 
       
    68 end
       
    69