Paper.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Sun, 13 Jan 2013 11:29:33 +0000
changeset 40 a37a21f7ccf4
parent 38 8f8db701f69f
permissions -rw-r--r--
updated test
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     1
(*<*)
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     2
theory Paper
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     3
imports UTM
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     4
begin
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     5
25
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
     6
hide_const (open) s 
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
     7
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
     8
abbreviation 
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
     9
  "update p a == new_tape a p"
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    10
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    11
lemma fetch_def2: 
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
    12
  shows "fetch p 0 b == (Nop, 0)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
    13
  and "fetch p (Suc s) Bk == 
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    14
     (case nth_of p (2 * s) of
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    15
        Some i \<Rightarrow> i
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    16
      | None \<Rightarrow> (Nop, 0))"
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
    17
  and "fetch p (Suc s) Oc == 
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    18
     (case nth_of p ((2 * s) + 1) of
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    19
         Some i \<Rightarrow> i
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    20
       | None \<Rightarrow> (Nop, 0))"
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
    21
apply -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
    22
apply(rule_tac [!] eq_reflection)
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    23
by (auto split: block.splits simp add: fetch.simps)
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    24
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    25
lemma new_tape_def2: 
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    26
  shows "new_tape W0 (l, r) == (l, Bk#(tl r))" 
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    27
  and "new_tape W1 (l, r) == (l, Oc#(tl r))"
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    28
  and "new_tape L (l, r) == 
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    29
     (if l = [] then ([], Bk#r) else (tl l, (hd l) # r))" 
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    30
  and "new_tape R (l, r) ==
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    31
     (if r = [] then (Bk#l,[]) else ((hd r)#l, tl r))" 
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    32
  and "new_tape Nop (l, r) == (l, r)"
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    33
apply -
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    34
apply(rule_tac [!] eq_reflection)
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    35
apply(auto split: taction.splits simp add: new_tape.simps)
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    36
done
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    37
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
    38
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
    39
abbreviation 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
    40
  "read r == if (r = []) then Bk else hd r"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
    41
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    42
lemma tstep_def2:
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
    43
  shows "tstep (s, l, r) p == (let (a, s') = fetch p s (read r) in (s', new_tape a (l, r)))"
23
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 22
diff changeset
    44
apply -
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 22
diff changeset
    45
apply(rule_tac [!] eq_reflection)
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
    46
by (auto split: if_splits prod.split list.split simp add: tstep.simps)
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    47
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
    48
abbreviation
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
    49
 "run p inp out == \<exists>n. steps (1, inp) p n = (0, out)"
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
    50
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
    51
lemma haltP_def2:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
    52
  "haltP p n = (\<exists>k l m. 
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
    53
     run p ([], exponent Oc n) (exponent Bk k, exponent Oc l @ exponent Bk m))"
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
    54
unfolding haltP_def 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
    55
apply(auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
    56
done
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
    57
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
    58
lemma tape_of_nat_list_def2:
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
    59
  shows "tape_of_nat_list [] = []" 
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
    60
  and "tape_of_nat_list [n] = exponent Oc (n+1)"
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
    61
  and "ns\<noteq> [] ==> tape_of_nat_list (n#ns) = (exponent Oc (n+1)) @ [Bk] @ (tape_of_nat_list ns)"
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
    62
apply(auto simp add: tape_of_nat_list.simps)
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
    63
apply(case_tac ns)
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
    64
apply(auto simp add: tape_of_nat_list.simps)
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
    65
done
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
    66
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
    67
lemma tshift_def2:
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
    68
  fixes p::"tprog"
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
    69
  shows "tshift p n == (map (\<lambda> (a, s). (a, (if s = 0 then 0 else s + n))) p)"
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
    70
apply(rule eq_reflection)
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
    71
apply(auto simp add: tshift.simps)
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
    72
done
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
    73
35
839e37b75d9a removed second definition of tshift in abacus.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 34
diff changeset
    74
lemma change_termi_state_def2:
839e37b75d9a removed second definition of tshift in abacus.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 34
diff changeset
    75
 "change_termi_state p  == 
36
4b35e0e0784b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
    76
  (map (\<lambda> (a, s). (a, if s = 0 then ((length p) div 2) + 1 else s)) p)"
35
839e37b75d9a removed second definition of tshift in abacus.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 34
diff changeset
    77
apply(rule eq_reflection)
839e37b75d9a removed second definition of tshift in abacus.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 34
diff changeset
    78
apply(auto simp add: change_termi_state.simps)
839e37b75d9a removed second definition of tshift in abacus.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 34
diff changeset
    79
done
839e37b75d9a removed second definition of tshift in abacus.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 34
diff changeset
    80
839e37b75d9a removed second definition of tshift in abacus.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 34
diff changeset
    81
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
    82
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    83
consts DUMMY::'a
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    84
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    85
notation (latex output)
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    86
  Cons ("_::_" [78,77] 73) and
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
    87
  set ("") and
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    88
  W0 ("W\<^bsub>\<^raw:\hspace{-2pt}>Bk\<^esub>") and
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    89
  W1 ("W\<^bsub>\<^raw:\hspace{-2pt}>Oc\<^esub>") and
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
    90
  t_correct ("twf") and 
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
    91
  tstep ("step") and
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
    92
  steps ("nsteps") and
27
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
    93
  abc_lm_v ("lookup") and
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
    94
  abc_lm_s ("set") and
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
    95
  haltP ("stdhalt") and 
36
4b35e0e0784b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
    96
  tshift ("shift") and 
37
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
    97
  tcopy ("copy") and 
35
839e37b75d9a removed second definition of tshift in abacus.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 34
diff changeset
    98
  change_termi_state ("adjust") and
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
    99
  tape_of_nat_list ("\<ulcorner>_\<urcorner>") and 
36
4b35e0e0784b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   100
  t_add ("_ \<oplus> _") and 
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   101
  DUMMY  ("\<^raw:\mbox{$\_$}>")
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   102
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   103
declare [[show_question_marks = false]]
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   104
(*>*)
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   105
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   106
section {* Introduction *}
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   107
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   108
text {*
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   109
8
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   110
\noindent
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   111
We formalised in earlier work the correctness proofs for two
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   112
algorithms in Isabelle/HOL---one about type-checking in
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   113
LF~\cite{UrbanCheneyBerghofer11} and another about deciding requests
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   114
in access control~\cite{WuZhangUrban12}.  The formalisations
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   115
uncovered a gap in the informal correctness proof of the former and
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   116
made us realise that important details were left out in the informal
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   117
model for the latter. However, in both cases we were unable to
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   118
formalise in Isabelle/HOL computability arguments about the
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   119
algorithms. The reason is that both algorithms are formulated in terms
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   120
of inductive predicates. Suppose @{text "P"} stands for one such
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   121
predicate.  Decidability of @{text P} usually amounts to showing
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   122
whether \mbox{@{term "P \<or> \<not>P"}} holds. But this does \emph{not} work
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   123
in Isabelle/HOL, since it is a theorem prover based on classical logic
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   124
where the law of excluded middle ensures that \mbox{@{term "P \<or> \<not>P"}}
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   125
is always provable no matter whether @{text P} is constructed by
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   126
computable means. The same problem would arise if we had formulated
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   127
the algorithms as recursive functions, because internally in
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   128
Isabelle/HOL, like in all HOL-based theorem provers, functions are
10
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 9
diff changeset
   129
represented as inductively defined predicates too.
8
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   130
12
dd400b5797e1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 10
diff changeset
   131
The only satisfying way out of this problem in a theorem prover based on classical
10
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 9
diff changeset
   132
logic is to formalise a theory of computability. Norrish provided such
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 9
diff changeset
   133
a formalisation for the HOL4 theorem prover. He choose the
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 9
diff changeset
   134
$\lambda$-calculus as the starting point for his formalisation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 9
diff changeset
   135
of computability theory,
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 9
diff changeset
   136
because of its ``simplicity'' \cite[Page 297]{Norrish11}.  Part of his
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 9
diff changeset
   137
formalisation is a clever infrastructure for reducing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 9
diff changeset
   138
$\lambda$-terms. He also established the computational equivalence
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 9
diff changeset
   139
between the $\lambda$-calculus and recursive functions.  Nevertheless he
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 9
diff changeset
   140
concluded that it would be ``appealing'' to have formalisations for more
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 9
diff changeset
   141
operational models of computations, such as Turing machines or register
12
dd400b5797e1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 10
diff changeset
   142
machines.  One reason is that many proofs in the literature use 
10
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 9
diff changeset
   143
them.  He noted however that in the context of theorem provers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 9
diff changeset
   144
\cite[Page 310]{Norrish11}:
8
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   145
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   146
\begin{quote}
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   147
\it``If register machines are unappealing because of their 
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   148
general fiddliness, Turing machines are an even more 
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   149
daunting prospect.''
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   150
\end{quote}
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   151
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   152
\noindent
33
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   153
In this paper we take on this daunting prospect and provide a
13
a7ec585d7f20 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   154
formalisation of Turing machines, as well as abacus machines (a kind
a7ec585d7f20 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   155
of register machines) and recursive functions. To see the difficulties
a7ec585d7f20 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   156
involved with this work, one has to understand that interactive
a7ec585d7f20 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   157
theorem provers, like Isabelle/HOL, are at their best when the
a7ec585d7f20 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   158
data-structures at hand are ``structurally'' defined, like lists,
a7ec585d7f20 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   159
natural numbers, regular expressions, etc. Such data-structures come
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   160
with convenient reasoning infrastructures (for example induction
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   161
principles, recursion combinators and so on).  But this is \emph{not}
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   162
the case with Turing machines (and also not with register machines):
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   163
underlying their definitions are sets of states together with 
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   164
transition functions, all of which are not structurally defined.  This
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   165
means we have to implement our own reasoning infrastructure in order
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   166
to prove properties about them. This leads to annoyingly fiddly
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   167
formalisations.  We noticed first the difference between both,
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   168
structural and non-structural, ``worlds'' when formalising the
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   169
Myhill-Nerode theorem, where regular expressions fared much better
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   170
than automata \cite{WuZhangUrban11}.  However, with Turing machines
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   171
there seems to be no alternative if one wants to formalise the great
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   172
many proofs from the literature that use them.  We will analyse one
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
   173
example---undecidability of Wang's tiling problem---in Section~\ref{Wang}. The
38
8f8db701f69f updated contribution section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   174
standard proof of this property uses the notion of universal
8f8db701f69f updated contribution section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   175
Turing machines.
12
dd400b5797e1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 10
diff changeset
   176
13
a7ec585d7f20 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   177
We are not the first who formalised Turing machines in a theorem
a7ec585d7f20 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   178
prover: we are aware of the preliminary work by Asperti and Ricciotti
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   179
\cite{AspertiRicciotti12}. They describe a complete formalisation of
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   180
Turing machines in the Matita theorem prover, including a universal
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   181
Turing machine. They report that the informal proofs from which they
37
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   182
started are \emph{not} ``sufficiently accurate to be directly usable as a
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   183
guideline for formalization'' \cite[Page 2]{AspertiRicciotti12}. For
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   184
our formalisation we followed mainly the proofs from the textbook
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   185
\cite{Boolos87} and found that the description there is quite
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   186
detailed. Some details are left out however: for example, it is only
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   187
shown how the universal Turing machine is constructed for Turing
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   188
machines computing unary functions. We had to figure out a way to
37
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   189
generalise this result to $n$-ary functions. Similarly, when compiling
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   190
recursive functions to abacus machines, the textbook again only shows
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   191
how it can be done for 2- and 3-ary functions, but in the
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   192
formalisation we need arbitrary functions. But the general ideas for
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   193
how to do this are clear enough in \cite{Boolos87}. However, one
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   194
aspect that is completely left out from the informal description in
31
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 30
diff changeset
   195
\cite{Boolos87}, and similar ones we are aware of, is arguments why certain Turing
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   196
machines are correct. We will introduce Hoare-style proof rules
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   197
which help us with such correctness arguments of Turing machines.
10
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 9
diff changeset
   198
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   199
The main difference between our formalisation and the one by Asperti
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   200
and Ricciotti is that their universal Turing machine uses a different
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   201
alphabet than the machines it simulates. They write \cite[Page
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   202
23]{AspertiRicciotti12}:
10
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 9
diff changeset
   203
15
90bc8cccc218 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 13
diff changeset
   204
\begin{quote}\it
13
a7ec585d7f20 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   205
``In particular, the fact that the universal machine operates with a
a7ec585d7f20 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   206
different alphabet with respect to the machines it simulates is
a7ec585d7f20 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   207
annoying.'' 
a7ec585d7f20 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   208
\end{quote}
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   209
15
90bc8cccc218 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 13
diff changeset
   210
\noindent
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   211
In this paper we follow the approach by Boolos et al \cite{Boolos87},
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   212
which goes back to Post \cite{Post36}, where all Turing machines
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
   213
operate on tapes that contain only \emph{blank} or \emph{occupied} cells
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   214
(represented by @{term Bk} and @{term Oc}, respectively, in our
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   215
formalisation). Traditionally the content of a cell can be any
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   216
character from a finite alphabet. Although computationally equivalent,
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   217
the more restrictive notion of Turing machines in \cite{Boolos87} makes
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   218
the reasoning more uniform. In addition some proofs \emph{about} Turing
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   219
machines are simpler.  The reason is that one often needs to encode
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
   220
Turing machines---consequently if the Turing machines are simpler, then the coding
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
   221
functions are simpler too. Unfortunately, the restrictiveness also makes
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   222
it harder to design programs for these Turing machines. In order
38
8f8db701f69f updated contribution section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   223
to construct a universal Turing machine we therefore do not follow 
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   224
\cite{AspertiRicciotti12}, instead follow the proof in
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   225
\cite{Boolos87} by relating abacus machines to Turing machines and in
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   226
turn recursive functions to abacus machines. The universal Turing
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
   227
machine can then be constructed as a recursive function.
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   228
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   229
\smallskip
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   230
\noindent
38
8f8db701f69f updated contribution section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   231
{\bf Contributions:} We formalised in Isabelle/HOL Turing machines following the
8f8db701f69f updated contribution section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   232
description of Boolos et al \cite{Boolos87} where tapes only have blank or
8f8db701f69f updated contribution section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   233
occupied cells. We mechanise the undecidability of the halting problem and
8f8db701f69f updated contribution section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   234
prove the correctness of concrete Turing machines that are needed
8f8db701f69f updated contribution section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   235
in this proof; such correctness proofs are left out in the informal literature.  
8f8db701f69f updated contribution section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   236
We construct the universal Turing machine from \cite{Boolos87} by
8f8db701f69f updated contribution section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   237
relating recursive functions to abacus machines and abacus machines to
8f8db701f69f updated contribution section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   238
Turing machines. Since we have set up in Isabelle/HOL a very general computability
8f8db701f69f updated contribution section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   239
model and undecidability result, we are able to formalise the
8f8db701f69f updated contribution section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   240
undecidability of Wang's tiling problem. We are not aware of any other
8f8db701f69f updated contribution section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   241
formalisation of a substantial undecidability problem.
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   242
*}
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   243
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   244
section {* Turing Machines *}
9
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
   245
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
   246
text {* \noindent
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   247
  Turing machines can be thought of as having a read-write-unit, also
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   248
  referred to as \emph{head},
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   249
  ``gliding'' over a potentially infinite tape. Boolos et
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   250
  al~\cite{Boolos87} only consider tapes with cells being either blank
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
   251
  or occupied, which we represent by a datatype having two
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
   252
  constructors, namely @{text Bk} and @{text Oc}.  One way to
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   253
  represent such tapes is to use a pair of lists, written @{term "(l,
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
   254
  r)"}, where @{term l} stands for the tape on the left-hand side of the
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   255
  head and @{term r} for the tape on the right-hand side. We have the
33
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   256
  convention that the head, abbreviated @{term hd}, of the right-list is
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   257
  the cell on which the head of the Turing machine currently operates. This can
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   258
  be pictured as follows:
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   259
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   260
  \begin{center}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   261
  \begin{tikzpicture}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   262
  \draw[very thick] (-3.0,0)   -- ( 3.0,0);
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   263
  \draw[very thick] (-3.0,0.5) -- ( 3.0,0.5);
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   264
  \draw[very thick] (-0.25,0)   -- (-0.25,0.5);
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   265
  \draw[very thick] ( 0.25,0)   -- ( 0.25,0.5);
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   266
  \draw[very thick] (-0.75,0)   -- (-0.75,0.5);
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   267
  \draw[very thick] ( 0.75,0)   -- ( 0.75,0.5);
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   268
  \draw[very thick] (-1.25,0)   -- (-1.25,0.5);
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   269
  \draw[very thick] ( 1.25,0)   -- ( 1.25,0.5);
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   270
  \draw[very thick] (-1.75,0)   -- (-1.75,0.5);
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   271
  \draw[very thick] ( 1.75,0)   -- ( 1.75,0.5);
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   272
  \draw[rounded corners=1mm] (-0.35,-0.1) rectangle (0.35,0.6);
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   273
  \draw[fill]     (1.35,0.1) rectangle (1.65,0.4);
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   274
  \draw[fill]     (0.85,0.1) rectangle (1.15,0.4);
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   275
  \draw[fill]     (-0.35,0.1) rectangle (-0.65,0.4);
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   276
  \draw (-0.25,0.8) -- (-0.25,-0.8);
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   277
  \draw[<->] (-1.25,-0.7) -- (0.75,-0.7);
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   278
  \node [anchor=base] at (-0.8,-0.5) {\small left list};
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   279
  \node [anchor=base] at (0.35,-0.5) {\small right list};
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   280
  \node [anchor=base] at (0.1,0.7) {\small head};
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   281
  \node [anchor=base] at (-2.2,0.2) {\ldots};
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   282
  \node [anchor=base] at ( 2.3,0.2) {\ldots};
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   283
  \end{tikzpicture}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   284
  \end{center}
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   285
  
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   286
  \noindent
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   287
  Note that by using lists each side of the tape is only finite. The
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   288
  potential infinity is achieved by adding an appropriate blank or occupied cell 
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   289
  whenever the head goes over the ``edge'' of the tape. To 
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   290
  make this formal we define five possible \emph{actions} 
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   291
  the Turing machine can perform:
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   292
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   293
  \begin{center}
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   294
  \begin{tabular}{rcll}
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   295
  @{text "a"} & $::=$  & @{term "W0"} & write blank (@{term Bk})\\
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   296
  & $\mid$ & @{term "W1"} & write occupied (@{term Oc})\\
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   297
  & $\mid$ & @{term L} & move left\\
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   298
  & $\mid$ & @{term R} & move right\\
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   299
  & $\mid$ & @{term Nop} & do-nothing operation\\
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   300
  \end{tabular}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   301
  \end{center}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   302
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   303
  \noindent
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
   304
  We slightly deviate
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
   305
  from the presentation in \cite{Boolos87} by using the @{term Nop} operation; however its use
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   306
  will become important when we formalise halting computations and also universal Turing 
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   307
  machines. Given a tape and an action, we can define the
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   308
  following tape updating function:
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   309
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   310
  \begin{center}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   311
  \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   312
  @{thm (lhs) new_tape_def2(1)} & @{text "\<equiv>"} & @{thm (rhs) new_tape_def2(1)}\\
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   313
  @{thm (lhs) new_tape_def2(2)} & @{text "\<equiv>"} & @{thm (rhs) new_tape_def2(2)}\\
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   314
  @{thm (lhs) new_tape_def2(3)} & @{text "\<equiv>"} & \\
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
   315
  \multicolumn{3}{l}{\hspace{1cm}@{thm (rhs) new_tape_def2(3)}}\\
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   316
  @{thm (lhs) new_tape_def2(4)} & @{text "\<equiv>"} & \\
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
   317
  \multicolumn{3}{l}{\hspace{1cm}@{thm (rhs) new_tape_def2(4)}}\\
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   318
  @{thm (lhs) new_tape_def2(5)} & @{text "\<equiv>"} & @{thm (rhs) new_tape_def2(5)}\\
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   319
  \end{tabular}
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   320
  \end{center}
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   321
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   322
  \noindent
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   323
  The first two clauses replace the head of the right-list
37
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   324
  with a new @{term Bk} or @{term Oc}, respectively. To see that
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   325
  these two clauses make sense in case where @{text r} is the empty
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   326
  list, one has to know that the tail function, @{term tl}, is defined in
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   327
  Isabelle/HOL
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   328
  such that @{term "tl [] == []"} holds. The third clause 
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   329
  implements the move of the head one step to the left: we need
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   330
  to test if the left-list @{term l} is empty; if yes, then we just prepend a 
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   331
  blank cell to the right-list; otherwise we have to remove the
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   332
  head from the left-list and prepend it to the right-list. Similarly
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   333
  in the fourth clause for a right move action. The @{term Nop} operation
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   334
  leaves the the tape unchanged (last clause).
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   335
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   336
  Note that our treatment of the tape is rather ``unsymmetric''---we
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   337
  have the convention that the head of the right-list is where the
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   338
  head is currently positioned. Asperti and Ricciotti
33
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   339
  \cite{AspertiRicciotti12} also considered such a representation, but
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   340
  dismiss it as it complicates their definition for \emph{tape
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   341
  equality}. The reason is that moving the head one step to
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   342
  the left and then back to the right might change the tape (in case
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   343
  of going over the ``edge''). Therefore they distinguish four types
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   344
  of tapes: one where the tape is empty; another where the head
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   345
  is on the left edge, respectively right edge, and in the middle
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   346
  of the tape. The reading, writing and moving of the tape is then
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   347
  defined in terms of these four cases.  In this way they can keep the
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   348
  tape in a ``normalised'' form, and thus making a left-move followed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   349
  by a right-move being the identity on tapes. Since we are not using
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   350
  the notion of tape equality, we can get away with the unsymmetric
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   351
  definition above, and by using the @{term update} function
37
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   352
  cover uniformly all cases including corner cases.
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   353
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   354
  Next we need to define the \emph{states} of a Turing machine.  Given
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   355
  how little is usually said about how to represent them in informal
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   356
  presentations, it might be surprising that in a theorem prover we
37
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   357
  have to select carefully a representation. If we use the naive
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   358
  representation where a Turing machine consists of a finite set of
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   359
  states, then we will have difficulties composing two Turing
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   360
  machines: we would need to combine two finite sets of states,
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   361
  possibly renaming states apart whenever both machines share
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   362
  states.\footnote{The usual disjoint union operation in Isabelle/HOL
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   363
  cannot be used as it does not preserve types.} This renaming can be
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   364
  quite cumbersome to reason about. Therefore we made the choice of
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   365
  representing a state by a natural number and the states of a Turing
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   366
  machine will always consist of the initial segment of natural
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   367
  numbers starting from @{text 0} up to the number of states of the
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   368
  machine. In doing so we can compose two Turing machine by
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   369
  shifting the states of one by an appropriate amount to a higher
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   370
  segment and adjusting some ``next states'' in the other.
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   371
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   372
  An \emph{instruction} @{term i} of a Turing machine is a pair consisting of 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   373
  an action and a natural number (the next state). A \emph{program} @{term p} of a Turing
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   374
  machine is then a list of such pairs. Using as an example the following Turing machine
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   375
  program, which consists of four instructions
29
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 28
diff changeset
   376
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   377
  \begin{equation}
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   378
  \begin{tikzpicture}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   379
  \node [anchor=base] at (0,0) {@{thm dither_def}};
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   380
  \node [anchor=west] at (-1.5,-0.42) {$\underbrace{\hspace{21mm}}_{\text{1st state}}$};
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   381
  \node [anchor=west] at ( 1.1,-0.42) {$\underbrace{\hspace{17mm}}_{\text{2nd state}}$};
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   382
  \node [anchor=west] at (-1.5,0.65) {$\overbrace{\hspace{10mm}}^{\text{@{term Bk}-case}}$};
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   383
  \node [anchor=west] at (-0.1,0.65) {$\overbrace{\hspace{6mm}}^{\text{@{term Oc}-case}}$};
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   384
  \end{tikzpicture}
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   385
  \label{dither}
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   386
  \end{equation}
29
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 28
diff changeset
   387
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 28
diff changeset
   388
  \noindent
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   389
  the reader can see we have organised our Turing machine programs so
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   390
  that segments of two belong to a state. The first component of the
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   391
  segment determines what action should be taken and which next state
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   392
  should be transitioned to in case the head reads a @{term Bk};
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   393
  similarly the second component determines what should be done in
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   394
  case of reading @{term Oc}. We have the convention that the first
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   395
  state is always the \emph{starting state} of the Turing machine.
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   396
  The zeroth state is special in that it will be used as the
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   397
  ``halting state''.  There are no instructions for the @{text
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   398
  0}-state, but it will always perform a @{term Nop}-operation and
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   399
  remain in the @{text 0}-state.  Unlike Asperti and Riccioti
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   400
  \cite{AspertiRicciotti12}, we have chosen a very concrete
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   401
  representation for programs, because when constructing a universal
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   402
  Turing machine, we need to define a coding function for programs.
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   403
  This can be easily done for our programs-as-lists, but is more
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   404
  difficult for the functions used by Asperti and Ricciotti.
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   405
29
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 28
diff changeset
   406
  Given a program @{term p}, a state
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   407
  and the cell being read by the head, we need to fetch
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   408
  the corresponding instruction from the program. For this we define 
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   409
  the function @{term fetch}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   410
 
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   411
  \begin{center}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   412
  \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   413
  \multicolumn{3}{l}{@{thm fetch_def2(1)[where b=DUMMY]}}\\
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   414
  @{thm (lhs) fetch_def2(2)} & @{text "\<equiv>"} & \\
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   415
  \multicolumn{3}{@ {\hspace{1cm}}l}{@{text "case nth_of p (2 * s) of"}}\\
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   416
  \multicolumn{3}{@ {\hspace{1.4cm}}l}{@{text "None \<Rightarrow> (Nop, 0) |"}}\\
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   417
  \multicolumn{3}{@ {\hspace{1.4cm}}l}{@{text "Some i \<Rightarrow> i"}}\\
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   418
  @{thm (lhs) fetch_def2(3)} & @{text "\<equiv>"} & \\
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   419
  \multicolumn{3}{@ {\hspace{1cm}}l}{@{text "case nth_of p (2 * s + 1) of"}}\\
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   420
  \multicolumn{3}{@ {\hspace{1.4cm}}l}{@{text "None \<Rightarrow> (Nop, 0) |"}}\\
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   421
  \multicolumn{3}{@ {\hspace{1.4cm}}l}{@{text "Some i \<Rightarrow> i"}}
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   422
  \end{tabular}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   423
  \end{center}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   424
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   425
  \noindent
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   426
  In this definition the function @{term nth_of} returns the @{text n}th element
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   427
  from a list, provided it exists (@{term Some}-case), or if it does not, it
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   428
  returns the default action @{term Nop} and the default state @{text 0} 
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   429
  (@{term None}-case). In doing so we slightly deviate from the description
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   430
  in \cite{Boolos87}: if their Turing machines transition to a non-existing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   431
  state, then the computation is halted. We will transition in such cases
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   432
  to the @{text 0}-state. However, with introducing the
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   433
  notion of \emph{well-formed} Turing machine programs we will later exclude such
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   434
  cases and make the  @{text 0}-state the only ``halting state''. A program 
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   435
  @{term p} is said to be well-formed if it satisfies
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   436
  the following three properties:
33
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   437
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   438
  \begin{center}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   439
  \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   440
  @{term "t_correct p"} & @{text "\<equiv>"} & @{term "2 <= length p"}\\
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   441
                        & @{text "\<and>"} & @{term "iseven (length p)"}\\
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   442
                        & @{text "\<and>"} & @{term "\<forall> (a, s) \<in> set p. s <= length p div 2"}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   443
  \end{tabular}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   444
  \end{center}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   445
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   446
  \noindent
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   447
  The first says that @{text p} must have at least an instruction for the starting 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   448
  state; the second that @{text p} has a @{term Bk} and @{term Oc} instruction for every 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   449
  state, and the third that every next-state is one of the states mentioned in
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   450
  the program or being the @{text 0}-state.
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   451
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   452
  A \emph{configuration} @{term c} of a Turing machine is a state together with 
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   453
  a tape. This is written as @{text "(s, (l, r))"}. If we have a 
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   454
  configuration and a program, we can calculate
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   455
  what the next configuration is by fetching the appropriate action and next state
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   456
  from the program, and by updating the state and tape accordingly. 
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   457
  This single step of execution is defined as the function @{term tstep}
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   458
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   459
  \begin{center}
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   460
  \begin{tabular}{l}
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   461
  @{text "step (s, (l, r)) p"} @{text "\<equiv>"}\\
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   462
  \hspace{10mm}@{text "let (a, s) = fetch p s (read r)"}\\
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   463
  \hspace{10mm}@{text "in (s', update (l, r) a)"}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   464
  \end{tabular}
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   465
  \end{center}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   466
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   467
  \noindent
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   468
  where @{term "read r"} returns the head of the list @{text r}, or if @{text r} is
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   469
  empty it returns @{term Bk}.
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   470
  It is impossible in Isabelle/HOL to lift the @{term step}-function realising
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   471
  a general evaluation function for Turing machines. The reason is that functions in HOL-based
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   472
  provers need to be terminating, and clearly there are Turing machine 
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   473
  programs that are not. We can however define an evaluation
33
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   474
  function so that it performs exactly @{text n} steps:
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   475
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   476
  \begin{center}
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   477
  \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   478
  @{thm (lhs) steps.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) steps.simps(1)}\\
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   479
  @{thm (lhs) steps.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) steps.simps(2)}\\
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   480
  \end{tabular}
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   481
  \end{center}
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   482
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   483
  \noindent
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   484
  Recall our definition of @{term fetch} with the default value for
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   485
  the @{text 0}-state. In case a Turing program takes in \cite{Boolos87} less 
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   486
  then @{text n} steps before it halts, then in our setting the @{term steps}-evaluation
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   487
  does not actually halt, but rather transitions to the @{text 0}-state and 
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   488
  remains there performing @{text Nop}-actions until @{text n} is reached. 
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   489
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   490
  Given some input tape @{text "(l\<^isub>i,r\<^isub>i)"}, we can define when a program 
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   491
  @{term p} generates a specific  output tape @{text "(l\<^isub>o,r\<^isub>o)"}
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   492
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   493
  \begin{center}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   494
  \begin{tabular}{l}
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   495
  @{term "runs p (l\<^isub>i, r\<^isub>i) (l\<^isub>o,r\<^isub>o)"} @{text "\<equiv>"}\\
33
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   496
  \hspace{6mm}@{text "\<exists>n. nsteps (1, (l\<^isub>i,r\<^isub>i)) p n = (0, (l\<^isub>o,r\<^isub>o))"}
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   497
  \end{tabular}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   498
  \end{center}
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   499
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   500
  \noindent
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   501
  where @{text 1} stands for the starting state and @{text 0} for our final state.
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   502
  A program @{text p} with input tape @{term "(l\<^isub>i, r\<^isub>i)"} \emph{halts} iff
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   503
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   504
  \begin{center}
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   505
  @{term "halts p (l\<^isub>i, r\<^isub>i) \<equiv>
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   506
  \<exists>l\<^isub>o r\<^isub>o. runs p (l\<^isub>i, r\<^isub>i) (l\<^isub>o,r\<^isub>o)"}
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   507
  \end{center}
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   508
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   509
  \noindent
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   510
  Later on we need to consider specific Turing machines that 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   511
  start with a tape in standard form and halt the computation
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   512
  in standard form. To define a tape in standard form, it is
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   513
  useful to have an operation @{term "tape_of_nat_list DUMMY"} that 
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   514
  translates 
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   515
  lists of natural numbers into tapes. 
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   516
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   517
  \begin{center}
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   518
  \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   519
  @{thm (lhs) tape_of_nat_list_def2(1)} & @{text "\<equiv>"} & @{thm (rhs) tape_of_nat_list_def2(1)}\\
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   520
  @{thm (lhs) tape_of_nat_list_def2(2)} & @{text "\<equiv>"} & @{thm (rhs) tape_of_nat_list_def2(2)}\\
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   521
  @{thm (lhs) tape_of_nat_list_def2(3)} & @{text "\<equiv>"} & @{thm (rhs) tape_of_nat_list_def2(3)}\\
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   522
  \end{tabular}
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   523
  \end{center}
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   524
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   525
  
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   526
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   527
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   528
  By this we mean
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   529
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   530
  \begin{center}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   531
  @{thm haltP_def2[where p="p" and n="n", THEN eq_reflection]}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   532
  \end{center}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   533
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   534
  \noindent
33
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   535
  This means the Turing machine starts with a tape containg @{text n} @{term Oc}s
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   536
  and the head pointing to the first one; the Turing machine
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   537
  halts with a tape consisting of some @{term Bk}s, followed by a 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   538
  ``cluster'' of @{term Oc}s and after that by some @{term Bk}s.
33
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   539
  The head in the output is pointing again at the first @{term Oc}.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   540
  The intuitive meaning of this definition is to start the Turing machine with a
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   541
  tape corresponding to a value @{term n} and producing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   542
  a new tape corresponding to the value @{term l} (the number of @{term Oc}s
33
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   543
  clustered on the output tape).
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   544
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   545
  Before we can prove the undecidability of the halting problem for Turing machines, 
36
4b35e0e0784b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   546
  we have to define how to compose two Turing machines. Given our setup, this is 
37
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   547
  relatively straightforward, if slightly fiddly. We use the following two
36
4b35e0e0784b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   548
  auxiliary functions:
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   549
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   550
  \begin{center}
36
4b35e0e0784b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   551
  \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}}
4b35e0e0784b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   552
  @{thm (lhs) tshift_def2} @{text "\<equiv>"}\\
4b35e0e0784b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   553
  \hspace{4mm}@{thm (rhs) tshift_def2}\\
4b35e0e0784b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   554
  @{thm (lhs) change_termi_state_def2} @{text "\<equiv>"}\\
4b35e0e0784b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   555
  \hspace{4mm}@{text "map (\<lambda> (a, s)."}\\
4b35e0e0784b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   556
  \hspace{14mm}@{text "(a, if s = 0 then length p div 2 + 1 else s)) p"}\\
4b35e0e0784b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   557
  \end{tabular}
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   558
  \end{center}
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   559
36
4b35e0e0784b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   560
  \noindent
37
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   561
  The first adds @{text n} to all states, exept the @{text 0}-state,
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   562
  thus moving all ``regular'' states to the segment starting at @{text
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   563
  n}; the second adds @{term "length p div 2 + 1"} to the @{text
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   564
  0}-state, thus ridirecting all references to the ``halting state''
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   565
  to the first state after the program @{text p}.  With these two
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   566
  functions in place, we can define the \emph{sequential composition}
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   567
  of two Turing machine programs @{text "p\<^isub>1"} and @{text "p\<^isub>2"}
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   568
36
4b35e0e0784b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   569
  \begin{center}
37
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   570
  @{thm t_add.simps[where ?t1.0="p\<^isub>1" and ?t2.0="p\<^isub>2", THEN eq_reflection]}
36
4b35e0e0784b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   571
  \end{center}
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   572
37
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   573
  \noindent
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   574
  This means @{text "p\<^isub>1"} is executed first. Whenever it originally
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   575
  transitioned to the @{text 0}-state, it will in the composed program transition to the starting
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   576
  state of @{text "p\<^isub>2"} instead. All the states of @{text "p\<^isub>2"}
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   577
  have been shifted in order to make sure that the states of the composed 
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   578
  program @{text "p\<^isub>1 \<oplus> p\<^isub>2"} still only ``occupy'' 
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   579
  an initial segment of the natural numbers.
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   580
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   581
  \begin{center}
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   582
  \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}p{6.9cm}@ {}}
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   583
  @{thm (lhs) tcopy_def} & @{text "\<equiv>"} & @{thm (rhs) tcopy_def}
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   584
  \end{tabular}
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   585
  \end{center}
36
4b35e0e0784b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   586
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   587
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   588
  assertion holds for all tapes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   589
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   590
  Hoare rule for composition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   591
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   592
  For showing the undecidability of the halting problem, we need to consider
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   593
  two specific Turing machines. copying TM and dithering TM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   594
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   595
  correctness of the copying TM
19
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 18
diff changeset
   596
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   597
  measure for the copying TM, which we however omit.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   598
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   599
  halting problem
9
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
   600
*}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
   601
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   602
section {* Abacus Machines *}
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   603
25
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   604
text {*
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   605
  \noindent
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   606
  Boolos et al \cite{Boolos87} use abacus machines as a 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   607
  stepping stone for making it less laborious to write
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   608
  programs for Turing machines. Abacus machines operate
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   609
  over an unlimited number of registers $R_0$, $R_1$, \ldots
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   610
  each being able to hold an arbitrary large natural number.
27
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   611
  We use natural numbers to refer to registers, but also
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   612
  to refer to \emph{opcodes} of abacus 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   613
  machines. Obcodes are given by the datatype
25
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   614
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   615
  \begin{center}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   616
  \begin{tabular}{rcll}
27
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   617
  @{text "o"} & $::=$  & @{term "Inc R\<iota>"} & increment register $R$ by one\\
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   618
  & $\mid$ & @{term "Dec R\<iota> o\<iota>"} & if content of $R$ is non-zero,\\
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   619
  & & & then decrement it by one\\
26
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 25
diff changeset
   620
  & & & otherwise jump to opcode $o$\\
27
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   621
  & $\mid$ & @{term "Goto o\<iota>"} & jump to opcode $o$
25
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   622
  \end{tabular}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   623
  \end{center}
27
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   624
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   625
  \noindent
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   626
  A \emph{program} of an abacus machine is a list of such
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   627
  obcodes. For example the program clearing the register
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   628
  $R$ (setting it to 0) can be defined as follows:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   629
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   630
  \begin{center}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   631
  @{thm clear.simps[where n="R\<iota>" and e="o\<iota>", THEN eq_reflection]}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   632
  \end{center}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   633
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   634
  \noindent
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   635
  The second opcode @{term "Goto 0"} in this programm means we 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   636
  jump back to the first opcode, namely @{text "Dec R o"}.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   637
  The \emph{memory} $m$ of an abacus machine holding the values
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   638
  of the registers is represented as a list of natural numbers.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   639
  We have a lookup function for this memory, written @{term "abc_lm_v m R\<iota>"},
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   640
  which looks up the content of register $R$; if $R$
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   641
  is not in this list, then we return 0. Similarly we
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   642
  have a setting function, written @{term "abc_lm_s m R\<iota> n"}, which
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   643
  sets the value of $R$ to $n$, and if $R$ was not yet in $m$
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   644
  it pads it approriately with 0s.
29
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 28
diff changeset
   645
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 28
diff changeset
   646
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 28
diff changeset
   647
  Abacus machine halts when it jumps out of range.
25
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   648
*}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   649
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   650
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   651
section {* Recursive Functions *}
7
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   652
13
a7ec585d7f20 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   653
section {* Wang Tiles\label{Wang} *}
7
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   654
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   655
text {*
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   656
  Used in texture mapings - graphics
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   657
*}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   658
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   659
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   660
section {* Related Work *}
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   661
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   662
text {*
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   663
  The most closely related work is by Norrish \cite{Norrish11}, and Asperti and 
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   664
  Ricciotti \cite{AspertiRicciotti12}. Norrish bases his approach on 
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   665
  lambda-terms. For this he introduced a clever rewriting technology
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   666
  based on combinators and de-Bruijn indices for
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   667
  rewriting modulo $\beta$-equivalence (to keep it manageable)
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   668
*}
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   669
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   670
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   671
(*
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   672
Questions:
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   673
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   674
Can this be done: Ackerman function is not primitive 
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   675
recursive (Nora Szasz)
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   676
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   677
Tape is represented as two lists (finite - usually infinite tape)?
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   678
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   679
*)
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   680
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   681
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   682
(*<*)
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   683
end
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   684
(*>*)