thys2/Turing2.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Fri, 24 May 2013 20:35:28 +0100
changeset 261 ca1fe315cb0a
parent 259 4524c5edcafb
child 288 a9003e6d0463
permissions -rwxr-xr-x
completed the UF-simulation lemmas
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
257
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     1
(* Title: thys/Turing.thy
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     2
   Author: Jian Xu, Xingyuan Zhang, and Christian Urban
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     3
*)
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     4
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     5
header {* Turing Machines *}
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     6
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     7
theory Turing2
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     8
imports Main
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     9
begin
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    10
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    11
section {* Basic definitions of Turing machine *}
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    12
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    13
datatype action = W0 | W1 | L | R | Nop
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    14
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    15
datatype cell = Bk | Oc
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    16
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    17
type_synonym tape = "cell list \<times> cell list"
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    18
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    19
type_synonym state = nat
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    20
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    21
type_synonym instr = "action \<times> state"
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    22
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    23
type_synonym tprog = "(instr \<times> instr) list"
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    24
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    25
type_synonym config = "state \<times> tape"
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    26
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    27
fun nth_of where
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    28
  "nth_of xs i = (if i \<ge> length xs then None else Some (xs ! i))"
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    29
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    30
fun 
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    31
  fetch :: "tprog \<Rightarrow> state \<Rightarrow> cell \<Rightarrow> instr"
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    32
where
259
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    33
  "fetch tm 0 b = (Nop, 0)"
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    34
| "fetch tm (Suc s) b = 
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    35
     (case nth_of tm s of
257
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    36
        Some i \<Rightarrow> (case b of Bk \<Rightarrow> fst i | Oc \<Rightarrow> snd i)
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    37
      | None \<Rightarrow> (Nop, 0))"
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    38
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    39
fun 
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    40
  update :: "action \<Rightarrow> tape \<Rightarrow> tape"
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    41
where 
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    42
  "update W0 (l, r) = (l, Bk # (tl r))" 
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    43
| "update W1 (l, r) = (l, Oc # (tl r))"
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    44
| "update L  (l, r) = (if l = [] then ([], Bk # r) else (tl l, (hd l) # r))" 
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    45
| "update R  (l, r) = (if r = [] then (Bk # l, []) else ((hd r) # l, tl r))" 
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    46
| "update Nop (l, r) = (l, r)"
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    47
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    48
abbreviation 
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    49
  "read r == if (r = []) then Bk else hd r"
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    50
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    51
fun step :: "config \<Rightarrow> tprog \<Rightarrow> config"
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    52
  where 
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    53
  "step (s, l, r) p = 
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    54
     (let (a, s') = fetch p s (read r) in (s', update a (l, r)))"
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    55
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    56
fun steps :: "config \<Rightarrow> tprog \<Rightarrow> nat \<Rightarrow> config"
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    57
  where
259
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    58
  "steps cf p 0 = cf" |
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    59
  "steps cf p (Suc n) = steps (step cf p) p n"
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    60
257
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    61
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    62
(* well-formedness of Turing machine programs *)
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    63
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    64
fun 
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    65
  tm_wf :: "tprog \<Rightarrow> bool"
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    66
where
259
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    67
  "tm_wf p = (1 \<le> length p \<and> (\<forall>((_, s1), (_, s2)) \<in> set p. s1 \<le> length p \<and> s2 \<le> length p))"
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    68
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    69
(* short-hand notation for tapes *)
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    70
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    71
abbreviation cell_replicate :: "'a \<Rightarrow> nat \<Rightarrow> 'a list" ("_ \<up> _" [100, 99] 100)
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    72
  where "x \<up> n \<equiv> replicate n x"
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    73
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    74
class tape_of =
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    75
  fixes tape_of :: "'a \<Rightarrow> cell list" ("<_>" [64] 67)
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    76
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    77
instantiation nat :: tape_of 
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    78
begin
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    79
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    80
fun tape_of_nat where 
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    81
  "<(n::nat)> = Oc \<up> (Suc n)"
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    82
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    83
instance ..
257
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    84
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    85
end
df6e8cb22995 more about the UF
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    86
259
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    87
instantiation list :: (tape_of) tape_of 
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    88
begin
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    89
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    90
fun tape_of_list :: "'a list \<Rightarrow> cell list" 
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    91
  where 
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    92
  "<[]> = []" |
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    93
  "<[n]> = <n>" |
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    94
  "<n # ns> = <n> @ [Bk] @ <ns>"
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    95
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    96
instance ..
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    97
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    98
end
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
    99
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
   100
instantiation prod :: (tape_of, tape_of) tape_of 
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
   101
begin
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
   102
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
   103
fun tape_of_prod :: "'a \<times> 'b \<Rightarrow> cell list" 
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
   104
  where 
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
   105
  "<(n, m)> = <n> @ [Bk] @ <m>" 
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
   106
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
   107
instance ..
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
   108
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
   109
end
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
   110
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
   111
definition 
261
ca1fe315cb0a completed the UF-simulation lemmas
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 259
diff changeset
   112
  "std_tape tp \<equiv> \<exists>k (n::nat) l. snd tp = (Bk \<up> k, <n> @ Bk \<up> l)"
ca1fe315cb0a completed the UF-simulation lemmas
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 259
diff changeset
   113
ca1fe315cb0a completed the UF-simulation lemmas
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 259
diff changeset
   114
fun
ca1fe315cb0a completed the UF-simulation lemmas
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 259
diff changeset
   115
  is_final :: "config \<Rightarrow> bool"
ca1fe315cb0a completed the UF-simulation lemmas
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 259
diff changeset
   116
where
ca1fe315cb0a completed the UF-simulation lemmas
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 259
diff changeset
   117
  "is_final cf = (fst cf = 0)"
ca1fe315cb0a completed the UF-simulation lemmas
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 259
diff changeset
   118
259
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
   119
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
   120
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
   121
end
4524c5edcafb moved new theries into a separate directory
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 257
diff changeset
   122