Paper/Paper.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Sat, 26 Jan 2013 12:00:21 +0000
changeset 86 046c9ecf9150
parent 85 e6f395eccfe7
child 87 cf6e89b5f702
permissions -rw-r--r--
updated paper
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
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
     3
imports "../thys/uncomputable"
6
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
48
559e5c6e5113 updated to ITP and updated directories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
     6
(*
25
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
     7
hide_const (open) s 
48
559e5c6e5113 updated to ITP and updated directories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
     8
*)
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
     9
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
    10
hide_const (open) Divides.adjust
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
    11
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
    12
abbreviation
48
559e5c6e5113 updated to ITP and updated directories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
    13
  "update2 p a \<equiv> update a p"
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
    14
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    15
consts DUMMY::'a
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    16
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    17
notation (latex output)
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    18
  Cons ("_::_" [78,77] 73) and
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
    19
  set ("") and
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    20
  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
    21
  W1 ("W\<^bsub>\<^raw:\hspace{-2pt}>Oc\<^esub>") and
48
559e5c6e5113 updated to ITP and updated directories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
    22
  update2 ("update") and
63
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
    23
  tm_wf0 ("wf") and
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
    24
  (*is_even ("iseven") and*)
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
    25
  tcopy_init ("copy\<^bsub>begin\<^esub>") and
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
    26
  tcopy_loop ("copy\<^bsub>loop\<^esub>") and
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
    27
  tcopy_end ("copy\<^bsub>end\<^esub>") and
75
97eaf7514988 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
    28
  step0 ("step") and
97eaf7514988 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
    29
  steps0 ("steps") and
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
    30
  exponent ("_\<^bsup>_\<^esup>") and
48
559e5c6e5113 updated to ITP and updated directories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
    31
(*  abc_lm_v ("lookup") and
559e5c6e5113 updated to ITP and updated directories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
    32
  abc_lm_s ("set") and*)
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
    33
  haltP ("stdhalt") and 
37
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
    34
  tcopy ("copy") and 
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
    35
  tape_of ("\<langle>_\<rangle>") and 
48
559e5c6e5113 updated to ITP and updated directories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
    36
  tm_comp ("_ \<oplus> _") and 
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    37
  DUMMY  ("\<^raw:\mbox{$\_$}>")
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    38
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    39
declare [[show_question_marks = false]]
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    40
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    41
(* THEOREMS *)
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    42
notation (Rule output)
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    43
  "==>"  ("\<^raw:\mbox{}\inferrule{\mbox{>_\<^raw:}}>\<^raw:{\mbox{>_\<^raw:}}>")
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    44
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    45
syntax (Rule output)
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    46
  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    47
  ("\<^raw:\mbox{}\inferrule{>_\<^raw:}>\<^raw:{\mbox{>_\<^raw:}}>")
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    48
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    49
  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" 
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    50
  ("\<^raw:\mbox{>_\<^raw:}\\>/ _")
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    51
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    52
  "_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>")
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    53
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    54
notation (Axiom output)
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    55
  "Trueprop"  ("\<^raw:\mbox{}\inferrule{\mbox{}}{\mbox{>_\<^raw:}}>")
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    56
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    57
notation (IfThen output)
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    58
  "==>"  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    59
syntax (IfThen output)
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    60
  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    61
  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    62
  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}> /\<^raw:{\normalsize \,>and\<^raw:\,}>/ _")
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    63
  "_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>")
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    64
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    65
notation (IfThenNoBox output)
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    66
  "==>"  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    67
syntax (IfThenNoBox output)
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    68
  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    69
  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    70
  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("_ /\<^raw:{\normalsize \,>and\<^raw:\,}>/ _")
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    71
  "_asm" :: "prop \<Rightarrow> asms" ("_")
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    72
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
    73
lemma nats2tape:
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
    74
  shows "<([]::nat list)> = []"
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
    75
  and "<[n]> = Oc \<up> (n + 1)"
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
    76
  and "ns \<noteq> [] \<Longrightarrow> <n#ns> = Oc \<up> (n + 1) @ [Bk] @ <ns>"
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
    77
apply(auto simp add: tape_of_nl_abv tape_of_nat_list.simps)
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
    78
apply(case_tac ns)
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
    79
apply(auto)
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
    80
done
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
    81
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    82
(*>*)
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    83
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    84
section {* Introduction *}
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    85
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
    86
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
    87
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    88
text {*
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    89
49
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
    90
%\noindent
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
    91
%We formalised in earlier work the correctness proofs for two
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
    92
%algorithms in Isabelle/HOL---one about type-checking in
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
    93
%LF~\cite{UrbanCheneyBerghofer11} and another about deciding requests
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
    94
%in access control~\cite{WuZhangUrban12}.  The formalisations
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
    95
%uncovered a gap in the informal correctness proof of the former and
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
    96
%made us realise that important details were left out in the informal
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
    97
%model for the latter. However, in both cases we were unable to
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
    98
%formalise in Isabelle/HOL computability arguments about the
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
    99
%algorithms. 
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   100
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   101
8
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   102
\noindent
79
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   103
Suppose you want to mechanise a proof for whether a predicate @{term
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   104
P}, say, is decidable or not. Decidability of @{text P} usually
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   105
amounts to showing whether \mbox{@{term "P \<or> \<not>P"}} holds. But this
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   106
does \emph{not} work in Isabelle/HOL and other HOL theorem provers,
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   107
since they are based on classical logic where the law of excluded
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   108
middle ensures that \mbox{@{term "P \<or> \<not>P"}} is always provable no
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   109
matter whether @{text P} is constructed by computable means. We hit on
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   110
this limitation previously when we mechanised the correctness proofs
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   111
of two algorithms \cite{UrbanCheneyBerghofer11,WuZhangUrban12}, but
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   112
were unable to formalise arguments about decidability.
49
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   113
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   114
%The same problem would arise if we had formulated
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   115
%the algorithms as recursive functions, because internally in
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   116
%Isabelle/HOL, like in all HOL-based theorem provers, functions are
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   117
%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
   118
49
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   119
The only satisfying way out of this problem in a theorem prover based
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   120
on classical logic is to formalise a theory of computability. Norrish
61
7edbd5657702 updated files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 52
diff changeset
   121
provided such a formalisation for the HOL. He choose
49
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   122
the $\lambda$-calculus as the starting point for his formalisation of
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   123
computability theory, because of its ``simplicity'' \cite[Page
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   124
297]{Norrish11}.  Part of his formalisation is a clever infrastructure
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   125
for reducing $\lambda$-terms. He also established the computational
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   126
equivalence between the $\lambda$-calculus and recursive functions.
52
2cb1e4499983 updated before_final
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 50
diff changeset
   127
Nevertheless he concluded that it would be appealing
49
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   128
 to have formalisations for more operational models of
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   129
computations, such as Turing machines or register machines.  One
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   130
reason is that many proofs in the literature use them.  He noted
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   131
however that \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
   132
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   133
\begin{quote}
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   134
\it``If register machines are unappealing because of their 
49
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   135
general fiddliness,\\ Turing machines are an even more 
8
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   136
daunting prospect.''
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   137
\end{quote}
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   138
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   139
\noindent
33
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   140
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
   141
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
   142
of register machines) and recursive functions. To see the difficulties
49
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   143
involved with this work, one has to understand that Turing machine
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   144
programs can be completely \emph{unstructured}, behaving similar to
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   145
Basic programs involving the infamous goto \cite{Dijkstra68}. This
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   146
precludes in the general case a compositional Hoare-style reasoning
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   147
about Turing programs.  We provide such Hoare-rules for when it
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   148
\emph{is} possible to reason in a compositional manner (which is
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   149
fortunately quite often), but also tackle the more complicated case
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   150
when we translate abacus programs into Turing programs.  This
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   151
reasoning about concrete Turing machine programs is usually
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   152
left out in the informal literature, e.g.~\cite{Boolos87}.
12
dd400b5797e1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 10
diff changeset
   153
49
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   154
%To see the difficulties
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   155
%involved with this work, one has to understand that interactive
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   156
%theorem provers, like Isabelle/HOL, are at their best when the
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   157
%data-structures at hand are ``structurally'' defined, like lists,
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   158
%natural numbers, regular expressions, etc. Such data-structures come
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   159
%with convenient reasoning infrastructures (for example induction
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   160
%principles, recursion combinators and so on).  But this is \emph{not}
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   161
%the case with Turing machines (and also not with register machines):
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   162
%underlying their definitions are sets of states together with 
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   163
%transition functions, all of which are not structurally defined.  This
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   164
%means we have to implement our own reasoning infrastructure in order
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   165
%to prove properties about them. This leads to annoyingly fiddly
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   166
%formalisations.  We noticed first the difference between both,
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   167
%structural and non-structural, ``worlds'' when formalising the
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   168
%Myhill-Nerode theorem, where regular expressions fared much better
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   169
%than automata \cite{WuZhangUrban11}.  However, with Turing machines
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   170
%there seems to be no alternative if one wants to formalise the great
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   171
%many proofs from the literature that use them.  We will analyse one
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   172
%example---undecidability of Wang's tiling problem---in Section~\ref{Wang}. The
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   173
%standard proof of this property uses the notion of universal
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   174
%Turing machines.
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   175
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   176
We are not the first who formalised Turing machines: we are aware 
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   177
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
   178
\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
   179
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
   180
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
   181
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
   182
guideline for formalization'' \cite[Page 2]{AspertiRicciotti12}. For
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   183
our formalisation we follow mainly the proofs from the textbook
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   184
\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
   185
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
   186
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
   187
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
   188
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
   189
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
   190
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
   191
formalisation we need arbitrary functions. But the general ideas for
49
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   192
how to do this are clear enough in \cite{Boolos87}. 
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   193
%However, one
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   194
%aspect that is completely left out from the informal description in
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   195
%\cite{Boolos87}, and similar ones we are aware of, is arguments why certain Turing
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   196
%machines are correct. We will introduce Hoare-style proof rules
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
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
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   213
operate on tapes that contain only \emph{blank} or \emph{occupied} cells. 
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   214
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
   215
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
   216
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
   217
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
   218
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
   219
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
   220
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
   221
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
   222
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
   223
\cite{AspertiRicciotti12}, instead follow the proof in
75
97eaf7514988 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   224
\cite{Boolos87} by translating abacus machines to Turing machines and in
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   225
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
   226
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
   227
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   228
\smallskip
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   229
\noindent
38
8f8db701f69f updated contribution section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   230
{\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
   231
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
   232
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
   233
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
   234
in this proof; such correctness proofs are left out in the informal literature.  
72
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   235
For reasoning about Turing machine programs we derive Hoare-rules.
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   236
We also construct the universal Turing machine from \cite{Boolos87} by
79
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   237
translating recursive functions to abacus machines and abacus machines to
38
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
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   239
model and undecidability result, we are able to formalise other
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   240
results: we describe a proof of the computational equivalence
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   241
of single-sided Turing machines, which is not given in \cite{Boolos87},
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   242
but needed for example for formalising the undecidability proof of 
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   243
Wang's tiling problem \cite{Robinson71}.
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   244
%We are not aware of any other
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   245
%formalisation of a substantial undecidability problem.
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   246
*}
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   247
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   248
section {* Turing Machines *}
9
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
   249
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
   250
text {* \noindent
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   251
  Turing machines can be thought of as having a \emph{head},
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   252
  ``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
   253
  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
   254
  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
   255
  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
   256
  represent such tapes is to use a pair of lists, written @{term "(l,
75
97eaf7514988 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   257
  r)"}, where @{term l} stands for the tape on the left-hand side of
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   258
  the head and @{term r} for the tape on the right-hand side.  We use
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   259
  the notation @{term "Bk \<up> n"} (similarly @{term "Oc \<up> n"}) for lists
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   260
  composed of @{term n} elements of @{term Bk}.  We also have the
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   261
  convention that the head, abbreviated @{term hd}, of the right-list
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   262
  is the cell on which the head of the Turing machine currently
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   263
  scannes. This can be pictured as follows:
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   264
  %
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   265
  \begin{center}
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   266
  \begin{tikzpicture}[scale=0.9]
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   267
  \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
   268
  \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
   269
  \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
   270
  \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
   271
  \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
   272
  \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
   273
  \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
   274
  \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
   275
  \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
   276
  \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
   277
  \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
   278
  \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
   279
  \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
   280
  \draw[fill]     (-0.35,0.1) rectangle (-0.65,0.4);
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   281
  \draw[fill]     (-1.65,0.1) rectangle (-1.35,0.4);
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   282
  \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
   283
  \draw[<->] (-1.25,-0.7) -- (0.75,-0.7);
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   284
  \node [anchor=base] at (-0.85,-0.5) {\small left list};
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   285
  \node [anchor=base] at (0.40,-0.5) {\small right list};
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   286
  \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
   287
  \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
   288
  \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
   289
  \end{tikzpicture}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   290
  \end{center}
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   291
  
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   292
  \noindent
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   293
  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
   294
  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
   295
  whenever the head goes over the ``edge'' of the tape. To 
79
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   296
  make this formal we define five possible \emph{actions} 
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   297
  the Turing machine can perform:
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   298
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   299
  \begin{center}
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   300
  \begin{tabular}[t]{@ {}rcl@ {\hspace{2mm}}l}
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   301
  @{text "a"} & $::=$  & @{term "W0"} & (write blank, @{term Bk})\\
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   302
  & $\mid$ & @{term "W1"} & (write occupied, @{term Oc})\\
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   303
  \end{tabular}
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   304
  \begin{tabular}[t]{rcl@ {\hspace{2mm}}l}
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   305
  & $\mid$ & @{term L} & (move left)\\
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   306
  & $\mid$ & @{term R} & (move right)\\
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   307
  \end{tabular}
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   308
  \begin{tabular}[t]{rcl@ {\hspace{2mm}}l@ {}}
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   309
  & $\mid$ & @{term Nop} & (do-nothing operation)\\
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   310
  \end{tabular}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   311
  \end{center}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   312
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   313
  \noindent
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
   314
  We slightly deviate
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
   315
  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
   316
  will become important when we formalise halting computations and also universal Turing 
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   317
  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
   318
  following tape updating function:
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   319
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   320
  \begin{center}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   321
  \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
48
559e5c6e5113 updated to ITP and updated directories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   322
  @{thm (lhs) update.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) update.simps(1)}\\
559e5c6e5113 updated to ITP and updated directories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   323
  @{thm (lhs) update.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) update.simps(2)}\\
49
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   324
  @{thm (lhs) update.simps(3)} & @{text "\<equiv>"} & @{thm (rhs) update.simps(3)}\\
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   325
  @{thm (lhs) update.simps(4)} & @{text "\<equiv>"} & @{thm (rhs) update.simps(4)}\\
48
559e5c6e5113 updated to ITP and updated directories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   326
  @{thm (lhs) update.simps(5)} & @{text "\<equiv>"} & @{thm (rhs) update.simps(5)}\\
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   327
  \end{tabular}
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   328
  \end{center}
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   329
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   330
  \noindent
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   331
  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
   332
  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
   333
  these two clauses make sense in case where @{text r} is the empty
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   334
  list, one has to know that the tail function, @{term tl}, is defined
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   335
  such that @{term "tl [] == []"} holds. The third clause 
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   336
  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
   337
  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
   338
  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
   339
  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
   340
  in the fourth clause for a right move action. The @{term Nop} operation
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   341
  leaves the the tape unchanged.
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   342
49
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   343
  %Note that our treatment of the tape is rather ``unsymmetric''---we
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   344
  %have the convention that the head of the right-list is where the
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   345
  %head is currently positioned. Asperti and Ricciotti
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   346
  %\cite{AspertiRicciotti12} also considered such a representation, but
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   347
  %dismiss it as it complicates their definition for \emph{tape
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   348
  %equality}. The reason is that moving the head one step to
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   349
  %the left and then back to the right might change the tape (in case
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   350
  %of going over the ``edge''). Therefore they distinguish four types
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   351
  %of tapes: one where the tape is empty; another where the head
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   352
  %is on the left edge, respectively right edge, and in the middle
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   353
  %of the tape. The reading, writing and moving of the tape is then
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   354
  %defined in terms of these four cases.  In this way they can keep the
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   355
  %tape in a ``normalised'' form, and thus making a left-move followed
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   356
  %by a right-move being the identity on tapes. Since we are not using
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   357
  %the notion of tape equality, we can get away with the unsymmetric
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   358
  %definition above, and by using the @{term update} function
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   359
  %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
   360
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   361
  Next we need to define the \emph{states} of a Turing machine.  
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   362
  %Given
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   363
  %how little is usually said about how to represent them in informal
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   364
  %presentations, it might be surprising that in a theorem prover we
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   365
  %have to select carefully a representation. If we use the naive
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   366
  %representation where a Turing machine consists of a finite set of
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   367
  %states, then we will have difficulties composing two Turing
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   368
  %machines: we would need to combine two finite sets of states,
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   369
  %possibly renaming states apart whenever both machines share
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   370
  %states.\footnote{The usual disjoint union operation in Isabelle/HOL
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   371
  %cannot be used as it does not preserve types.} This renaming can be
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   372
  %quite cumbersome to reason about. 
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   373
  We follow the choice made in \cite{AspertiRicciotti12} 
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   374
  by representing a state with a natural number and the states in a Turing
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   375
  machine program by the initial segment of natural numbers starting from @{text 0}.
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   376
  In doing so we can compose two Turing machine programs by
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   377
  shifting the states of one by an appropriate amount to a higher
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   378
  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
   379
63
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
   380
  An \emph{instruction} of a Turing machine is a pair consisting of 
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   381
  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
   382
  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
   383
  program, which consists of four instructions
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   384
  %
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   385
  \begin{equation}
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   386
  \begin{tikzpicture}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   387
  \node [anchor=base] at (0,0) {@{thm dither_def}};
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   388
  \node [anchor=west] at (-1.5,-0.64) 
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   389
  {$\underbrace{\hspace{21mm}}_{\text{\begin{tabular}{@ {}l@ {}}1st state\\[-2mm] 
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   390
  = starting state\end{tabular}}}$};
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   391
  
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   392
  \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
   393
  \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
   394
  \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
   395
  \end{tikzpicture}
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   396
  \label{dither}
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   397
  \end{equation}
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   398
  %
29
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 28
diff changeset
   399
  \noindent
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   400
  the reader can see we have organised our Turing machine programs so
79
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   401
  that segments of two belong to a state. The first component of such a
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   402
  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
   403
  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
   404
  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
   405
  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
   406
  state is always the \emph{starting state} of the Turing machine.
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   407
  The @{text 0}-state is special in that it will be used as the
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   408
  ``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
   409
  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
   410
  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
   411
  \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
   412
  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
   413
  Turing machine, we need to define a coding function for programs.
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   414
  This can be directly done for our programs-as-lists, but is
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   415
  slightly more difficult for the functions used by Asperti and Ricciotti.
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   416
29
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 28
diff changeset
   417
  Given a program @{term p}, a state
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   418
  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
   419
  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
   420
  the function @{term fetch}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   421
 
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   422
  \begin{equation}\label{fetch}
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   423
  \mbox{\begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
48
559e5c6e5113 updated to ITP and updated directories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   424
  \multicolumn{3}{l}{@{thm fetch.simps(1)[where b=DUMMY]}}\\
49
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   425
  @{thm (lhs) fetch.simps(2)} & @{text "\<equiv>"} & @{text "case nth_of p (2 * s) of"}\\
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   426
  \multicolumn{3}{@ {\hspace{4cm}}l}{@{text "None \<Rightarrow> (Nop, 0) | Some i \<Rightarrow> i"}}\\
49
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   427
  @{thm (lhs) fetch.simps(3)} & @{text "\<equiv>"} & @{text "case nth_of p (2 * s + 1) of"}\\
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   428
  \multicolumn{3}{@ {\hspace{4cm}}l}{@{text "None \<Rightarrow> (Nop, 0) | Some i \<Rightarrow> i"}}
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   429
  \end{tabular}}
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   430
  \end{equation}
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   431
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   432
  \noindent
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   433
  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
   434
  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
   435
  returns the default action @{term Nop} and the default state @{text 0} 
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   436
  (@{term None}-case). We often need to restrict Turing machine programs 
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   437
  to be well-formed: a program @{term p} is \emph{well-formed} if it 
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   438
  satisfies the following three properties:
33
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   439
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   440
  \begin{center}
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   441
  @{thm tm_wf.simps[where p="p" and off="0::nat", simplified, THEN eq_reflection]}
33
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   442
  \end{center}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   443
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   444
  \noindent
75
97eaf7514988 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   445
  The first states that @{text p} must have at least an instruction for the starting 
33
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   446
  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
   447
  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
   448
  the program or being the @{text 0}-state.
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   449
72
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   450
  We need to be able to sequentially compose Turing machine programs. Given our
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   451
  concrete representation, this is relatively straightforward, if
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   452
  slightly fiddly. We use the following two auxiliary functions:
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   453
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   454
  \begin{center}
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   455
  \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}}
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   456
  @{thm (lhs) shift.simps} @{text "\<equiv>"} @{thm (rhs) shift.simps}\\
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   457
  @{thm (lhs) adjust.simps} @{text "\<equiv>"} @{thm (rhs) adjust.simps}\\
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   458
  \end{tabular}
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   459
  \end{center}
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   460
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   461
  \noindent
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   462
  The first adds @{text n} to all states, exept the @{text 0}-state,
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   463
  thus moving all ``regular'' states to the segment starting at @{text
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   464
  n}; the second adds @{term "Suc(length p div 2)"} to the @{text
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   465
  0}-state, thus redirecting all references to the ``halting state''
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   466
  to the first state after the program @{text p}.  With these two
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   467
  functions in place, we can define the \emph{sequential composition}
79
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   468
  of two Turing machine programs @{text "p\<^isub>1"} and @{text "p\<^isub>2"} as
72
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   469
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   470
  \begin{center}
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   471
  @{thm tm_comp.simps[where ?p1.0="p\<^isub>1" and ?p2.0="p\<^isub>2", THEN eq_reflection]}
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   472
  \end{center}
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   473
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   474
  \noindent
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   475
  %This means @{text "p\<^isub>1"} is executed first. Whenever it originally
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   476
  %transitioned to the @{text 0}-state, it will in the composed program transition to the starting
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   477
  %state of @{text "p\<^isub>2"} instead. All the states of @{text "p\<^isub>2"}
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   478
  %have been shifted in order to make sure that the states of the composed 
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   479
  %program @{text "p\<^isub>1 \<oplus> p\<^isub>2"} still only ``occupy'' 
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   480
  %an initial segment of the natural numbers.
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   481
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   482
  A \emph{configuration} @{term c} of a Turing machine is a state
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   483
  together with a tape. This is written as @{text "(s, (l, r))"}.  We
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   484
  say a configuration \emph{is final} if @{term "s = (0::nat)"} and we
72
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   485
  say a predicate @{text P} \emph{holds for} a configuration if @{text
79
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   486
  "P"} holds for the tape @{text "(l, r)"}.  If we have a configuration and a program, we can
72
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   487
  calculate what the next configuration is by fetching the appropriate
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   488
  action and next state from the program, and by updating the state
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   489
  and tape accordingly.  This single step of execution is defined as
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   490
  the function @{term step}
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   491
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   492
  \begin{center}
63
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
   493
  \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
   494
  @{text "step (s, (l, r)) p"} & @{text "\<equiv>"} & @{text "let (a, s') = fetch p s (read r)"}\\
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
   495
  & & @{text "in (s', update (l, r) a)"}
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   496
  \end{tabular}
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   497
  \end{center}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   498
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   499
  \noindent
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   500
  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
   501
  empty it returns @{term Bk}.
79
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   502
  It is impossible in Isabelle/HOL to lift the @{term step}-function to realise
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   503
  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
   504
  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
   505
  programs that are not. We can however define an evaluation
79
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   506
  function that performs exactly @{text n} steps:
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   507
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   508
  \begin{center}
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   509
  \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   510
  @{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
   511
  @{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
   512
  \end{tabular}
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   513
  \end{center}
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   514
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   515
  \noindent Recall our definition of @{term fetch} (shown in
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   516
  \eqref{fetch}) with the default value for the @{text 0}-state. In
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   517
  case a Turing program takes according to the usual textbook
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   518
  definition \cite{Boolos87} less than @{text n} steps before it
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   519
  halts, then in our setting the @{term steps}-evaluation does not
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   520
  actually halt, but rather transitions to the @{text 0}-state (the
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   521
  final state) and remains there performing @{text Nop}-actions until
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   522
  @{text n} is reached.
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   523
  
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   524
  \begin{figure}[t]
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   525
  \begin{center}
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   526
  \begin{tabular}{@ {}c@ {\hspace{3mm}}c@ {\hspace{3mm}}c}
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   527
  \begin{tabular}{@ {}l@ {}}
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   528
  @{thm (lhs) tcopy_init_def} @{text "\<equiv>"}\\
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   529
  @{text "["}@{text "(W0, 0), (R, 2), (R, 3),"}\\
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   530
  \phantom{@{text "["}}@{text "(R, 2), (W1, 3), (L, 4),"}\\
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   531
  \phantom{@{text "["}}@{text "(L, 4), (L, 0)"}@{text "]"} 
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   532
  \end{tabular}
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   533
  &
86
046c9ecf9150 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 85
diff changeset
   534
  \begin{tabular}{@ {}l@ {}}
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   535
  @{thm (lhs) tcopy_loop_def} @{text "\<equiv>"}\\
86
046c9ecf9150 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 85
diff changeset
   536
  @{text "["}@{text "(R, 0), (R, 2), (R, 3),"}\\
046c9ecf9150 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 85
diff changeset
   537
  \phantom{@{text "["}}@{text "(W0, 2), (R, 3), (R, 4),"}\\
046c9ecf9150 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 85
diff changeset
   538
  \phantom{@{text "["}}@{text "(W1, 5), (R, 4), (L, 6),"}\\
046c9ecf9150 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 85
diff changeset
   539
  \phantom{@{text "["}}@{text "(L, 5), (L, 6), (L, 1)"}@{text "]"} 
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   540
  \end{tabular}
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   541
  &
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   542
  \begin{tabular}{@ {}p{3.6cm}@ {}}
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   543
  @{thm (lhs) tcopy_end_def} @{text "\<equiv>"}\\
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   544
  @{thm (rhs) tcopy_end_def}
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   545
  \end{tabular}
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   546
  \end{tabular}
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   547
  \end{center}
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   548
  \caption{Copy machine}\label{copy}
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   549
  \end{figure}
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   550
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   551
  We often need to restrict tapes to be in \emph{standard form}, which means 
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   552
  the left list of the tape is either empty or only contains @{text "Bk"}s, and 
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   553
  the right list contains some ``clusters'' of @{text "Oc"}s separted by single 
84
4c8325c64dab updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 81
diff changeset
   554
  blanks. To make this formal we define the following function
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   555
  
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   556
  \begin{center}
84
4c8325c64dab updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 81
diff changeset
   557
  \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}}
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   558
  @{thm (lhs) nats2tape(1)} & @{text "\<equiv>"} & @{thm (rhs) nats2tape(1)}\\
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   559
  @{thm (lhs) nats2tape(2)} & @{text "\<equiv>"} & @{thm (rhs) nats2tape(2)}\\
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   560
  @{thm (lhs) nats2tape(3)} & @{text "\<equiv>"} & @{thm (rhs) nats2tape(3)}
84
4c8325c64dab updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 81
diff changeset
   561
  \end{tabular}
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   562
  \end{center}
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   563
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   564
  \noindent
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   565
  A standard tape is then of the form @{text "(Bk\<^isup>l,\<langle>[n\<^isub>1,...,n\<^isub>m]\<rangle>)"} for some @{text l} 
84
4c8325c64dab updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 81
diff changeset
   566
  and @{text "n\<^isub>i"}. Note that the head in a standard tape ``points'' to the 
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   567
  leftmost @{term "Oc"} on the tape. Note also that @{text 0} is represented by
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   568
  a single filled cell, @{text 1} by two filled cells and so on.
79
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   569
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   570
  Before we can prove the undecidability of the halting problem for our
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   571
  Turing machines, we need to analyse two concrete Turing machine
79
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   572
  programs and establish that they are correct---that means they are
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   573
  ``doing what they are supposed to be doing''.  Such correctness proofs are usually left
75
97eaf7514988 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   574
  out in the informal literature, for example \cite{Boolos87}.  One program 
76
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   575
  we need to prove correct is the @{term dither} program shown in \eqref{dither} 
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   576
  and the other program is @{term "tcopy"} defined as
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   577
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   578
  \begin{center}
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   579
  \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}}
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   580
  @{thm (lhs) tcopy_def} & @{text "\<equiv>"} & @{thm (rhs) tcopy_def}
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   581
  \end{tabular}
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   582
  \end{center}
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   583
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   584
  \noindent
79
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   585
  whose three components are given in Figure~\ref{copy}. To the prove
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   586
  correctness of these Turing machine programs, we introduce the
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   587
  notion of total correctness defined in terms of
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   588
  \emph{Hoare-triples}, written @{term "{P} p {Q}"}. They realise the
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   589
  idea that a program @{term p} started in state @{term "1::nat"} with
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   590
  a tape satisfying @{term P} will after some @{text n} steps halt (have
79
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   591
  transitioned into the halting state) with a tape satisfying @{term
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   592
  Q}. We also have \emph{Hoare-pairs} of the form @{term "{P} p \<up>"}
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   593
  realising the case that a program @{term p} started with a tape
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   594
  satisfying @{term P} will loop (never transition into the halting
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   595
  state). Both notion are formally defined as
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   596
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   597
  \begin{center}
76
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   598
  \begin{tabular}{@ {}c@ {\hspace{4mm}}c@ {}}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   599
  \begin{tabular}[t]{@ {}l@ {}}
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   600
  \colorbox{mygrey}{@{thm (lhs) Hoare_halt_def}} @{text "\<equiv>"}\\[1mm]
76
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   601
  \hspace{5mm}@{text "\<forall>"} @{term "(l, r)"}.\\
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   602
  \hspace{7mm}if @{term "P (l, r)"} holds then\\
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   603
  \hspace{7mm}@{text "\<exists>"} @{term n}. such that\\
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   604
  \hspace{7mm}@{text "is_final (steps (1, (l, r)) p n)"} \hspace{1mm}@{text "\<and>"}\\
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   605
  \hspace{7mm}@{text "Q holds_for (steps (1, (l, r)) p n)"}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   606
  \end{tabular} &
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   607
  \begin{tabular}[t]{@ {}l@ {}}
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   608
  \colorbox{mygrey}{@{thm (lhs) Hoare_unhalt_def}} @{text "\<equiv>"}\\[1mm]
76
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   609
  \hspace{5mm}@{text "\<forall>"} @{term "(l, r)"}.\\ 
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   610
  \hspace{7mm}if @{term "P (l, r)"} holds then\\
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   611
  \hspace{7mm}@{text "\<forall>"} @{term n}. @{text "\<not> is_final (steps (1, (l, r)) p n)"}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   612
  \end{tabular}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   613
  \end{tabular}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   614
  \end{center}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   615
  
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   616
  \noindent
79
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   617
  We have set up our Hoare-style reasoning so that we can deal explicitly 
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   618
  with total correctness and non-terminantion, rather than have notions for partial 
79
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   619
  correctness and termination. Although the latter would allow us to reason
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   620
  more uniformly (only using Hoare-triples), we prefer our definitions because 
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   621
  we can derive simple Hoare-rules for sequentially composed Turing programs. 
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   622
  In this way we can reason about the correctness of @{term "tcopy_init"},
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   623
  for example, completely separately from @{term "tcopy_loop"} and @{term "tcopy_end"}.
76
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   624
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   625
  \begin{center}
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   626
  \begin{tabular}{l@ {\hspace{3mm}}lcl}
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   627
  & \multicolumn{1}{c}{start tape}\\[1mm]
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   628
  \raisebox{2.5mm}{halting case:} &
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   629
  \begin{tikzpicture}[scale=0.9]
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   630
  \draw[very thick] (-2,0)   -- ( 0.75,0);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   631
  \draw[very thick] (-2,0.5) -- ( 0.75,0.5);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   632
  \draw[very thick] (-0.25,0)   -- (-0.25,0.5);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   633
  \draw[very thick] ( 0.25,0)   -- ( 0.25,0.5);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   634
  \draw[very thick] (-0.75,0)   -- (-0.75,0.5);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   635
  \draw[very thick] ( 0.75,0)   -- ( 0.75,0.5);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   636
  \draw[very thick] (-1.25,0)   -- (-1.25,0.5);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   637
  \draw[rounded corners=1mm] (-0.35,-0.1) rectangle (0.35,0.6);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   638
  \draw[fill]     (-0.15,0.1) rectangle (0.15,0.4);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   639
  \draw[fill]     ( 0.35,0.1) rectangle (0.65,0.4);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   640
  \node [anchor=base] at (-1.7,0.2) {\ldots};
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   641
  \end{tikzpicture} 
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   642
  & \raisebox{2.5mm}{$\;\;\large\Rightarrow\;\;$} &
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   643
  \begin{tikzpicture}[scale=0.9]
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   644
  \draw[very thick] (-2,0)   -- ( 0.75,0);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   645
  \draw[very thick] (-2,0.5) -- ( 0.75,0.5);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   646
  \draw[very thick] (-0.25,0)   -- (-0.25,0.5);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   647
  \draw[very thick] ( 0.25,0)   -- ( 0.25,0.5);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   648
  \draw[very thick] (-0.75,0)   -- (-0.75,0.5);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   649
  \draw[very thick] ( 0.75,0)   -- ( 0.75,0.5);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   650
  \draw[very thick] (-1.25,0)   -- (-1.25,0.5);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   651
  \draw[rounded corners=1mm] (-0.35,-0.1) rectangle (0.35,0.6);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   652
  \draw[fill]     (-0.15,0.1) rectangle (0.15,0.4);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   653
  \draw[fill]     ( 0.35,0.1) rectangle (0.65,0.4);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   654
  \node [anchor=base] at (-1.7,0.2) {\ldots};
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   655
  \end{tikzpicture}\\
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   656
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   657
  \raisebox{2.5mm}{non-halting case:} &
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   658
  \begin{tikzpicture}[scale=0.9]
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   659
  \draw[very thick] (-2,0)   -- ( 0.25,0);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   660
  \draw[very thick] (-2,0.5) -- ( 0.25,0.5);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   661
  \draw[very thick] (-0.25,0)   -- (-0.25,0.5);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   662
  \draw[very thick] ( 0.25,0)   -- ( 0.25,0.5);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   663
  \draw[very thick] (-0.75,0)   -- (-0.75,0.5);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   664
  \draw[very thick] (-1.25,0)   -- (-1.25,0.5);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   665
  \draw[rounded corners=1mm] (-0.35,-0.1) rectangle (0.35,0.6);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   666
  \draw[fill]     (-0.15,0.1) rectangle (0.15,0.4);
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   667
  \node [anchor=base] at (-1.7,0.2) {\ldots};
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   668
  \end{tikzpicture} 
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   669
  & \raisebox{2.5mm}{$\;\;\large\Rightarrow\;\;$} &
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   670
  \raisebox{2.5mm}{loops}
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   671
  \end{tabular}
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   672
  \end{center}
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   673
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   674
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   675
  It is straightforward to prove that the Turing program 
76
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   676
  @{term "dither"} satisfies the following correctness properties
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   677
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   678
  \begin{center}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   679
  \begin{tabular}{l}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   680
  @{thm dither_halts}\\
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   681
  @{thm dither_loops}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   682
  \end{tabular}
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   683
  \end{center}
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   684
77
04e5850818c8 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   685
  \noindent
76
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   686
  {\it unfold} The first states that on a tape @{term "(Bk \<up> n,  [Oc, Oc])"}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   687
  halts in tree steps leaving the tape unchanged. In the other states
78
2669235c4b1e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 77
diff changeset
   688
  that @{term dither} started with tape @{term "(Bk \<up> n,  [Oc])"} loops.
76
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   689
  
75
97eaf7514988 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   690
97eaf7514988 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   691
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   692
  In the following we will consider the following Turing machine program
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   693
  that makes a copies a value on the tape.
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   694
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   695
 
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   696
37
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
   697
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   698
  
36
4b35e0e0784b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
   699
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   700
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   701
  assertion holds for all tapes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   702
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   703
  Hoare rule for composition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   704
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   705
  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
   706
  two specific Turing machines. copying TM and dithering TM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   707
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   708
  correctness of the copying TM
19
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 18
diff changeset
   709
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   710
  measure for the copying TM, which we however omit.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   711
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   712
  halting problem
9
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
   713
*}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
   714
63
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
   715
text {*
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
   716
  \begin{center}
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
   717
  \begin{tabular}{@ {}p{3cm}p{3cm}@ {}}
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
   718
  @{thm[mode=Rule] 
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
   719
  Hoare_plus_halt[where ?P1.0="P\<^isub>1" and ?P2.0="P\<^isub>2" and ?Q1.0="Q\<^isub>1" and ?Q2.0="Q\<^isub>2" and ?A="p\<^isub>1" and B="p\<^isub>2"]}
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
   720
  &
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
   721
  @{thm[mode=Rule] 
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
   722
  Hoare_plus_unhalt[where ?P1.0="P\<^isub>1" and ?P2.0="P\<^isub>2" and ?Q1.0="Q\<^isub>1" and ?A="p\<^isub>1" and B="p\<^isub>2"]}
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
   723
  \end{tabular}
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
   724
  \end{center}
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
   725
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
   726
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
   727
*}
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
   728
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   729
section {* Abacus Machines *}
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   730
25
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   731
text {*
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   732
  \noindent
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   733
  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
   734
  stepping stone for making it less laborious to write
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   735
  programs for Turing machines. Abacus machines operate
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   736
  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
   737
  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
   738
  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
   739
  to refer to \emph{opcodes} of abacus 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   740
  machines. Obcodes are given by the datatype
25
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   741
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   742
  \begin{center}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   743
  \begin{tabular}{rcll}
27
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   744
  @{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
   745
  & $\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
   746
  & & & then decrement it by one\\
26
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 25
diff changeset
   747
  & & & otherwise jump to opcode $o$\\
27
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   748
  & $\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
   749
  \end{tabular}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   750
  \end{center}
27
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   751
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   752
  \noindent
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   753
  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
   754
  obcodes. For example the program clearing the register
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   755
  $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
   756
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   757
  \begin{center}
48
559e5c6e5113 updated to ITP and updated directories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
   758
  %@ {thm clear.simps[where n="R\<iota>" and e="o\<iota>", THEN eq_reflection]}
27
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   759
  \end{center}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   760
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   761
  \noindent
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   762
  The second opcode @{term "Goto 0"} in this program means we 
27
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   763
  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
   764
  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
   765
  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
   766
  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
   767
  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
   768
  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
   769
  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
   770
  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
   771
  it pads it approriately with 0s.
29
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 28
diff changeset
   772
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 28
diff changeset
   773
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 28
diff changeset
   774
  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
   775
*}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   776
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   777
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   778
section {* Recursive Functions *}
7
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   779
13
a7ec585d7f20 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   780
section {* Wang Tiles\label{Wang} *}
7
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   781
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   782
text {*
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   783
  Used in texture mapings - graphics
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   784
*}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   785
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
   786
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   787
section {* Related Work *}
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   788
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   789
text {*
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   790
  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
   791
  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
   792
  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
   793
  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
   794
  rewriting modulo $\beta$-equivalence (to keep it manageable)
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   795
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   796
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   797
  
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   798
  Given some input tape @{text "(l\<^isub>i,r\<^isub>i)"}, we can define when a program 
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   799
  @{term p} generates a specific  output tape @{text "(l\<^isub>o,r\<^isub>o)"}
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   800
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   801
  \begin{center}
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   802
  \begin{tabular}{l}
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   803
  @{term "runs p (l\<^isub>i, r\<^isub>i) (l\<^isub>o,r\<^isub>o)"} @{text "\<equiv>"}\\
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   804
  \hspace{6mm}@{text "\<exists>n. nsteps (1, (l\<^isub>i,r\<^isub>i)) p n = (0, (l\<^isub>o,r\<^isub>o))"}
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   805
  \end{tabular}
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   806
  \end{center}
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   807
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   808
  \noindent
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   809
  where @{text 1} stands for the starting state and @{text 0} for our final state.
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   810
  A program @{text p} with input tape @{term "(l\<^isub>i, r\<^isub>i)"} \emph{halts} iff
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   811
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   812
  \begin{center}
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   813
  @{term "halts p (l\<^isub>i, r\<^isub>i) \<equiv>
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   814
  \<exists>l\<^isub>o r\<^isub>o. runs p (l\<^isub>i, r\<^isub>i) (l\<^isub>o,r\<^isub>o)"}
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   815
  \end{center}
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   816
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   817
  \noindent
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   818
  Later on we need to consider specific Turing machines that 
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   819
  start with a tape in standard form and halt the computation
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   820
  in standard form. To define a tape in standard form, it is
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   821
  useful to have an operation %@{ term "tape_of_nat_list DUMMY"} 
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   822
  that translates lists of natural numbers into tapes. 
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   823
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   824
  
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   825
  \begin{center}
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   826
  \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   827
  %@ { thm (lhs) tape_of_nat_list_def2(1)} & @{text "\<equiv>"} & @ { thm (rhs) tape_of_nat_list_def2(1)}\\
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   828
  %@ { thm (lhs) tape_of_nat_list_def2(2)} & @{text "\<equiv>"} & @ { thm (rhs) tape_of_nat_list_def2(2)}\\
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   829
  %@ { thm (lhs) tape_of_nat_list_def2(3)} & @{text "\<equiv>"} & @ { thm (rhs) tape_of_nat_list_def2(3)}\\
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   830
  \end{tabular}
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   831
  \end{center}
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   832
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   833
  
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   834
 By this we mean
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   835
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   836
  \begin{center}
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   837
  %@ {thm haltP_def2[where p="p" and n="n", THEN eq_reflection]}
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   838
  \end{center}
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   839
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   840
  \noindent
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   841
  This means the Turing machine starts with a tape containg @{text n} @{term Oc}s
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   842
  and the head pointing to the first one; the Turing machine
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   843
  halts with a tape consisting of some @{term Bk}s, followed by a 
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   844
  ``cluster'' of @{term Oc}s and after that by some @{term Bk}s.
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   845
  The head in the output is pointing again at the first @{term Oc}.
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   846
  The intuitive meaning of this definition is to start the Turing machine with a
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   847
  tape corresponding to a value @{term n} and producing
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   848
  a new tape corresponding to the value @{term l} (the number of @{term Oc}s
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   849
  clustered on the output tape).
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   850
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   851
*}
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   852
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   853
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   854
(*
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   855
Questions:
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   856
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   857
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
   858
recursive (Nora Szasz)
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   859
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   860
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
   861
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   862
*)
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   863
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   864
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   865
(*<*)
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   866
end
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   867
(*>*)