Paper/Paper.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 10 Jan 2019 12:51:24 +0000
changeset 294 6836da75b3ac
parent 284 a21fb87bb0bd
permissions -rw-r--r--
updated to Isabelle 2016-1
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
198
d93cc4295306 tuned some files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 194
diff changeset
     3
imports "../thys/UTM" "../thys/Abacus_Defs"
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
232
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
     6
(*
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
     7
value "map (steps (1,[],[Oc]) ([(L,0),(L,2),(R,2),(R,0)],0)) [0 ..< 4]"
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
     8
*)
25
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
     9
hide_const (open) s 
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
    10
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    11
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    12
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
    13
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
    14
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
    15
abbreviation
48
559e5c6e5113 updated to ITP and updated directories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
    16
  "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
    17
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    18
consts DUMMY::'a
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    19
169
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 164
diff changeset
    20
(* Theorems as inference rules from LaTeXsugar.thy *)
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    21
notation (Rule output)
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    22
  "==>"  ("\<^raw:\mbox{}\inferrule{\mbox{>_\<^raw:}}>\<^raw:{\mbox{>_\<^raw:}}>")
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    23
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    24
syntax (Rule output)
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    25
  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    26
  ("\<^raw:\mbox{}\inferrule{>_\<^raw:}>\<^raw:{\mbox{>_\<^raw:}}>")
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    27
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    28
  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" 
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    29
  ("\<^raw:\mbox{>_\<^raw:}\\>/ _")
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    30
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    31
  "_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>")
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    32
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    33
notation (Axiom output)
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    34
  "Trueprop"  ("\<^raw:\mbox{}\inferrule{\mbox{}}{\mbox{>_\<^raw:}}>")
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    35
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    36
notation (IfThen output)
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    37
  "==>"  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    38
syntax (IfThen output)
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    39
  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    40
  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    41
  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}> /\<^raw:{\normalsize \,>and\<^raw:\,}>/ _")
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    42
  "_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>")
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    43
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    44
notation (IfThenNoBox output)
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    45
  "==>"  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    46
syntax (IfThenNoBox output)
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    47
  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    48
  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    49
  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("_ /\<^raw:{\normalsize \,>and\<^raw:\,}>/ _")
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    50
  "_asm" :: "prop \<Rightarrow> asms" ("_")
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    51
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    52
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    53
context uncomputable
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    54
begin
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    55
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    56
notation (latex output)
100
dfe852aacae6 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 99
diff changeset
    57
  Cons ("_::_" [48,47] 48) and
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
    58
  set ("") and
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    59
  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
    60
  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
    61
  update2 ("update") and
63
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
    62
  tm_wf0 ("wf") and
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
    63
  tcopy_begin ("cbegin") and
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
    64
  tcopy_loop ("cloop") and
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
    65
  tcopy_end ("cend") and
75
97eaf7514988 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
    66
  step0 ("step") and
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    67
  tcontra ("contra") and
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    68
  code_tcontra ("code contra") and
75
97eaf7514988 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
    69
  steps0 ("steps") and
190
f1ecb4a68a54 renamed sete definition to adjust and old special case of adjust to adjust0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 187
diff changeset
    70
  adjust0 ("adjust") and
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
    71
  exponent ("_\<^bsup>_\<^esup>") and
37
c9b689bb4156 added new version of uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 36
diff changeset
    72
  tcopy ("copy") and 
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
    73
  tape_of ("\<langle>_\<rangle>") and 
232
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
    74
  tm_comp ("_ ; _") and 
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
    75
  DUMMY  ("\<^raw:\mbox{$\_\!\_\,$}>") and
284
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
    76
  inv_begin0 ("I\<^sub>0") and
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
    77
  inv_begin1 ("I\<^sub>1") and
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
    78
  inv_begin2 ("I\<^sub>2") and
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
    79
  inv_begin3 ("I\<^sub>3") and
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
    80
  inv_begin4 ("I\<^sub>4") and 
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
    81
  inv_begin ("I\<^bsub>cbegin\<^esub>") and
284
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
    82
  inv_loop1 ("J\<^sub>1") and
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
    83
  inv_loop0 ("J\<^sub>0") and
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
    84
  inv_end1 ("K\<^sub>1") and
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
    85
  inv_end0 ("K\<^sub>0") and
117
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
    86
  measure_begin_step ("M\<^bsub>cbegin\<^esub>") and
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
    87
  layout_of ("layout") and
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
    88
  findnth ("find'_nth") and
284
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
    89
  recf.id ("id\<^raw:\makebox[0mm]{\,\,\,\,>\<^sup>_\<^raw:}>\<^sub>_") and
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
    90
  Pr ("Pr\<^sup>_") and
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
    91
  Cn ("Cn\<^sup>_") and
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
    92
  Mn ("Mn\<^sup>_") and
230
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
    93
  rec_exec ("eval") and
202
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
    94
  tm_of_rec ("translate") and
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
    95
  listsum ("\<Sigma>")
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
    96
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
    97
 
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
    98
lemma inv_begin_print:
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
    99
  shows "s = 0 \<Longrightarrow> inv_begin n (s, tp) = inv_begin0 n tp" and
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   100
        "s = 1 \<Longrightarrow> inv_begin n (s, tp) = inv_begin1 n tp" and 
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   101
        "s = 2 \<Longrightarrow> inv_begin n (s, tp) = inv_begin2 n tp" and 
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   102
        "s = 3 \<Longrightarrow> inv_begin n (s, tp) = inv_begin3 n tp" and 
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   103
        "s = 4 \<Longrightarrow> inv_begin n (s, tp) = inv_begin4 n tp" and
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   104
        "s \<notin> {0,1,2,3,4} \<Longrightarrow> inv_begin n (s, l, r) = False"
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   105
apply(case_tac [!] tp)
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   106
by (auto)
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   107
230
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
   108
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   109
lemma inv1: 
230
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
   110
  shows "0 < (n::nat) \<Longrightarrow> Turing_Hoare.assert_imp (inv_begin0 n) (inv_loop1 n)"
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
   111
unfolding Turing_Hoare.assert_imp_def
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   112
unfolding inv_loop1.simps inv_begin0.simps
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   113
apply(auto)
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   114
apply(rule_tac x="1" in exI)
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   115
apply(auto simp add: replicate.simps)
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   116
done
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   117
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   118
lemma inv2: 
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   119
  shows "0 < n \<Longrightarrow> inv_loop0 n = inv_end1 n"
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   120
apply(rule ext)
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   121
apply(case_tac x)
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   122
apply(simp add: inv_end1.simps)
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   123
done
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   124
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   125
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
   126
lemma measure_begin_print:
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
   127
  shows "s = 2 \<Longrightarrow> measure_begin_step (s, l, r) = length r" and
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
   128
        "s = 3 \<Longrightarrow> measure_begin_step (s, l, r) = (if r = [] \<or> r = [Bk] then 1 else 0)" and
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
   129
        "s = 4 \<Longrightarrow> measure_begin_step (s, l, r) = length l" and 
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
   130
        "s \<notin> {2,3,4} \<Longrightarrow> measure_begin_step (s, l, r) = 0"
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
   131
by (simp_all)
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
   132
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   133
declare [[show_question_marks = false]]
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   134
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   135
lemma nats2tape:
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   136
  shows "<([]::nat list)> = []"
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   137
  and "<[n]> = <n>"
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   138
  and "ns \<noteq> [] \<Longrightarrow> <n#ns> = <(n::nat, ns)>"
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   139
  and "<(n, m)> = <n> @ [Bk] @ <m>"
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   140
  and "<[n, m]> = <(n, m)>"
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   141
  and "<n> = Oc \<up> (n + 1)" 
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   142
apply(auto simp add: tape_of_nat_pair tape_of_nl_abv tape_of_nat_abv tape_of_nat_list.simps)
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   143
apply(case_tac ns)
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   144
apply(auto simp add: tape_of_nat_pair tape_of_nat_abv)
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   145
done 
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   146
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   147
lemmas HR1 = 
284
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
   148
  Hoare_plus_halt[where ?S.0="R\<iota>" and ?A="p\<^sub>1" and B="p\<^sub>2"]
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   149
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   150
lemmas HR2 =
284
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
   151
  Hoare_plus_unhalt[where ?A="p\<^sub>1" and B="p\<^sub>2"]
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   152
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   153
lemma inv_begin01:
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   154
  assumes "n > 1"
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   155
  shows "inv_begin0 n (l, r) = (n > 1 \<and> (l, r) = (Oc \<up> (n - 2), [Oc, Oc, Bk, Oc]))"
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   156
using assms by auto                          
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   157
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   158
lemma inv_begin02:
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   159
  assumes "n = 1"
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   160
  shows "inv_begin0 n (l, r) = (n = 1 \<and> (l, r) = ([], [Bk, Oc, Bk, Oc]))"
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   161
using assms by auto
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   162
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
   163
115
653426ed4b38 started with abacus section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 114
diff changeset
   164
lemma layout:
653426ed4b38 started with abacus section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 114
diff changeset
   165
  shows "layout_of [] = []"
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   166
  and   "layout_of ((Inc R\<iota>)#is) = (2 * R\<iota> + 9)#(layout_of is)"
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   167
  and   "layout_of ((Dec R\<iota> l)#is) = (2 * R\<iota> + 16)#(layout_of is)"
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   168
  and   "layout_of ((Goto l)#is) = 1#(layout_of is)"
115
653426ed4b38 started with abacus section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 114
diff changeset
   169
by(auto simp add: layout_of.simps length_of.simps)
653426ed4b38 started with abacus section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 114
diff changeset
   170
190
f1ecb4a68a54 renamed sete definition to adjust and old special case of adjust to adjust0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 187
diff changeset
   171
f1ecb4a68a54 renamed sete definition to adjust and old special case of adjust to adjust0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 187
diff changeset
   172
lemma adjust_simps:
f1ecb4a68a54 renamed sete definition to adjust and old special case of adjust to adjust0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 187
diff changeset
   173
  shows "adjust0 p = map (\<lambda>(a, s). (a, if s = 0 then (Suc (length p div 2)) else s)) p"
f1ecb4a68a54 renamed sete definition to adjust and old special case of adjust to adjust0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 187
diff changeset
   174
by(simp add: adjust.simps)
f1ecb4a68a54 renamed sete definition to adjust and old special case of adjust to adjust0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 187
diff changeset
   175
185
2fad78b479a3 added clear-definition to paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 171
diff changeset
   176
fun clear :: "nat \<Rightarrow> nat \<Rightarrow> abc_prog"
2fad78b479a3 added clear-definition to paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 171
diff changeset
   177
  where
2fad78b479a3 added clear-definition to paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 171
diff changeset
   178
  "clear n e = [Dec n e, Goto 0]"
202
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   179
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   180
partial_function (tailrec)
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   181
  run
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   182
where
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   183
  "run p cf = (if (is_final cf) then cf else run p (step0 cf p))"
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   184
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   185
lemma numeral2:
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   186
  shows "11 = Suc 10"
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   187
  and   "12 = Suc 11"
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   188
  and   "13 = Suc 12"
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   189
  and   "14 = Suc 13"
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   190
  and   "15 = Suc 14"
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   191
apply(arith)+
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   192
done
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   193
239
ac3309722536 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 237
diff changeset
   194
lemma tm_wf_simps:
ac3309722536 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 237
diff changeset
   195
  "tm_wf0 p = (2 <=length p \<and> is_even(length p) \<and> (\<forall>(a,s) \<in> set p. s <= (length p) div 2))"
ac3309722536 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 237
diff changeset
   196
apply(simp add: tm_wf.simps)
202
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   197
done
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   198
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   199
(*>*)
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   200
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   201
section {* Introduction *}
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   202
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   203
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   204
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   205
text {*
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   206
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
   207
%\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
   208
%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
   209
%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
   210
%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
   211
%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
   212
%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
   213
%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
   214
%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
   215
%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
   216
%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
   217
233
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   218
%Suppose you want to mechanise a proof for whether a predicate @{term P}, 
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   219
%say, is decidable or not. Decidability of @{text P} usually
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   220
%amounts to showing whether \mbox{@{term "P \<or> \<not>P"}} holds. But this
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   221
%does \emph{not} work in 
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
   222
233
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   223
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   224
%Since Isabelle/HOL and other HOL theorem provers,
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   225
%are based on classical logic where the law of excluded
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   226
%middle ensures that \mbox{@{term "P \<or> \<not>P"}} is always provable no
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   227
%matter whether @{text P} is constructed by computable means. We hit on
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   228
%this limitation previously when we mechanised the correctness proofs
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   229
%of two algorithms \cite{UrbanCheneyBerghofer11,WuZhangUrban12}, but
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   230
%were unable to formalise arguments about decidability or undecidability.
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
   231
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
   232
%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
   233
%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
   234
%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
   235
%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
   236
233
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   237
\noindent
237
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
   238
We like to enable Isabelle/HOL users to reason about computability
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
   239
theory.  Reasoning about decidability of predicates, for example, is not as
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
   240
straightforward as one might think in Isabelle/HOL and other HOL
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
   241
theorem provers, since they are based on classical logic where the law
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
   242
of excluded middle ensures that \mbox{@{term "P \<or> \<not>P"}} is always
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
   243
provable no matter whether the predicate @{text P} is constructed by
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
   244
computable means.
236
6b6d71d14e75 polished the intro
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 234
diff changeset
   245
6b6d71d14e75 polished the intro
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 234
diff changeset
   246
Norrish formalised computability
233
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   247
theory in HOL4. He choose the $\lambda$-calculus as the starting point
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   248
for his formalisation because of its ``simplicity'' \cite[Page
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   249
297]{Norrish11}.  Part of his formalisation is a clever infrastructure
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   250
for reducing $\lambda$-terms. He also established the computational
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   251
equivalence between the $\lambda$-calculus and recursive functions.
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   252
Nevertheless he concluded that it would be appealing to have
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   253
formalisations for more operational models of computations, such as
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   254
Turing machines or register machines.  One reason is that many proofs
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   255
in the literature use them.  He noted however that \cite[Page
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   256
310]{Norrish11}:
8
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   257
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   258
\begin{quote}
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   259
\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
   260
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
   261
daunting prospect.''
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   262
\end{quote}
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   263
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   264
\noindent
33
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   265
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
   266
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
   267
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
   268
involved with this work, one has to understand that Turing machine
202
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   269
programs (similarly abacus programs) can be completely
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   270
\emph{unstructured}, behaving similar to Basic programs containing the
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   271
infamous goto \cite{Dijkstra68}. This precludes in the general case a
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   272
compositional Hoare-style reasoning about Turing programs.  We provide
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   273
such Hoare-rules for when it \emph{is} possible to reason in a
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   274
compositional manner (which is fortunately quite often), but also
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   275
tackle the more complicated case when we translate abacus programs
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   276
into Turing programs.  This reasoning about concrete Turing machine
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   277
programs is usually left out in the informal literature,
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
   278
e.g.~\cite{Boolos87}.
12
dd400b5797e1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 10
diff changeset
   279
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
   280
%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
   281
%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
   282
%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
   283
%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
   284
%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
   285
%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
   286
%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
   287
%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
   288
%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
   289
%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
   290
%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
   291
%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
   292
%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
   293
%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
   294
%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
   295
%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
   296
%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
   297
%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
   298
%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
   299
%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
   300
%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
   301
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   302
We are not the first who formalised Turing machines: we are aware of
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   303
the work by Asperti and Ricciotti \cite{AspertiRicciotti12}. They
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   304
describe a complete formalisation of Turing machines in the Matita
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   305
theorem prover, including a universal Turing machine. However, they do
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   306
\emph{not} formalise the undecidability of the halting problem since
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   307
their main focus is complexity, rather than computability theory. They
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   308
also report that the informal proofs from which they started are not
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   309
``sufficiently accurate to be directly usable as a guideline for
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   310
formalization'' \cite[Page 2]{AspertiRicciotti12}. For our
104
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   311
formalisation we follow mainly the proofs from the textbook by Boolos
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   312
et al \cite{Boolos87} and found that the description there is quite
88
904f9351ab94 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 87
diff changeset
   313
detailed. Some details are left out however: for example, constructing
130
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   314
the \emph{copy Turing machine} is left as an exercise to the
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   315
reader---a corresponding correctness proof is not mentioned at all; also \cite{Boolos87}
104
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   316
only shows how the universal Turing machine is constructed for Turing
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   317
machines computing unary functions. We had to figure out a way to
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   318
generalise this result to $n$-ary functions. Similarly, when compiling
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   319
recursive functions to abacus machines, the textbook again only shows
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   320
how it can be done for 2- and 3-ary functions, but in the
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   321
formalisation we need arbitrary functions. But the general ideas for
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   322
how to do this are clear enough in \cite{Boolos87}.
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
   323
%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
   324
%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
   325
%\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
   326
%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
   327
%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
   328
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   329
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
   330
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
   331
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
   332
23]{AspertiRicciotti12}:
10
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 9
diff changeset
   333
15
90bc8cccc218 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 13
diff changeset
   334
\begin{quote}\it
13
a7ec585d7f20 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   335
``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
   336
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
   337
annoying.'' 
a7ec585d7f20 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   338
\end{quote}
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   339
15
90bc8cccc218 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 13
diff changeset
   340
\noindent
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   341
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
   342
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
   343
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
   344
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
   345
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
   346
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
   347
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
   348
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
   349
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
   350
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
   351
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
   352
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
   353
\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
   354
\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
   355
turn recursive functions to abacus machines. The universal Turing
237
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
   356
machine can then be constructed by translating from a (universal) recursive function. 
233
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   357
The part of mechanising the translation of recursive function to register 
239
ac3309722536 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 237
diff changeset
   358
machines has already been done by Zammit in HOL4 \cite{Zammit99}, 
233
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   359
although his register machines use a slightly different instruction
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
   360
set than the one described in \cite{Boolos87}.
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   361
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   362
\smallskip
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   363
\noindent
141
4d7a568bd911 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 140
diff changeset
   364
{\bf Contributions:} We formalised in Isabelle/HOL Turing machines
4d7a568bd911 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 140
diff changeset
   365
following the description of Boolos et al \cite{Boolos87} where tapes
4d7a568bd911 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 140
diff changeset
   366
only have blank or occupied cells. We mechanise the undecidability of
4d7a568bd911 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 140
diff changeset
   367
the halting problem and prove the correctness of concrete Turing
4d7a568bd911 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 140
diff changeset
   368
machines that are needed in this proof; such correctness proofs are
4d7a568bd911 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 140
diff changeset
   369
left out in the informal literature.  For reasoning about Turing
4d7a568bd911 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 140
diff changeset
   370
machine programs we derive Hoare-rules.  We also construct the
4d7a568bd911 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 140
diff changeset
   371
universal Turing machine from \cite{Boolos87} by translating recursive
4d7a568bd911 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 140
diff changeset
   372
functions to abacus machines and abacus machines to Turing
158
6a584d61820f updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 157
diff changeset
   373
machines. This works essentially like a small, verified compiler 
6a584d61820f updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 157
diff changeset
   374
from recursive functions to Turing machine programs.
6a584d61820f updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 157
diff changeset
   375
When formalising the universal Turing machine,
159
b4b789a59086 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 158
diff changeset
   376
we stumbled in \cite{Boolos87} upon an inconsistent use of the definition of
158
6a584d61820f updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 157
diff changeset
   377
what partial function a Turing machine calculates. 
6a584d61820f updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 157
diff changeset
   378
152
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
   379
%Since we have set up in
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
   380
%Isabelle/HOL a very general computability model and undecidability
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
   381
%result, we are able to formalise other results: we describe a proof of
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
   382
%the computational equivalence of single-sided Turing machines, which
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
   383
%is not given in \cite{Boolos87}, but needed for example for
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
   384
%formalising the undecidability proof of Wang's tiling problem
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
   385
%\cite{Robinson71}.  %We are not aware of any other %formalisation of a
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
   386
%substantial undecidability problem.
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   387
*}
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   388
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   389
section {* Turing Machines *}
9
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
   390
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
   391
text {* \noindent
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   392
  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
   393
  ``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
   394
  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
   395
  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
   396
  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
   397
  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
   398
  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
   399
  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
   400
  the notation @{term "Bk \<up> n"} (similarly @{term "Oc \<up> n"}) for lists
88
904f9351ab94 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 87
diff changeset
   401
  composed of @{term n} elements of @{term Bk}s.  We also have the
89
c67e8ed4c865 updated uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 88
diff changeset
   402
  convention that the head, abbreviated @{term hd}, of the right list
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   403
  is the cell on which the head of the Turing machine currently
130
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   404
  scans. 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
   405
  %
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   406
  \begin{center}
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   407
  \begin{tikzpicture}[scale=0.9]
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   408
  \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
   409
  \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
   410
  \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
   411
  \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
   412
  \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
   413
  \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
   414
  \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
   415
  \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
   416
  \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
   417
  \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
   418
  \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
   419
  \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
   420
  \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
   421
  \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
   422
  \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
   423
  \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
   424
  \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
   425
  \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
   426
  \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
   427
  \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
   428
  \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
   429
  \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
   430
  \end{tikzpicture}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   431
  \end{center}
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   432
  
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   433
  \noindent
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   434
  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
   435
  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
   436
  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
   437
  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
   438
  the Turing machine can perform:
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   439
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   440
  \begin{center}
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   441
  \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
   442
  @{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
   443
  & $\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
   444
  & $\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
   445
  & $\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
   446
  & $\mid$ & @{term Nop} & (do-nothing operation)\\
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   447
  \end{tabular}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   448
  \end{center}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   449
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   450
  \noindent
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
   451
  We slightly deviate
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   452
  from the presentation in \cite{Boolos87} (and also \cite{AspertiRicciotti12}) 
91
2068654bdf54 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 90
diff changeset
   453
  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
   454
  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
   455
  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
   456
  following tape updating function:
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   457
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   458
  \begin{center}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   459
  \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
   460
  @{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
   461
  @{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
   462
  @{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
   463
  @{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
   464
  @{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
   465
  \end{tabular}
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   466
  \end{center}
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   467
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   468
  \noindent
89
c67e8ed4c865 updated uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 88
diff changeset
   469
  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
   470
  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
   471
  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
   472
  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
   473
  such that @{term "tl [] == []"} holds. The third clause 
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   474
  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
   475
  to test if the left-list @{term l} is empty; if yes, then we just prepend a 
89
c67e8ed4c865 updated uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 88
diff changeset
   476
  blank cell to the right list; otherwise we have to remove the
c67e8ed4c865 updated uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 88
diff changeset
   477
  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
   478
  in the fourth clause for a right move action. The @{term Nop} operation
232
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   479
  leaves the tape unchanged.
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   480
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
   481
  %Note that our treatment of the tape is rather ``unsymmetric''---we
89
c67e8ed4c865 updated uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 88
diff changeset
   482
  %have the convention that the head of the right list is where the
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
   483
  %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
   484
  %\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
   485
  %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
   486
  %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
   487
  %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
   488
  %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
   489
  %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
   490
  %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
   491
  %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
   492
  %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
   493
  %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
   494
  %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
   495
  %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
   496
  %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
   497
  %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
   498
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   499
  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
   500
  %Given
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   501
  %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
   502
  %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
   503
  %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
   504
  %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
   505
  %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
   506
  %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
   507
  %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
   508
  %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
   509
  %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
   510
  %quite cumbersome to reason about. 
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   511
  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
   512
  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
   513
  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
   514
  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
   515
  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
   516
  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
   517
63
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
   518
  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
   519
  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
   520
  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
   521
  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
   522
  %
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   523
  \begin{equation}
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   524
  \begin{tikzpicture}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   525
  \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
   526
  \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
   527
  {$\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
   528
  = starting state\end{tabular}}}$};
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   529
  
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   530
  \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
   531
  \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
   532
  \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
   533
  \end{tikzpicture}
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   534
  \label{dither}
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   535
  \end{equation}
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   536
  %
29
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 28
diff changeset
   537
  \noindent
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   538
  the reader can see we have organised our Turing machine programs so
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   539
  that segments of two pairs 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
   540
  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
   541
  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
   542
  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
   543
  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
   544
  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
   545
  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
   546
  ``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
   547
  0}-state, but it will always perform a @{term Nop}-operation and
232
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   548
  remain in the @{text 0}-state.  
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   549
  We have chosen a very concrete
159
b4b789a59086 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 158
diff changeset
   550
  representation for Turing machine programs, because when constructing a universal
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   551
  Turing machine, we need to define a coding function for programs.
232
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   552
  %This can be directly done for our programs-as-lists, but is
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   553
  %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
   554
29
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 28
diff changeset
   555
  Given a program @{term p}, a state
159
b4b789a59086 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 158
diff changeset
   556
  and the cell being scanned by the head, we need to fetch
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   557
  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
   558
  the function @{term fetch}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   559
 
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   560
  \begin{equation}\label{fetch}
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   561
  \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
   562
  \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
   563
  @{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
   564
  \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
   565
  @{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
   566
  \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
   567
  \end{tabular}}
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   568
  \end{equation}
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   569
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   570
  \noindent
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   571
  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
   572
  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
   573
  returns the default action @{term Nop} and the default state @{text 0} 
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   574
  (@{term None}-case). We often have to restrict Turing machine programs 
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   575
  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
   576
  satisfies the following three properties:
33
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   577
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   578
  \begin{center}
239
ac3309722536 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 237
diff changeset
   579
  @{thm tm_wf_simps[where p="p", THEN eq_reflection]}
33
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   580
  \end{center}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   581
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   582
  \noindent
75
97eaf7514988 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   583
  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
   584
  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
   585
  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
   586
  the program or being the @{text 0}-state.
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   587
232
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   588
 
72
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   589
  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
   590
  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
   591
  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
   592
  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
   593
  "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
   594
  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
   595
  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
   596
  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
   597
  the function @{term step}
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   598
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   599
  \begin{center}
63
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
   600
  \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
   601
  @{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
   602
  & & @{text "in (s', update (l, r) a)"}
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   603
  \end{tabular}
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   604
  \end{center}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   605
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   606
  \noindent
104
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   607
  where @{term "read r"} returns the head of the list @{text r}, or if
237
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
   608
  @{text r} is empty it returns @{term Bk}.  
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
   609
  We lift this definition to an evaluation function that performs 
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
   610
  exactly @{text n} steps:
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   611
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   612
  \begin{center}
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   613
  \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   614
  @{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
   615
  @{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
   616
  \end{tabular}
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   617
  \end{center}
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   618
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   619
  \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
   620
  \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
   621
  case a Turing program takes according to the usual textbook
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
   622
  definition, say \cite{Boolos87}, less than @{text n} steps before it
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   623
  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
   624
  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
   625
  final state) and remains there performing @{text Nop}-actions until
91
2068654bdf54 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 90
diff changeset
   626
  @{text n} is reached. 
232
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   627
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   628
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   629
  We often need to restrict tapes to be in standard form, which means 
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   630
  the left list of the tape is either empty or only contains @{text "Bk"}s, and 
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   631
  the right list contains some ``clusters'' of @{text "Oc"}s separated by single 
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   632
  blanks. To make this formal we define the following overloaded function
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   633
  encoding natural numbers into lists of @{term "Oc"}s and @{term Bk}s.
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   634
  % 
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   635
  \begin{equation}
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   636
  \mbox{\begin{tabular}[t]{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}}
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   637
  @{thm (lhs) nats2tape(6)} & @{text "\<equiv>"} & @{thm (rhs) nats2tape(6)}\\
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   638
  @{thm (lhs) nats2tape(4)} & @{text "\<equiv>"} & @{thm (rhs) nats2tape(4)}\\
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   639
  \end{tabular}\hspace{6mm}
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   640
  \begin{tabular}[t]{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}}
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   641
  @{thm (lhs) nats2tape(1)} & @{text "\<equiv>"} & @{thm (rhs) nats2tape(1)}\\
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   642
  @{thm (lhs) nats2tape(2)} & @{text "\<equiv>"} & @{thm (rhs) nats2tape(2)}\\
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   643
  @{thm (lhs) nats2tape(3)} & @{text "\<equiv>"} & @{thm (rhs) nats2tape(3)}
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   644
  \end{tabular}}\label{standard}
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   645
  \end{equation}
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   646
  %
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   647
  \noindent
284
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
   648
  A \emph{standard tape} is then of the form @{text "(Bk\<^sup>k,\<langle>[n\<^sub>1,...,n\<^sub>m]\<rangle> @ Bk\<^sup>l)"} for some @{text k}, 
232
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   649
  @{text l} 
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   650
  and @{text "n\<^bsub>1...m\<^esub>"}. Note that the head in a standard tape ``points'' to the 
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   651
  leftmost @{term "Oc"} on the tape. Note also that the natural number @{text 0} 
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   652
  is represented by a single filled cell on a standard tape, @{text 1} by two filled cells and so on.
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   653
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   654
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   655
 We need to be able to sequentially compose Turing machine programs. Given our
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   656
  concrete representation, this is relatively straightforward, if
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   657
  slightly fiddly. We use the following two auxiliary functions:
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   658
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   659
  \begin{center}
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   660
  \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}}
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   661
  @{thm (lhs) shift.simps} @{text "\<equiv>"} @{thm (rhs) shift.simps}\\
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   662
  @{thm (lhs) adjust_simps} @{text "\<equiv>"} @{thm (rhs) adjust_simps}\\
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   663
  \end{tabular}
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   664
  \end{center}
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   665
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   666
  \noindent
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   667
  The first adds @{text n} to all states, except the @{text 0}-state,
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   668
  thus moving all ``regular'' states to the segment starting at @{text
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   669
  n}; the second adds @{term "Suc(length p div 2)"} to the @{text
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   670
  0}-state, thus redirecting all references to the ``halting state''
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   671
  to the first state after the program @{text p}.  With these two
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   672
  functions in place, we can define the \emph{sequential composition}
284
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
   673
  of two Turing machine programs @{text "p\<^sub>1"} and @{text "p\<^sub>2"} as
232
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   674
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   675
  \begin{center}
284
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
   676
  @{thm tm_comp.simps[where ?p1.0="p\<^sub>1" and ?p2.0="p\<^sub>2", THEN eq_reflection]}
232
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   677
  \end{center}
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   678
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   679
  \noindent
284
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
   680
  %This means @{text "p\<^sub>1"} is executed first. Whenever it originally
232
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   681
  %transitioned to the @{text 0}-state, it will in the composed program transition to the starting
284
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
   682
  %state of @{text "p\<^sub>2"} instead. All the states of @{text "p\<^sub>2"}
232
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   683
  %have been shifted in order to make sure that the states of the composed 
284
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
   684
  %program @{text "p\<^sub>1 \<oplus> p\<^sub>2"} still only ``occupy'' 
232
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   685
  %an initial segment of the natural numbers.
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   686
  
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   687
  \begin{figure}[t]
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   688
  \begin{center}
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   689
  \begin{tabular}{@ {}c@ {\hspace{3mm}}c@ {\hspace{3mm}}c}
87
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   690
  \begin{tabular}[t]{@ {}l@ {}}
92
b9d0dd18c81e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 91
diff changeset
   691
  @{thm (lhs) tcopy_begin_def} @{text "\<equiv>"}\\
87
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   692
  \hspace{2mm}@{text "["}@{text "(W\<^bsub>Bk\<^esub>, 0), (R, 2), (R, 3),"}\\
217
ebe8fd1fb26f small typo in the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 214
diff changeset
   693
  \hspace{2mm}\phantom{@{text "["}}@{text "(R, 2), (W\<^bsub>Oc\<^esub>, 3), (L, 4),"}\\
87
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   694
  \hspace{2mm}\phantom{@{text "["}}@{text "(L, 4), (L, 0)"}@{text "]"} 
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   695
  \end{tabular}
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   696
  &
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   697
  \begin{tabular}[t]{@ {}l@ {}}
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   698
  @{thm (lhs) tcopy_loop_def} @{text "\<equiv>"}\\
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   699
  \hspace{2mm}@{text "["}@{text "(R, 0), (R, 2), (R, 3),"}\\
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   700
  \hspace{2mm}\phantom{@{text "["}}@{text "(W\<^bsub>Bk\<^esub>, 2), (R, 3), (R, 4),"}\\
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   701
  \hspace{2mm}\phantom{@{text "["}}@{text "(W\<^bsub>Oc\<^esub>, 5), (R, 4), (L, 6),"}\\
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   702
  \hspace{2mm}\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
   703
  \end{tabular}
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   704
  &
87
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   705
  \begin{tabular}[t]{@ {}l@ {}}
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   706
  @{thm (lhs) tcopy_end_def} @{text "\<equiv>"}\\
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   707
  \hspace{2mm}@{text "["}@{text "(L, 0), (R, 2), (W\<^bsub>Oc\<^esub>, 3),"}\\
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   708
  \hspace{2mm}\phantom{@{text "["}}@{text "(L, 4), (R, 2), (R, 2),"}\\
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   709
  \hspace{2mm}\phantom{@{text "["}}@{text "(L, 5), (W\<^bsub>Bk\<^esub>, 4), (R, 0),"}\\
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   710
  \hspace{2mm}\phantom{@{text "["}}@{text "(L, 5)"}@{text "]"} 
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   711
  \end{tabular}
87
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   712
  \end{tabular}\\[2mm]
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   713
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   714
  \begin{tikzpicture}[scale=0.7]
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   715
  \node [anchor=base] at (2.2,0.1) {\small$\Rightarrow$};
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   716
  \node [anchor=base] at (5.6,0.1) {\small$\Rightarrow$};
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   717
  \node [anchor=base] at (10.5,0.1) {\small$\Rightarrow$};
98
860f05037c36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 97
diff changeset
   718
  \node [anchor=base] at (2.2,-0.6) {\small$\overbrace{@{term "tcopy_begin"}}^{}$};
88
904f9351ab94 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 87
diff changeset
   719
  \node [anchor=base] at (5.6,-0.6) {\small$\overbrace{@{term "tcopy_loop"}}^{}$};
904f9351ab94 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 87
diff changeset
   720
  \node [anchor=base] at (10.5,-0.6) {\small$\overbrace{@{term "tcopy_end"}}^{}$};
87
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   721
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   722
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   723
  \begin{scope}[shift={(0.5,0)}]
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   724
  \draw[very thick] (-0.25,0)   -- ( 1.25,0);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   725
  \draw[very thick] (-0.25,0.5) -- ( 1.25,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   726
  \draw[very thick] (-0.25,0)   -- (-0.25,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   727
  \draw[very thick] ( 0.25,0)   -- ( 0.25,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   728
  \draw[very thick] ( 0.75,0)   -- ( 0.75,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   729
  \draw[very thick] ( 1.25,0)   -- ( 1.25,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   730
  \draw[rounded corners=1mm] (-0.35,-0.1) rectangle (0.35,0.6);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   731
  \draw[fill]     (-0.15,0.1) rectangle (0.15,0.4);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   732
  \draw[fill]     ( 0.35,0.1) rectangle (0.65,0.4);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   733
  \draw[fill]     ( 0.85,0.1) rectangle (1.15,0.4);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   734
  \end{scope}
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   735
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   736
  \begin{scope}[shift={(2.9,0)}]
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   737
  \draw[very thick] (-0.25,0)   -- ( 2.25,0);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   738
  \draw[very thick] (-0.25,0.5) -- ( 2.25,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   739
  \draw[very thick] (-0.25,0)   -- (-0.25,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   740
  \draw[very thick] ( 0.25,0)   -- ( 0.25,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   741
  \draw[very thick] ( 0.75,0)   -- ( 0.75,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   742
  \draw[very thick] ( 1.25,0)   -- ( 1.25,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   743
  \draw[very thick] ( 1.75,0)   -- ( 1.75,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   744
  \draw[very thick] ( 2.25,0)   -- ( 2.25,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   745
  \draw[rounded corners=1mm] (0.15,-0.1) rectangle (0.85,0.6);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   746
  \draw[fill]     (-0.15,0.1) rectangle (0.15,0.4);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   747
  \draw[fill]     ( 0.35,0.1) rectangle (0.65,0.4);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   748
  \draw[fill]     ( 0.85,0.1) rectangle (1.15,0.4);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   749
  \draw[fill]     ( 1.85,0.1) rectangle (2.15,0.4);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   750
  \end{scope}
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   751
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   752
  \begin{scope}[shift={(6.8,0)}]
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   753
  \draw[very thick] (-0.75,0)   -- ( 3.25,0);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   754
  \draw[very thick] (-0.75,0.5) -- ( 3.25,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   755
  \draw[very thick] (-0.75,0)   -- (-0.75,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   756
  \draw[very thick] (-0.25,0)   -- (-0.25,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   757
  \draw[very thick] ( 0.25,0)   -- ( 0.25,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   758
  \draw[very thick] ( 0.75,0)   -- ( 0.75,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   759
  \draw[very thick] ( 1.25,0)   -- ( 1.25,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   760
  \draw[very thick] ( 1.75,0)   -- ( 1.75,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   761
  \draw[very thick] ( 2.25,0)   -- ( 2.25,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   762
  \draw[very thick] ( 2.75,0)   -- ( 2.75,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   763
  \draw[very thick] ( 3.25,0)   -- ( 3.25,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   764
  \draw[rounded corners=1mm] (-0.35,-0.1) rectangle (0.35,0.6);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   765
  \draw[fill]     (-0.15,0.1) rectangle (0.15,0.4);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   766
  \draw[fill]     ( 2.35,0.1) rectangle (2.65,0.4);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   767
  \draw[fill]     ( 2.85,0.1) rectangle (3.15,0.4);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   768
  \draw[fill]     ( 1.85,0.1) rectangle (2.15,0.4);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   769
  \end{scope}
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   770
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   771
  \begin{scope}[shift={(11.7,0)}]
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   772
  \draw[very thick] (-0.75,0)   -- ( 3.25,0);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   773
  \draw[very thick] (-0.75,0.5) -- ( 3.25,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   774
  \draw[very thick] (-0.75,0)   -- (-0.75,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   775
  \draw[very thick] (-0.25,0)   -- (-0.25,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   776
  \draw[very thick] ( 0.25,0)   -- ( 0.25,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   777
  \draw[very thick] ( 0.75,0)   -- ( 0.75,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   778
  \draw[very thick] ( 1.25,0)   -- ( 1.25,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   779
  \draw[very thick] ( 1.75,0)   -- ( 1.75,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   780
  \draw[very thick] ( 2.25,0)   -- ( 2.25,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   781
  \draw[very thick] ( 2.75,0)   -- ( 2.75,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   782
  \draw[very thick] ( 3.25,0)   -- ( 3.25,0.5);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   783
  \draw[rounded corners=1mm] (-0.35,-0.1) rectangle (0.35,0.6);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   784
  \draw[fill]     (-0.15,0.1) rectangle (0.15,0.4);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   785
  \draw[fill]     ( 2.35,0.1) rectangle (2.65,0.4);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   786
  \draw[fill]     ( 2.85,0.1) rectangle (3.15,0.4);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   787
  \draw[fill]     ( 1.85,0.1) rectangle (2.15,0.4);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   788
  \draw[fill]     ( 0.35,0.1) rectangle (0.65,0.4);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   789
  \draw[fill]     ( 0.85,0.1) rectangle (1.15,0.4);
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   790
  \end{scope}
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   791
  \end{tikzpicture}\\[-8mm]\mbox{} 
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   792
  \end{center}
94
aeaf1374dc67 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 93
diff changeset
   793
  \caption{The three components of the \emph{copy Turing machine} (above). If started 
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   794
  (below) with the tape @{term "([], <(2::nat)>)"} the first machine appends @{term "[Bk, Oc]"} at 
87
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   795
  the end of the right tape; the second then ``moves'' all @{term Oc}s except the 
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   796
  first from the beginning of the tape to the end; the third ``refills'' the original 
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   797
  block of @{term "Oc"}s. The resulting tape is @{term "([Bk], <(2::nat, 2::nat)>)"}.}
87
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   798
  \label{copy}
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   799
  \end{figure}
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   800
87
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   801
232
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
   802
79
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   803
104
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   804
  Before we can prove the undecidability of the halting problem for
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   805
  our Turing machines working on standard tapes, we have to analyse
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   806
  two concrete Turing machine programs and establish that they are
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   807
  correct---that means they are ``doing what they are supposed to be
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   808
  doing''.  Such correctness proofs are usually left out in the
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   809
  informal literature, for example \cite{Boolos87}.  The first program
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   810
  we need to prove correct is the @{term dither} program shown in
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   811
  \eqref{dither} and the second 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
   812
91
2068654bdf54 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 90
diff changeset
   813
  \begin{equation}
2068654bdf54 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 90
diff changeset
   814
  \mbox{\begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}}
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   815
  @{thm (lhs) tcopy_def} & @{text "\<equiv>"} & @{thm (rhs) tcopy_def}
91
2068654bdf54 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 90
diff changeset
   816
  \end{tabular}}\label{tcopy}
2068654bdf54 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 90
diff changeset
   817
  \end{equation}
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   818
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   819
  \noindent
102
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   820
  whose three components are given in Figure~\ref{copy}. For our
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   821
  correctness proofs, we introduce the notion of total correctness
230
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
   822
  defined in terms of \emph{Hoare-triples}, written @{term "{P} (p::tprog0)
102
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   823
  {Q}"}.  They implement the idea that a program @{term
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   824
  p} started in state @{term "1::nat"} with a tape satisfying @{term
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   825
  P} will after some @{text n} steps halt (have transitioned into the
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   826
  halting state) with a tape satisfying @{term Q}. This idea is very
104
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   827
  similar to the notion of \emph{realisability} in \cite{AspertiRicciotti12}. We
230
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
   828
  also have \emph{Hoare-pairs} of the form @{term "{P} (p::tprog0) \<up>"}
102
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   829
  implementing the case that a program @{term p} started with a tape
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   830
  satisfying @{term P} will loop (never transition into the halting
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   831
  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
   832
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   833
  \begin{center}
76
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   834
  \begin{tabular}{@ {}c@ {\hspace{4mm}}c@ {}}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   835
  \begin{tabular}[t]{@ {}l@ {}}
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   836
  \colorbox{mygrey}{@{thm (lhs) Hoare_halt_def}} @{text "\<equiv>"}\\[1mm]
239
ac3309722536 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 237
diff changeset
   837
  \hspace{5mm}@{text "\<forall>"}@{term "tp"}.\\
ac3309722536 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 237
diff changeset
   838
  \hspace{7mm}if @{term "P tp"} holds then\\
ac3309722536 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 237
diff changeset
   839
  \hspace{7mm}@{text "\<exists>"}@{term n}. such that\\
ac3309722536 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 237
diff changeset
   840
  \hspace{7mm}@{text "is_final (steps (1, tp) p n)"} \hspace{1mm}@{text "\<and>"}\\
ac3309722536 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 237
diff changeset
   841
  \hspace{7mm}@{text "Q holds_for (steps (1, tp) p n)"}
76
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   842
  \end{tabular} &
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   843
  \begin{tabular}[t]{@ {}l@ {}}
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   844
  \colorbox{mygrey}{@{thm (lhs) Hoare_unhalt_def}} @{text "\<equiv>"}\\[1mm]
239
ac3309722536 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 237
diff changeset
   845
  \hspace{5mm}@{text "\<forall>"}@{term "tp"}.\\ 
ac3309722536 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 237
diff changeset
   846
  \hspace{7mm}if @{term "P tp"} holds then\\
ac3309722536 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 237
diff changeset
   847
  \hspace{7mm}@{text "\<forall>"}@{term n}. @{text "\<not> is_final (steps (1, tp) p n)"}
76
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   848
  \end{tabular}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   849
  \end{tabular}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   850
  \end{center}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   851
  
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   852
  \noindent
102
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   853
  For our Hoare-triples we can easily prove the following Hoare-consequence rule
99
fe7a257bdff4 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 98
diff changeset
   854
fe7a257bdff4 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 98
diff changeset
   855
  \begin{equation}
fe7a257bdff4 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 98
diff changeset
   856
  @{thm[mode=Rule] Hoare_consequence}
fe7a257bdff4 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 98
diff changeset
   857
  \end{equation}
fe7a257bdff4 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 98
diff changeset
   858
104
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   859
  \noindent
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   860
  where
230
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
   861
  @{term "Turing_Hoare.assert_imp P' P"} stands for the fact that for all tapes @{term "tp"},
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
   862
  @{term "P' tp"} implies @{term "P tp"} (similarly for @{text "Q"} and @{text "Q'"}).
99
fe7a257bdff4 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 98
diff changeset
   863
90
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   864
  Like Asperti and Ricciotti with their notion of realisability, we
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   865
  have set up our Hoare-rules so that we can deal explicitly
130
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   866
  with total correctness and non-termination, rather than have
90
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   867
  notions for partial correctness and termination. Although the latter
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   868
  would allow us to reason more uniformly (only using Hoare-triples),
96
bd320b5365e2 updated uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 94
diff changeset
   869
  we prefer our definitions because we can derive below some simple
90
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   870
  Hoare-rules for sequentially composed Turing programs.  In this way
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   871
  we can reason about the correctness of @{term "tcopy_begin"}, for
90
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   872
  example, completely separately from @{term "tcopy_loop"} and @{term
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   873
  "tcopy_end"}.
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   874
130
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   875
  It is relatively straightforward to prove that the Turing program 
90
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   876
  @{term "dither"} shown in \eqref{dither} is correct. This program
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   877
  should be the ``identity'' when started with a standard tape representing 
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   878
  @{text "1"} but loops when started with the @{text 0}-representation instead, as pictured 
94
aeaf1374dc67 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 93
diff changeset
   879
  below.
90
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   880
76
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   881
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   882
  \begin{center}
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   883
  \begin{tabular}{l@ {\hspace{3mm}}lcl}
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   884
  & \multicolumn{1}{c}{start tape}\\[1mm]
90
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   885
  \raisebox{2mm}{halting case:} &
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   886
  \begin{tikzpicture}[scale=0.8]
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   887
  \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
   888
  \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
   889
  \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
   890
  \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
   891
  \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
   892
  \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
   893
  \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
   894
  \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
   895
  \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
   896
  \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
   897
  \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
   898
  \end{tikzpicture} 
90
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   899
  & \raisebox{2mm}{$\;\;\large\Rightarrow\;\;$} &
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   900
  \begin{tikzpicture}[scale=0.8]
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   901
  \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
   902
  \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
   903
  \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
   904
  \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
   905
  \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
   906
  \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
   907
  \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
   908
  \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
   909
  \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
   910
  \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
   911
  \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
   912
  \end{tikzpicture}\\
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   913
90
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   914
  \raisebox{2mm}{non-halting case:} &
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   915
  \begin{tikzpicture}[scale=0.8]
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   916
  \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
   917
  \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
   918
  \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
   919
  \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
   920
  \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
   921
  \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
   922
  \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
   923
  \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
   924
  \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
   925
  \end{tikzpicture} 
90
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   926
  & \raisebox{2mm}{$\;\;\large\Rightarrow\;\;$} &
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   927
  \raisebox{2mm}{loops}
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   928
  \end{tabular}
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   929
  \end{center}
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   930
90
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   931
  \noindent
239
ac3309722536 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 237
diff changeset
   932
  We can prove the following two Hoare-statements:
90
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   933
 
76
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   934
  \begin{center}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   935
  \begin{tabular}{l}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   936
  @{thm dither_halts}\\
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   937
  @{thm dither_loops}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   938
  \end{tabular}
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   939
  \end{center}
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   940
77
04e5850818c8 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   941
  \noindent
102
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   942
  The first is by a simple calculation. The second is by an induction on the
91
2068654bdf54 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 90
diff changeset
   943
  number of steps we can perform starting from the input tape.
75
97eaf7514988 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   944
96
bd320b5365e2 updated uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 94
diff changeset
   945
  The program @{term tcopy} defined in \eqref{tcopy} has 15 states;
104
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   946
  its purpose is to produce the standard tape @{term "(Bks, <(n,
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   947
  n::nat)>)"} when started with @{term "(Bks, <(n::nat)>)"}, that is
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   948
  making a copy of a value @{term n} on the tape.  Reasoning about this program
96
bd320b5365e2 updated uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 94
diff changeset
   949
  is substantially harder than about @{term dither}. To ease the
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   950
  burden, we derive the following two Hoare-rules for sequentially
96
bd320b5365e2 updated uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 94
diff changeset
   951
  composed programs.
75
97eaf7514988 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   952
91
2068654bdf54 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 90
diff changeset
   953
  \begin{center}
94
aeaf1374dc67 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 93
diff changeset
   954
  \begin{tabular}{@ {\hspace{-10mm}}c@ {\hspace{14mm}}c@ {}}
aeaf1374dc67 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 93
diff changeset
   955
  $\inferrule*[Right=@{thm (prem 3) HR1}]
aeaf1374dc67 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 93
diff changeset
   956
  {@{thm (prem 1) HR1} \\ @{thm (prem 2) HR1}}
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   957
  {@{thm (concl) HR1}}
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   958
  $ &
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   959
  $
94
aeaf1374dc67 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 93
diff changeset
   960
  \inferrule*[Right=@{thm (prem 3) HR2}]
aeaf1374dc67 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 93
diff changeset
   961
  {@{thm (prem 1) HR2} \\ @{thm (prem 2) HR2}}
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   962
  {@{thm (concl) HR2}}
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   963
  $
91
2068654bdf54 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 90
diff changeset
   964
  \end{tabular}
2068654bdf54 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 90
diff changeset
   965
  \end{center}
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   966
91
2068654bdf54 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 90
diff changeset
   967
  \noindent
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   968
  The first corresponds to the usual Hoare-rule for composition of two
102
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   969
  terminating programs. The second rule gives the conditions for when
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   970
  the first program terminates generating a tape for which the second
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   971
  program loops.  The side-conditions about @{thm (prem 3) HR2} are
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   972
  needed in order to ensure that the redirection of the halting and
284
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
   973
  initial state in @{term "p\<^sub>1"} and @{term "p\<^sub>2"}, respectively, match
107
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
   974
  up correctly.  These Hoare-rules allow us to prove the correctness
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
   975
  of @{term tcopy} by considering the correctness of the components
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
   976
  @{term "tcopy_begin"}, @{term "tcopy_loop"} and @{term "tcopy_end"}
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
   977
  in isolation. This simplifies the reasoning considerably, for
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
   978
  example when designing decreasing measures for proving the termination
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
   979
  of the programs.  We will show the details for the program @{term
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
   980
  "tcopy_begin"}. For the two other programs we refer the reader to
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
   981
  our formalisation.
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   982
107
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
   983
  Given the invariants @{term "inv_begin0"},\ldots,
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
   984
  @{term "inv_begin4"} shown in Figure~\ref{invbegin}, which
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
   985
  correspond to each state of @{term tcopy_begin}, we define the
103
294576baaeed updated theories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 102
diff changeset
   986
  following invariant for the whole @{term tcopy_begin} program:
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   987
104
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   988
  \begin{figure}[t]
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   989
  \begin{center}
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
   990
  \begin{tabular}{@ {}lcl@ {\hspace{-0.5cm}}l@ {}}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
   991
  \hline
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   992
  @{thm (lhs) inv_begin1.simps} & @{text "\<equiv>"} & @{thm (rhs) inv_begin1.simps} & (starting state)\\
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   993
  @{thm (lhs) inv_begin2.simps} & @{text "\<equiv>"} & @{thm (rhs) inv_begin2.simps}\\
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   994
  @{thm (lhs) inv_begin3.simps} & @{text "\<equiv>"} & @{thm (rhs) inv_begin3.simps}\\
100
dfe852aacae6 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 99
diff changeset
   995
  @{thm (lhs) inv_begin4.simps} & @{text "\<equiv>"} & @{thm (rhs) inv_begin4.simps}\\
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
   996
  @{thm (lhs) inv_begin0.simps} & @{text "\<equiv>"} & @{thm (rhs) inv_begin01} @{text "\<or>"}& (halting state)\\
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
   997
                                &             & @{thm (rhs) inv_begin02}\smallskip \\
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
   998
   \hline
100
dfe852aacae6 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 99
diff changeset
   999
  @{thm (lhs) inv_loop1.simps} & @{text "\<equiv>"} & @{thm (rhs) inv_loop1_loop.simps} @{text "\<or>"}\\
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1000
                               &             & @{thm (rhs) inv_loop1_exit.simps} & (starting state)\\
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1001
  @{thm (lhs) inv_loop0.simps} & @{text "\<equiv>"} & @{thm (rhs) inv_loop0.simps}& (halting state)\smallskip\\
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1002
   \hline
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1003
  @{thm (lhs) inv_end1.simps} & @{text "\<equiv>"} & @{thm (rhs) inv_end1.simps} & (starting state)\\
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1004
  @{thm (lhs) inv_end0.simps} & @{text "\<equiv>"} & @{thm (rhs) inv_end0.simps} & (halting state)\smallskip\\
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1005
  \hline
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
  1006
  \end{tabular}
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
  1007
  \end{center}
106
c3155e9e4f63 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 105
diff changeset
  1008
  \caption{The invariants @{term inv_begin0},\ldots,@{term inv_begin4} are for the states of 
c3155e9e4f63 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 105
diff changeset
  1009
  @{term tcopy_begin}. Below, the invariants only for the starting and halting states of
232
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
  1010
  @{term tcopy_loop} and @{term tcopy_end} are shown. In each invariant, the parameter 
106
c3155e9e4f63 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 105
diff changeset
  1011
  @{term n} stands for the number
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1012
  of @{term Oc}s with which the Turing machine is started.}\label{invbegin}
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
  1013
  \end{figure}
96
bd320b5365e2 updated uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 94
diff changeset
  1014
bd320b5365e2 updated uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 94
diff changeset
  1015
  \begin{center}
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
  1016
  \begin{tabular}{rcl}
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
  1017
  @{thm (lhs) inv_begin.simps} & @{text "\<equiv>"} & 
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
  1018
  @{text "if"} @{thm (prem 1) inv_begin_print(1)} @{text then} @{thm (rhs) inv_begin_print(1)}\\
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
  1019
  & & @{text else} @{text "if"} @{thm (prem 1) inv_begin_print(2)} @{text then} @{thm (rhs) inv_begin_print(2)} \\
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
  1020
 & & @{text else} @{text "if"} @{thm (prem 1) inv_begin_print(3)} @{text then} @{thm (rhs) inv_begin_print(3)}\\
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
  1021
  & & @{text else} @{text "if"} @{thm (prem 1) inv_begin_print(4)} @{text then} @{thm (rhs) inv_begin_print(4)}\\
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
  1022
   & & @{text else} @{text "if"} @{thm (prem 1) inv_begin_print(5)} @{text then} @{thm (rhs) inv_begin_print(5)}\\
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
  1023
  & & @{text else} @{thm (rhs) inv_begin_print(6)}
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
  1024
  \end{tabular}
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
  1025
  \end{center}
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
  1026
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
  1027
  \noindent
102
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
  1028
  This invariant depends on @{term n} representing the number of
141
4d7a568bd911 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 140
diff changeset
  1029
  @{term Oc}s on the tape. It is not hard (26
104
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
  1030
  lines of automated proof script) to show that for @{term "n >
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1031
  (0::nat)"} this invariant is preserved under the computation rules
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1032
  @{term step} and @{term steps}. This gives us partial correctness
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1033
  for @{term "tcopy_begin"}. 
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
  1034
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1035
  We next need to show that @{term "tcopy_begin"} terminates. For this
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1036
  we introduce lexicographically ordered pairs @{term "(n, m)"}
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1037
  derived from configurations @{text "(s, (l, r))"} whereby @{text n} is
106
c3155e9e4f63 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 105
diff changeset
  1038
  the state @{text s}, but ordered according to how @{term tcopy_begin} executes
c3155e9e4f63 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 105
diff changeset
  1039
  them, that is @{text "1 > 2 > 3 > 4 > 0"}; in order to have
130
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1040
  a strictly decreasing measure, @{term m} takes the data on the tape into
107
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1041
  account and is calculated according to the following measure function:
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
  1042
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
  1043
  \begin{center}
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1044
  \begin{tabular}{rcl}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1045
  @{term measure_begin_step}@{text "(s, (l, r))"} & @{text "\<equiv>"} & 
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1046
  @{text "if"} @{thm (prem 1) measure_begin_print(1)} @{text then} @{thm (rhs) measure_begin_print(1)}\\
106
c3155e9e4f63 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 105
diff changeset
  1047
  & & @{text else} @{text "if"} @{thm (prem 1) measure_begin_print(2)} @{text then} 
c3155e9e4f63 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 105
diff changeset
  1048
  @{text "("}@{thm (rhs) measure_begin_print(2)}@{text ")"} \\
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1049
 & & @{text else} @{text "if"} @{thm (prem 1) measure_begin_print(3)} @{text then} @{thm (rhs) measure_begin_print(3)}\\
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1050
  & & @{text else} @{thm (rhs) measure_begin_print(4)}\\
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1051
  \end{tabular}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1052
  \end{center}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1053
  
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1054
  \noindent
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1055
  With this in place, we can show that for every starting tape of the
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1056
  form @{term "([], Oc \<up> n)"} with @{term "n > (0::nat)"}, the Turing
107
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1057
  machine @{term "tcopy_begin"} will eventually halt (the measure
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1058
  decreases in each step). Taking this and the partial correctness
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1059
  proof together, we obtain the Hoare-triple shown on the left for @{term tcopy_begin}:
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1060
   
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1061
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1062
  \begin{center}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1063
  @{thm (concl) begin_correct}\hspace{6mm}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1064
  @{thm (concl) loop_correct}\hspace{6mm}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1065
  @{thm (concl) end_correct}
96
bd320b5365e2 updated uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 94
diff changeset
  1066
  \end{center}
36
4b35e0e0784b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1067
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1068
  \noindent 
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1069
  where we assume @{text "0 < n"} (similar reasoning is needed for
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1070
  the Hoare-triples for @{term tcopy_loop} and @{term tcopy_end}). Since the invariant of
107
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1071
  the halting state of @{term tcopy_begin} implies the invariant of
230
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1072
  the starting state of @{term tcopy_loop}, that is @{term "Turing_Hoare.assert_imp (inv_begin0 n)
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1073
  (inv_loop1 n)"} holds, and also @{term "inv_loop0 n = inv_end1
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1074
  n"}, we can derive the following Hoare-triple for the correctness 
107
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1075
  of @{term tcopy}:
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
  1076
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1077
  \begin{center}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1078
  @{thm tcopy_correct}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1079
  \end{center}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1080
107
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1081
  \noindent
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1082
  That means if we start with a tape of the form @{term "([], <n::nat>)"} then 
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1083
  @{term tcopy} will halt with the tape \mbox{@{term "([Bk], <(n::nat, n::nat)>)"}}, as desired.
107
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1084
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1085
  Finally, we are in the position to prove the undecidability of the halting problem.
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1086
  A program @{term p} started with a standard tape containing the (encoded) numbers
130
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1087
  @{term ns} will \emph{halt} with a standard tape containing a single (encoded) 
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1088
  number is defined as
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
  1089
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1090
  \begin{center}
170
eccd79a974ae updated some files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 169
diff changeset
  1091
  @{thm halts_def}
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1092
  \end{center}
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1093
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1094
  \noindent
107
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1095
  This roughly means we considering only Turing machine programs
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1096
  representing functions that take some numbers as input and produce a
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1097
  single number as output. For undecidability, the property we are
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1098
  proving is that there is no Turing machine that can decide in
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1099
  general whether a Turing machine program halts (answer either @{text
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1100
  0} for halting or @{text 1} for looping). Given our correctness
107
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1101
  proofs for @{term dither} and @{term tcopy} shown above, this
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1102
  non-existence is now relatively straightforward to establish. We first
107
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1103
  assume there is a coding function, written @{term "code M"}, which
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1104
  represents a Turing machine @{term "M"} as a natural number.  No
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1105
  further assumptions are made about this coding function. Suppose a
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1106
  Turing machine @{term H} exists such that if started with the
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1107
  standard tape @{term "([Bk], <(code M, ns)>)"} returns @{text 0},
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1108
  respectively @{text 1}, depending on whether @{text M} halts or not when
107
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1109
  started with the input tape containing @{term "<ns>"}.  This
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1110
  assumption is formalised as follows---for all @{term M} and all lists of
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1111
  natural numbers @{term ns}:
106
c3155e9e4f63 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 105
diff changeset
  1112
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1113
  \begin{center}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1114
  \begin{tabular}{r}
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1115
  @{thm (prem 2) uncomputable.h_case} implies
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1116
  @{thm (concl) uncomputable.h_case}\\
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1117
  
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1118
  @{thm (prem 2) uncomputable.nh_case} implies
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1119
  @{thm (concl) uncomputable.nh_case}
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1120
  \end{tabular}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1121
  \end{center}
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1122
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1123
  \noindent
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1124
  The contradiction can be derived using the following Turing machine
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1125
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1126
  \begin{center}
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1127
  @{thm tcontra_def}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1128
  \end{center}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1129
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1130
  \noindent
284
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
  1131
  Suppose @{thm (prem 1) "tcontra_halt"} holds. Given the invariants @{text "P\<^sub>1"}\ldots@{text "P\<^sub>3"} 
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1132
  shown on the 
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1133
  left, we can derive the following Hoare-pair for @{term tcontra} on the right.
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1134
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1135
  \begin{center}\small
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1136
  \begin{tabular}{@ {}c@ {\hspace{-10mm}}c@ {}}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1137
  \begin{tabular}[t]{@ {}l@ {}}
284
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
  1138
  @{term "P\<^sub>1 \<equiv> \<lambda>tp. tp = ([]::cell list, <code_tcontra>)"}\\
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
  1139
  @{term "P\<^sub>2 \<equiv> \<lambda>tp. tp = ([Bk], <(code_tcontra, code_tcontra)>)"}\\
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
  1140
  @{term "P\<^sub>3 \<equiv> \<lambda>tp. \<exists>k. tp = (Bk \<up> k, <0::nat>)"}
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1141
  \end{tabular}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1142
  &
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1143
  \begin{tabular}[b]{@ {}l@ {}}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1144
  \raisebox{-20mm}{$\inferrule*{
284
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
  1145
    \inferrule*{@{term "{P\<^sub>1} tcopy {P\<^sub>2}"} \\ @{term "{P\<^sub>2} H {P\<^sub>3}"}}
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
  1146
    {@{term "{P\<^sub>1} (tcopy |+| H) {P\<^sub>3}"}}
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
  1147
    \\ @{term "{P\<^sub>3} dither \<up>"}
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1148
   }
284
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
  1149
   {@{term "{P\<^sub>1} tcontra \<up>"}}
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1150
  $}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1151
  \end{tabular}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1152
  \end{tabular}
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1153
  \end{center}
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1154
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1155
  \noindent
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1156
  This Hoare-pair contradicts our assumption that @{term tcontra} started
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1157
  with @{term "<(code tcontra)>"} halts.
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1158
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1159
  Suppose @{thm (prem 1) "tcontra_unhalt"} holds. Again, given the invariants 
284
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
  1160
  @{text "Q\<^sub>1"}\ldots@{text "Q\<^sub>3"}
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1161
  shown on the 
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1162
  left, we can derive the Hoare-triple for @{term tcontra} on the right.
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1163
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1164
  \begin{center}\small
116
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
  1165
  \begin{tabular}{@ {}c@ {\hspace{-18mm}}c@ {}}
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1166
  \begin{tabular}[t]{@ {}l@ {}}
284
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
  1167
  @{term "Q\<^sub>1 \<equiv> \<lambda>tp. tp = ([]::cell list, <code_tcontra>)"}\\
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
  1168
  @{term "Q\<^sub>2 \<equiv> \<lambda>tp. tp = ([Bk], <(code_tcontra, code_tcontra)>)"}\\
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
  1169
  @{term "Q\<^sub>3 \<equiv> \<lambda>tp. \<exists>k. tp = (Bk \<up> k, <1::nat>)"}
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1170
  \end{tabular}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1171
  &
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1172
  \begin{tabular}[t]{@ {}l@ {}}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1173
  \raisebox{-20mm}{$\inferrule*{
284
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
  1174
    \inferrule*{@{term "{Q\<^sub>1} tcopy {Q\<^sub>2}"} \\ @{term "{Q\<^sub>2} H {Q\<^sub>3}"}}
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
  1175
    {@{term "{Q\<^sub>1} (tcopy |+| H) {Q\<^sub>3}"}}
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
  1176
    \\ @{term "{Q\<^sub>3} dither {Q\<^sub>3}"}
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1177
   }
284
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
  1178
   {@{term "{Q\<^sub>1} tcontra {Q\<^sub>3}"}}
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1179
  $}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1180
  \end{tabular}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1181
  \end{tabular}
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1182
  \end{center}
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1183
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1184
  \noindent
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1185
  This time the Hoare-triple states that @{term tcontra} terminates 
233
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
  1186
  with the ``output'' @{term "<(1::nat)>"}. In both cases we come
130
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1187
  to a contradiction, which means we have to abandon our assumption 
112
fea23f9a9d85 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 111
diff changeset
  1188
  that there exists a Turing machine @{term H} which can in general decide 
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1189
  whether Turing machines terminate.
9
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
  1190
*}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
  1191
63
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
  1192
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
  1193
section {* Abacus Machines *}
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
  1194
25
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
  1195
text {*
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
  1196
  \noindent
112
fea23f9a9d85 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 111
diff changeset
  1197
  Boolos et al \cite{Boolos87} use abacus machines as a stepping stone
fea23f9a9d85 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 111
diff changeset
  1198
  for making it less laborious to write Turing machine
284
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
  1199
  programs. Abacus machines operate over a set of registers @{text "R\<^sub>0"},
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
  1200
  @{text "R\<^sub>1"}, \ldots{}, @{text "R\<^sub>n"} each being able to hold an arbitrary large natural
113
8ef94047e6e2 abacus section updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 112
diff changeset
  1201
  number.  We use natural numbers to refer to registers; we also use a natural number
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1202
  to represent a program counter and to represent jumping ``addresses'', for which we 
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1203
  use the letter @{text l}. An abacus 
117
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1204
  program is a list of \emph{instructions} defined by the datatype:
25
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
  1205
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
  1206
  \begin{center}
111
dfc629cd11de made uncomputable compatible with abacus
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 109
diff changeset
  1207
  \begin{tabular}{rcl@ {\hspace{10mm}}l}
142
21c7139ffa07 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 141
diff changeset
  1208
  @{text "i"} & $::=$  & @{term "Inc R\<iota>"} & increment register @{text "R"} by one\\
21c7139ffa07 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 141
diff changeset
  1209
  & $\mid$ & @{term "Dec R\<iota> l"} & if content of @{text R} is non-zero, then decrement it by one\\
21c7139ffa07 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 141
diff changeset
  1210
  & & & otherwise jump to instruction @{text l}\\
21c7139ffa07 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 141
diff changeset
  1211
  & $\mid$ & @{term "Goto l"} & jump to instruction @{text l}
25
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
  1212
  \end{tabular}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
  1213
  \end{center}
27
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
  1214
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
  1215
  \noindent
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1216
  For example the program clearing the register @{text R} (that is setting
113
8ef94047e6e2 abacus section updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 112
diff changeset
  1217
  it to @{term "(0::nat)"}) can be defined as follows:
27
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
  1218
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
  1219
  \begin{center}
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1220
  @{thm clear.simps[where n="R\<iota>" and e="l", THEN eq_reflection]}
27
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
  1221
  \end{center}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
  1222
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
  1223
  \noindent
113
8ef94047e6e2 abacus section updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 112
diff changeset
  1224
  Running such a program means we start with the first instruction
117
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1225
  then execute one instructions after the other, unless there is a jump.  For
143
c56f04b6a2f9 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 142
diff changeset
  1226
  example the second instruction @{term "Goto 0"} above means
122
db518aba152a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 121
diff changeset
  1227
  we jump back to the first instruction thereby closing the loop.  Like with our
117
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1228
  Turing machines, we fetch instructions from an abacus program such
113
8ef94047e6e2 abacus section updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 112
diff changeset
  1229
  that a jump out of ``range'' behaves like a @{term "Nop"}-action. In
122
db518aba152a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 121
diff changeset
  1230
  this way it is again easy to define a function @{term steps} that
120
844e149696d4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 119
diff changeset
  1231
  executes @{term n} instructions of an abacus program. A \emph{configuration}
122
db518aba152a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 121
diff changeset
  1232
  of an abacus machine is the current program counter together with a snapshot of 
120
844e149696d4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 119
diff changeset
  1233
  all registers.
844e149696d4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 119
diff changeset
  1234
  By convention
119
892a3d7493aa updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 118
diff changeset
  1235
  the value calculated by an abacus program is stored in the
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1236
  last register (the one with the highest index in the program).
113
8ef94047e6e2 abacus section updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 112
diff changeset
  1237
  
115
653426ed4b38 started with abacus section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 114
diff changeset
  1238
  The main point of abacus programs is to be able to translate them to 
117
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1239
  Turing machine programs. Registers and their content are represented by
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1240
  standard tapes (see definition shown in \eqref{standard}). Because of the jumps in abacus programs, it
187
326310016da9 polished some typos in the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 186
diff changeset
  1241
  is impossible to build Turing machine programs out of components 
239
ac3309722536 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 237
diff changeset
  1242
  using our @{text ";"}-operation shown in the previous section.
124
bda714532263 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
  1243
  To overcome this difficulty, we calculate a \emph{layout} of an
bda714532263 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
  1244
  abacus program as follows
117
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1245
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1246
  \begin{center}
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1247
  \begin{tabular}[t]{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}}
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1248
  @{thm (lhs) layout(1)} & @{text "\<equiv>"} & @{thm (rhs) layout(1)}\\
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1249
  @{thm (lhs) layout(2)} & @{text "\<equiv>"} & @{thm (rhs) layout(2)}\\
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1250
  @{thm (lhs) layout(3)} & @{text "\<equiv>"} & @{thm (rhs) layout(3)}\\
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1251
  @{thm (lhs) layout(4)} & @{text "\<equiv>"} & @{thm (rhs) layout(4)}\\
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1252
  \end{tabular}
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1253
  \end{center}
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1254
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1255
  \noindent
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1256
  This gives us a list of natural numbers specifying how many states
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1257
  are needed to translate each abacus instruction. This information
162
a63c3f8d7234 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 159
diff changeset
  1258
  is needed in order to calculate the state where the Turing machine program
202
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
  1259
  of an abacus instruction starts. This can be defined as
194
fc2a5e9fbb97 updated Scala files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 190
diff changeset
  1260
202
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
  1261
  \begin{center}
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
  1262
  @{thm address.simps[where x="n"]}
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
  1263
  \end{center}
194
fc2a5e9fbb97 updated Scala files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 190
diff changeset
  1264
202
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
  1265
  \noindent
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
  1266
  where @{text p} is an abacus program and @{term "take n"} takes the first
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
  1267
  @{text n} elements from a list.
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
  1268
  
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1269
  The @{text Goto}
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1270
  instruction is easiest to translate requiring only one state, namely
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1271
  the Turing machine program:
115
653426ed4b38 started with abacus section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 114
diff changeset
  1272
653426ed4b38 started with abacus section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 114
diff changeset
  1273
  \begin{center}
144
07730607b0ca updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 143
diff changeset
  1274
  @{text "translate_Goto l"} @{text "\<equiv>"} @{thm (rhs) tgoto.simps[where n="l"]}
117
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1275
  \end{center}
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1276
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1277
  \noindent
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1278
  where @{term "l"} is the state in the Turing machine program 
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1279
  to jump to. For translating the instruction @{term "Inc R\<iota>"}, 
117
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1280
  one has to remember that the content of the registers are encoded
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1281
  in the Turing machine as a standard tape. Therefore the translated Turing machine 
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1282
  needs to first find the number corresponding to the content of register 
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1283
  @{text "R"}. This needs a machine
117
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1284
  with @{term "(2::nat) * R\<iota>"} states and can be constructed as follows: 
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1285
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1286
  \begin{center}
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1287
  \begin{tabular}[t]{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}}
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1288
  @{thm (lhs) findnth.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) findnth.simps(1)}\\
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1289
  @{thm (lhs) findnth.simps(2)} & @{text "\<equiv>"}\\
202
7cfc83879fc9 added a coment about partial_function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 198
diff changeset
  1290
  \multicolumn{3}{@ {}l@ {}}{\hspace{6mm}@{thm (rhs) findnth.simps(2)}}\\
117
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1291
  \end{tabular}
115
653426ed4b38 started with abacus section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 114
diff changeset
  1292
  \end{center}
653426ed4b38 started with abacus section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 114
diff changeset
  1293
117
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1294
  \noindent
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1295
  Then we need to increase the ``number'' on the tape by one,
144
07730607b0ca updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 143
diff changeset
  1296
  and adjust the following ``registers''. For adjusting we only need to 
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1297
  change the first @{term Oc} of each number to @{term Bk} and the last
117
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1298
  one from @{term Bk} to @{term Oc}.
124
bda714532263 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
  1299
  Finally, we need to transition the head of the
117
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1300
  Turing machine back into the standard position. This requires a Turing machine
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1301
  with 9 states (we omit the details). Similarly for the translation of @{term "Dec R\<iota> l"}, where the 
117
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1302
  translated Turing machine needs to first check whether the content of the
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1303
  corresponding register is @{text 0}. For this we have a Turing machine program
187
326310016da9 polished some typos in the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 186
diff changeset
  1304
  with @{text 16} states (again the details are omitted). 
117
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1305
121
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
  1306
  Finally, having a Turing machine for each abacus instruction we need
117
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1307
  to ``stitch'' the Turing machines together into one so that each
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1308
  Turing machine component transitions to next one, just like in
119
892a3d7493aa updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 118
diff changeset
  1309
  the abacus programs. One last problem to overcome is that an abacus
892a3d7493aa updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 118
diff changeset
  1310
  program is assumed to calculate a value stored in the last
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1311
  register (the one with the highest register). That means we have to append a Turing machine that
119
892a3d7493aa updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 118
diff changeset
  1312
  ``mops up'' the tape (cleaning all @{text Oc}s) except for the
162
a63c3f8d7234 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 159
diff changeset
  1313
  @{term Oc}s of the last register represented on the tape. This needs
171
1bad3ec5bcd5 typo in the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 170
diff changeset
  1314
  a Turing machine program with @{text "2 * R + 6"} states, assuming @{text R}
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1315
  is the number of registers to be ``cleaned''.
119
892a3d7493aa updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 118
diff changeset
  1316
124
bda714532263 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
  1317
  While generating the Turing machine program for an abacus program is
bda714532263 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
  1318
  not too difficult to formalise, the problem is that it contains
bda714532263 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
  1319
  @{text Goto}s all over the place. The unfortunate result is that we
bda714532263 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
  1320
  cannot use our Hoare-rules for reasoning about sequentially composed
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1321
  programs (for this each component needs to be completely independent). Instead we
124
bda714532263 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
  1322
  have to treat the translated Turing machine as one ``big block'' and 
bda714532263 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
  1323
  prove as invariant that it performs
bda714532263 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
  1324
  the same operations as the abacus program. For this we have to show
bda714532263 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
  1325
  that for each configuration of an abacus machine the @{term
bda714532263 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
  1326
  step}-function is simulated by zero or more steps in our translated
bda714532263 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
  1327
  Turing machine. This leads to a rather large ``monolithic''
bda714532263 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
  1328
  correctness proof (4600 loc and 380 sublemmas) that on the
bda714532263 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
  1329
  conceptual level is difficult to break down into smaller components.
121
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
  1330
  
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
  1331
  %We were able to simplify the proof somewhat
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
  1332
*}
29
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 28
diff changeset
  1333
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 28
diff changeset
  1334
121
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
  1335
section {* Recursive Functions and a Universal Turing Machine *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
  1336
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
  1337
text {*
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1338
  The main point of recursive functions is that we can relatively 
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1339
  easily construct a universal Turing machine via a universal 
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1340
  function. This is different from Norrish \cite{Norrish11} who gives a universal 
232
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
  1341
  function for the lambda-calculus, and also from Asperti and Ricciotti 
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1342
  \cite{AspertiRicciotti12} who construct a universal Turing machine
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1343
  directly, but for simulating Turing machines with a more restricted alphabet.
214
763ed488fa79 added a comment about deeply embedding of recursive functions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 202
diff changeset
  1344
  Unlike Norrish \cite{Norrish11}, we need to represent recursive functions
763ed488fa79 added a comment about deeply embedding of recursive functions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 202
diff changeset
  1345
  ``deeply'' because we want to translate them to abacus programs.
763ed488fa79 added a comment about deeply embedding of recursive functions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 202
diff changeset
  1346
  Thus \emph{recursive functions} are defined as the datatype
121
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
  1347
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1348
  \begin{center}
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1349
  \begin{tabular}{c@ {\hspace{4mm}}c}
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1350
  \begin{tabular}{rcl@ {\hspace{4mm}}l}
164
8a3e63163910 fixed compilation of paper and typo
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 162
diff changeset
  1351
  @{term r} & @{text "::="} & @{term z} & (zero-function)\\
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1352
            & @{text "|"}   & @{term s} & (successor-function)\\
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1353
            & @{text "|"}   & @{term "id n m"} & (projection)\\
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1354
  \end{tabular} &
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1355
  \begin{tabular}{cl@ {\hspace{4mm}}l}
230
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1356
  @{text "|"} & @{term "Cn n f fs"} & (composition)\\
284
a21fb87bb0bd soem changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 239
diff changeset
  1357
  @{text "|"} & @{term "Pr n f\<^sub>1 f\<^sub>2"} & (primitive recursion)\\
230
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1358
  @{text "|"} & @{term "Mn n f"} & (minimisation)\\
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1359
  \end{tabular}
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1360
  \end{tabular}
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1361
  \end{center}
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1362
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1363
  \noindent 
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1364
  where @{text n} indicates the function expects @{term n} arguments
194
fc2a5e9fbb97 updated Scala files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 190
diff changeset
  1365
  (in \cite{Boolos87} both @{text z} and @{term s} expect one
230
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1366
  argument), and @{text fs} stands for a list of recursive
194
fc2a5e9fbb97 updated Scala files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 190
diff changeset
  1367
  functions. Since we know in each case the arity, say @{term l}, we
230
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1368
  can define an evaluation function, called @{term rec_exec}, that takes a recursive
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1369
  function @{text f} and a list @{term ns} of natural numbers of
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1370
  length @{text l} as arguments. Since this evaluation function uses 
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1371
  the minimisation operator
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1372
  from HOL, this function might not terminate always. As a result we also
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1373
  need to inductively characterise when @{term rec_exec} terminates.
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1374
  We omit the definitions for
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1375
  @{term "rec_exec f ns"} and @{term "terminate f ns"}. Because of 
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1376
  space reasons, we also omit the definition of
194
fc2a5e9fbb97 updated Scala files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 190
diff changeset
  1377
  translating recursive functions into abacus programs. We can prove,
fc2a5e9fbb97 updated Scala files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 190
diff changeset
  1378
  however, the following theorem about the translation: If @{thm (prem
230
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1379
  1) recursive_compile_to_tm_correct3[where recf="f" and args="ns"]} 
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1380
  holds for the recursive function @{text f} and arguments @{term ns}, then the
194
fc2a5e9fbb97 updated Scala files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 190
diff changeset
  1381
  following Hoare-triple holds
130
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1382
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1383
  \begin{center}
230
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1384
  @{thm (concl) recursive_compile_to_tm_correct3[where recf="f" and args="ns"]}
130
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1385
  \end{center}
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1386
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1387
  \noindent
230
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1388
  for the Turing machine generated by @{term "translate f"}. This
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1389
  means the translated Turing machine if started
130
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1390
  with the standard tape @{term "([Bk, Bk], <ns::nat list>)"} will terminate
230
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1391
  with the standard tape @{term "(Bk \<up> k, <(rec_exec f ns)::nat> @ Bk \<up> l)"} for some @{text k} and @{text l}.
130
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1392
153
a4601143db90 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 152
diff changeset
  1393
  Having recursive functions under our belt, we can construct a universal
a4601143db90 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 152
diff changeset
  1394
  function, written @{text UF}. This universal function acts like an interpreter for Turing machines.
a4601143db90 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 152
diff changeset
  1395
  It takes two arguments: one is the code of the Turing machine to be interpreted and the 
a4601143db90 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 152
diff changeset
  1396
  other is the ``packed version'' of the arguments of the Turing machine. 
a4601143db90 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 152
diff changeset
  1397
  We can then consider how this universal function is translated to a 
a4601143db90 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 152
diff changeset
  1398
  Turing machine and from this construct the universal Turing machine, 
239
ac3309722536 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 237
diff changeset
  1399
  written @{term UTM}. It is defined as 
153
a4601143db90 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 152
diff changeset
  1400
  the composition of the Turing machine that packages the arguments and
a4601143db90 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 152
diff changeset
  1401
  the translated recursive 
a4601143db90 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 152
diff changeset
  1402
  function @{text UF}:
a4601143db90 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 152
diff changeset
  1403
a4601143db90 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 152
diff changeset
  1404
  \begin{center}
239
ac3309722536 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 237
diff changeset
  1405
  @{text "UTM \<equiv> arg_coding ; (translate UF)"}
153
a4601143db90 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 152
diff changeset
  1406
  \end{center}
a4601143db90 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 152
diff changeset
  1407
a4601143db90 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 152
diff changeset
  1408
  \noindent
146
0f52b971cc03 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
  1409
  Suppose
152
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1410
  a Turing program @{term p} is well-formed and  when started with the standard 
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1411
  tape containing the arguments @{term args}, will produce a standard tape
146
0f52b971cc03 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
  1412
  with ``output'' @{term n}. This assumption can be written as the
0f52b971cc03 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
  1413
  Hoare-triple
145
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1414
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1415
  \begin{center}
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1416
  @{thm (prem 3) UTM_halt_lemma2}
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1417
  \end{center}
146
0f52b971cc03 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
  1418
  
0f52b971cc03 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
  1419
  \noindent
153
a4601143db90 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 152
diff changeset
  1420
  where we require that the @{term args} stand for a non-empty list. Then the universal Turing
146
0f52b971cc03 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
  1421
  machine @{term UTM} started with the code of @{term p} and the arguments
153
a4601143db90 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 152
diff changeset
  1422
  @{term args}, calculates the same result, namely
145
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1423
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1424
  \begin{center}
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1425
  @{thm (concl) UTM_halt_lemma2} 
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1426
  \end{center}
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1427
146
0f52b971cc03 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
  1428
  \noindent
0f52b971cc03 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
  1429
  Similarly, if a Turing program @{term p} started with the 
0f52b971cc03 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
  1430
  standard tape containing @{text args} loops, which is represented
0f52b971cc03 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
  1431
  by the Hoare-pair
145
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1432
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1433
  \begin{center}
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1434
  @{thm (prem 2) UTM_unhalt_lemma2}
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1435
  \end{center}
146
0f52b971cc03 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
  1436
0f52b971cc03 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
  1437
  \noindent
0f52b971cc03 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
  1438
  then the universal Turing machine started with the code of @{term p} and the arguments
0f52b971cc03 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
  1439
  @{term args} will also loop
145
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1440
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1441
  \begin{center}
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1442
  @{thm (concl) UTM_unhalt_lemma2} 
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1443
  \end{center}
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1444
152
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1445
  %Analysing the universal Turing machine constructed in \cite{Boolos87} more carefully
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1446
  %we can strengthen this result slightly by observing that @{text m} is at most
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1447
  %2 in the output tape. This observation allows one to construct a universal Turing machine that works
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1448
  %entirely on the left-tape by composing it with a machine that drags the tape
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1449
  %two cells to the right. A corollary is that one-sided Turing machines (where the
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1450
  %tape only extends to the right) are computationally as powerful as our two-sided 
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1451
  %Turing machines. So our undecidability proof for the halting problem extends
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1452
  %also to one-sided Turing machines, which is needed for example in order to
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1453
  %formalise the undecidability of Wang's tiling problem \cite{Robinson71}.
149
199cf7ce1169 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 146
diff changeset
  1454
154
9b9e0d37fc19 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 153
diff changeset
  1455
  While formalising the chapter in \cite{Boolos87} about universal
9b9e0d37fc19 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 153
diff changeset
  1456
  Turing machines, an unexpected outcome of our work is that we
9b9e0d37fc19 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 153
diff changeset
  1457
  identified an inconsistency in their use of a definition. This is
9b9e0d37fc19 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 153
diff changeset
  1458
  unexpected since \cite{Boolos87} is a classic textbook which has
186
455411d69c12 added link and comment to fourth edition of Boolos
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 185
diff changeset
  1459
  undergone several editions (we used the fifth edition; the material 
455411d69c12 added link and comment to fourth edition of Boolos
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 185
diff changeset
  1460
  containing the inconsistency was introduced in the fourth edition
233
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
  1461
  of this book). The central
154
9b9e0d37fc19 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 153
diff changeset
  1462
  idea about Turing machines is that when started with standard tapes
232
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
  1463
  they compute a partial arithmetic function.  The inconsistency arises
155
1834acc6fd76 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  1464
  when they define the case when this function should \emph{not} return a
187
326310016da9 polished some typos in the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 186
diff changeset
  1465
  result. Boolos at al write in Chapter 3, Page 32:
134
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1466
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1467
  \begin{quote}\it
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1468
  ``If the function that is to be computed assigns no value to the arguments that 
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1469
  are represented initially on the tape, then the machine either will never halt, 
156
7c9dbacc6c7c updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 155
diff changeset
  1470
  \colorbox{mygrey}{or} will halt in some nonstandard configuration\ldots''
134
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1471
  \end{quote}
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1472
  
146
0f52b971cc03 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 145
diff changeset
  1473
  \noindent
152
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1474
  Interestingly, they do not implement this definition when constructing
155
1834acc6fd76 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  1475
  their universal Turing machine. In Chapter 8, on page 93, a recursive function 
152
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1476
  @{term stdh} is defined as:
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1477
	
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1478
  \begin{equation}\label{stdh_def}
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1479
  @{text "stdh(m, x, t) \<equiv> stat(conf(m, x, t)) + nstd(conf(m, x, t))"}
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1480
  \end{equation}
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1481
  
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1482
  \noindent
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1483
  where @{text "stat(conf(m, x, t))"} computes the current state of the
155
1834acc6fd76 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  1484
  simulated Turing machine, and @{text "nstd(conf(m, x, t))"} returns @{text 1}
152
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1485
  if the tape content is non-standard. If either one evaluates to
155
1834acc6fd76 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  1486
  something that is not zero, then @{text "stdh(m, x, t)"} will be not zero, because of 
232
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
  1487
  the $+$-operation. On the same page, a function @{text "halt(m, x)"} is defined 
152
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1488
  in terms of @{text stdh} for computing the steps the Turing machine needs to
155
1834acc6fd76 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  1489
  execute before it halts (in case it halts at all). According to this definition, the simulated
152
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1490
  Turing machine will continue to run after entering the @{text 0}-state
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1491
  with a non-standard tape. The consequence of this inconsistency is
155
1834acc6fd76 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  1492
  that there exist Turing machines that given some arguments do not compute a value
1834acc6fd76 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  1493
  according to Chapter 3, but return a proper result according to
152
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1494
  the definition in Chapter 8. One such Turing machine is:
134
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1495
152
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1496
  %This means that if you encode the plus function but only give one argument,
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1497
  %then the TM will either loop {\bf or} stop with a non-standard tape
134
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1498
152
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1499
  %But in the definition of the universal function the TMs will never stop
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1500
  %with non-standard tapes. 
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1501
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1502
  %SO the following TM calculates something according to def from chap 8,
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1503
  %but not with chap 3. For example:
134
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1504
  
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1505
  \begin{center}
155
1834acc6fd76 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  1506
  @{term "counter_example \<equiv> [(L, (0::nat)), (L, 2), (R, 2), (R, 0)]"}
134
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1507
  \end{center}
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1508
152
2c0d79801e36 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 151
diff changeset
  1509
  \noindent
155
1834acc6fd76 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  1510
  If started with standard tape @{term "([], [Oc])"}, it halts with the
232
8f89170bb076 updated according to comments from reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 231
diff changeset
  1511
  non-standard tape @{term "([Oc, Bk], [])"} according to the definition in Chapter 3---so no
234
ca2ea835c363 fixed counterexample according to def in Chap 8
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 233
diff changeset
  1512
  result is calculated; but with the standard tape @{term "([Bk], [Oc])"} according to the 
ca2ea835c363 fixed counterexample according to def in Chap 8
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 233
diff changeset
  1513
  definition in Chapter 8. 
230
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1514
  We solve this inconsistency in our formalisation by
155
1834acc6fd76 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 154
diff changeset
  1515
  setting up our definitions so that the @{text counter_example} Turing machine does not 
157
fe0e6733b9e4 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 156
diff changeset
  1516
  produce any result by looping forever fetching @{term Nop}s in state @{text 0}. 
230
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1517
  This solution implements essentially the definition in Chapter 3; it  
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1518
  differs from the definition in Chapter 8, where perplexingly the instruction 
49dcc0b9b0b3 adapted paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 218
diff changeset
  1519
  from state @{text 1} is fetched.
140
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1520
*}
134
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1521
145
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1522
(*
140
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1523
section {* XYZ *}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1524
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1525
text {*
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1526
One of the main objectives of the paper is the construction and verification of Universal Turing machine (UTM). A UTM takes the code of any Turing machine $M$ and its arguments $a_1, a_2, \ldots, a_n$ as input and computes to the same effect as $M$ does on $a_1, a_2, \ldots, a_n$. That is to say:
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1527
\begin{enumerate}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1528
    \item If $M$ terminates and gives a result on $a_1, a_2, \ldots, a_n$, so does $UTM$ on input
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1529
        $
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1530
        code(M), a_1, a_1, a_2, \ldots, a_n
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1531
        $.
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1532
    \item If $M$ loops forever on $a_1, a_2, \ldots, a_n$, then $UTM$ does the same on $code (M), a_1, a_1, a_2, \ldots, a_n$.
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1533
\end{enumerate}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1534
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1535
The existence of UTM is the cornerstone of {\em Turing Thesis}, which says: any effectively computable function can be computed by a Turing Machine. The evaluation of Turing machine is obviously effective computable (otherwise, Turing machine is not an effect computation model). So, if the evaluation function of Turing machine can not be implemented by a Turing machine, the {\em Turing Thesis} would fail. Although people believe that UTM exists, few have gave one in close form and prove its correctness with the only exception of Asperti.
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1536
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1537
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1538
The method to obtain Universal Turing machine (UTM), as hinted by Boolos's book, is first constructing a recursive function recF (named Universal Function), which serves as an interpreter for Turing machine program, and then the UTM is obtained by $translate(recF)$. However, since any particular recursive function only takes fixed number of arguments determined by its construction,  no matter how recF is constructed, it can only server as interpret for Turing machines which take the fixed number of arguments as input. Our solution is to precede the $translate(recF)$ with a Turing machine which compacts multiple arguments into one argument using Wang's coding. Now, $recF$ is defined as a function taking two arguments, where the first is the code of Turing machine to be interpreted and the second is the packed arguments.
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1539
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1540
The construction of recF roughly follows the idea in the book. Since the book gives no correctness proof of the construction (not even an informal one), we have to formulate the correctness statements and as well as their formal proofs explicitly. As an unexpected outcome of this formulation, we identified one inconsistency in Boolos' book. Every Turing machine is supposed to compute an arithmetic function which is possibly partial. When the TM is started with an argument where the function is undefined, the definition on Chapter 3 (page 32) says:
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1541
\begin{quote}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1542
(e) If the function that is to be computed assigns no value to the arguments that are
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1543
represented initially on the tape, then the machine either will never halt, or will
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1544
halt in some nonstandard configuration such as $B_n11111$ or $B11_n111$ or $B11111_n$.
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1545
\end{quote}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1546
According to this definition, a TM can signify a non-result either by looping forever or get into state 0 with a nonstandard tape. However, when we were trying to formalize the universal function in Chapter 8, we found the definition given there is not in accordance. On page 93, an recursive function $stdh$ is defined as:
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1547
	\begin{equation}\label{stdh_def}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1548
stdh(m, x, t) = stat(conf(m, x, t)) + nstd(conf(m, x, t))
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1549
\end{equation}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1550
Where $ stat(conf(m, x, t)) $ computes the current state of the simulated Turing machine, and the $nstd(conf(m, x, t))$ returns $1$ if the tape content is nonstandard. If either one evaluates to nonzero, stdh(m, x, t) will be nonzero, because of the $+$ operation. One the same page, a function $halt(m, x)$ is defined to in terms of $stdh$ to computes the steps the Turing machine needs to execute before halt, which stipulates the TM halts when nstd(conf(m, x, t)) returns $0$. According to this definition, the simulated Turing machine will continue to run after getting into state $0$ with a nonstandard tape. The consequence of this inconsistency is that: there exists Turing machines which computes non-value according to Chapter 3, but returns a proper result according to Chapter 8. One such Truing machine is:
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1551
    \begin{equation}\label{contrived_tm}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1552
        [(L, 0), (L, 2), (R, 2), (R, 0)]
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1553
    \end{equation}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1554
Starting in a standard configuration (1, [], [Oc]), it goes through the following series of configurations leading to state 0:
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1555
\[
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1556
(1, [], [Oc]) \rightsquigarrow (L, 2) \rightsquigarrow  (2, [], [Bk, Oc]) \rightsquigarrow (R, 2)\rightsquigarrow (2, [Bk], [Oc]) \rightsquigarrow (R, 0)\rightsquigarrow  (0, [Bk, Oc], [])
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1557
\]
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1558
According to Chapter 3, this Turing machine halts and gives a non-result. According to Chapter 8, it will continue to fetch and execute the next instruction. The fetching function $actn$ and $newstat$ in \eqref{fetch-def} (taken from page 92) takes $q$ as current state and $r$ as the currently scanned cell.
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1559
\begin{equation}\label{fetch-def}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1560
\begin{aligned}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1561
	actn(m, q, r )  &= ent(m, 4(q - 1)  + 2 \cdot scan(r )) \\
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1562
newstat(m, q, r ) & = ent(m, (4(q - 1) + 2 \cdot scan(r )) + 1)
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1563
\end{aligned}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1564
\end{equation}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1565
For our instance, $q=0$ and $r = 1$. Because $q - 1 = 0 - 1 = 1 - 1 = 0$, the instruction fetched by \eqref{fetch-def} at state $0$ will be the same as if the machine is at state $0$. So the Turing machine will go through the follow execution and halt with a standard tape:
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1566
\[
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1567
(0, [Bk, Oc], []) \rightsquigarrow (L, 0) \rightsquigarrow (0, [Bk], [Oc])
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1568
\]
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1569
In summary, according to Chapter 3, the Turing machine in \eqref{contrived_tm} computes non-result and according to Chapter 8, it computes an identify function. 
25
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
  1570
*}
145
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1571
*)
25
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
  1572
121
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
  1573
(*
13
a7ec585d7f20 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  1574
section {* Wang Tiles\label{Wang} *}
7
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1575
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1576
text {*
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1577
  Used in texture mapings - graphics
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1578
*}
121
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
  1579
*)
7
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1580
121
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
  1581
section {* Conclusion *}
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1582
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1583
text {*
136
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1584
  In previous works we were unable to formalise results about
237
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1585
  computability because in Isabelle/HOL we cannot, for example,
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1586
  represent the decidability of a predicate @{text P}, say, as the
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1587
  formula @{term "P \<or> \<not>P"}. For reasoning about computability we need
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1588
  to formalise a concrete model of computations. We could have
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1589
  followed Norrish \cite{Norrish11} using the $\lambda$-calculus as
239
ac3309722536 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 237
diff changeset
  1590
  the starting point for formalising computability theory, but then we would have
237
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1591
  to reimplement on the ML-level his infrastructure for rewriting
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1592
  $\lambda$-terms modulo $\beta$-equality: HOL4 has a simplifer that
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1593
  can rewrite terms modulo an arbitrary equivalence relation, which
239
ac3309722536 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 237
diff changeset
  1594
  Isabelle unfortunately does not yet have.  Even though, we would
237
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1595
  still need to connect $\lambda$-terms somehow to Turing machines for
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1596
  proofs that make essential use of them (for example the
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1597
  undecidability proof for Wang's tiling problem \cite{Robinson71}).
136
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1598
158
6a584d61820f updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 157
diff changeset
  1599
  We therefore have formalised Turing machines in the first place and the main
136
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1600
  computability results from Chapters 3 to 8 in the textbook by Boolos
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1601
  et al \cite{Boolos87}.  For this we did not need to implement
145
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1602
  anything on the ML-level of Isabelle/HOL. While formalising the six chapters
162
a63c3f8d7234 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 159
diff changeset
  1603
  of \cite{Boolos87} we have found an inconsistency in Boolos et al's 
138
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1604
  definitions of what function a Turing machine calculates. In
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1605
  Chapter 3 they use a definition that states a function is undefined
136
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1606
  if the Turing machine loops \emph{or} halts with a non-standard
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1607
  tape. Whereas in Chapter 8 about the universal Turing machine, the
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1608
  Turing machines will \emph{not} halt unless the tape is in standard
237
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1609
  form. Like Nipkow \cite{Nipkow98} observed with his formalisation
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1610
  of a textbook, we found that Boolos et al are (almost)
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1611
  right. We have not attempted to formalise everything precisely as
136
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1612
  Boolos et al present it, but use definitions that make our
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1613
  mechanised proofs manageable. For example our definition of the
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1614
  halting state performing @{term Nop}-operations seems to be
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1615
  non-standard, but very much suited to a formalisation in a theorem
145
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1616
  prover where the @{term steps}-function needs to be total.
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
  1617
237
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1618
   Norrish mentions that formalising Turing machines would be a
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1619
  ``\emph{daunting prospect}'' \cite[Page 310]{Norrish11}. While
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1620
  $\lambda$-terms indeed lead to some slick mechanised proofs, our
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1621
  experience is that Turing machines are not too daunting if one is
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1622
  only concerned with formalising the undecidability of the halting
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1623
  problem for Turing machines.  As a point of comparison, the halting
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1624
  problem took us around 1500 loc of Isar-proofs, which is just
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1625
  one-and-a-half times of a mechanised proof pearl about the
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1626
  Myhill-Nerode theorem. So our conclusion is that this part is not as
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1627
  daunting as we estimated when reading the paper by Norrish
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1628
  \cite{Norrish11}. The work involved with constructing a universal
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1629
  Turing machine via recursive functions and abacus machines, we
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1630
  agree, is not a project one wants to undertake too many times (our
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1631
  formalisation of abacus machines and their correct translation is
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1632
  approximately 4600 loc; recursive functions 2800 loc and the
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1633
  universal Turing machine 10000 loc).
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
  1634
  
159
b4b789a59086 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 158
diff changeset
  1635
  Our work is also very much inspired by the formalisation of Turing
187
326310016da9 polished some typos in the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 186
diff changeset
  1636
  machines of Asperti and Ricciotti \cite{AspertiRicciotti12} in the
125
1ce74a77fa2a conclusion
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 124
diff changeset
  1637
  Matita theorem prover. It turns out that their notion of
1ce74a77fa2a conclusion
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 124
diff changeset
  1638
  realisability and our Hoare-triples are very similar, however we
1ce74a77fa2a conclusion
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 124
diff changeset
  1639
  differ in some basic definitions for Turing machines. Asperti and
1ce74a77fa2a conclusion
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 124
diff changeset
  1640
  Ricciotti are interested in providing a mechanised foundation for
130
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1641
  complexity theory. They formalised a universal Turing machine
125
1ce74a77fa2a conclusion
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 124
diff changeset
  1642
  (which differs from ours by using a more general alphabet), but did
1ce74a77fa2a conclusion
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 124
diff changeset
  1643
  not describe an undecidability proof. Given their definitions and
138
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1644
  infrastructure, we expect however this should not be too difficult 
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1645
  for them.
125
1ce74a77fa2a conclusion
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 124
diff changeset
  1646
  
159
b4b789a59086 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 158
diff changeset
  1647
  For us the most interesting aspects of our work are the correctness
138
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1648
  proofs for Turing machines. Informal presentations of computability
125
1ce74a77fa2a conclusion
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 124
diff changeset
  1649
  theory often leave the constructions of particular Turing machines
145
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1650
  as exercise to the reader, for example \cite{Boolos87}, deeming
138
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1651
  it to be just a chore. However, as far as we are aware all informal
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1652
  presentations leave out any arguments why these Turing machines
145
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1653
  should be correct.  This means the reader is left
138
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1654
  with the task of finding appropriate invariants and measures for
159
b4b789a59086 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 158
diff changeset
  1655
  showing the correctness and termination of these Turing machines.
125
1ce74a77fa2a conclusion
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 124
diff changeset
  1656
  Whenever we can use Hoare-style reasoning, the invariants are
231
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 230
diff changeset
  1657
  relatively straightforward and again as a point of comparison much smaller than for example the
138
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1658
  invariants used by Myreen in a correctness proof of a garbage collector
145
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1659
  written in machine code \cite[Page 76]{Myreen09}. However, the invariant 
138
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1660
  needed for the abacus proof, where Hoare-style reasoning does not work, is
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1661
  similar in size as the one by Myreen and finding a sufficiently
159
b4b789a59086 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 158
diff changeset
  1662
  strong one took us, like Myreen, something on the magnitude of
138
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1663
  weeks.
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
  1664
138
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1665
  Our reasoning about the invariants is not much supported by the
237
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1666
  automation beyond the standard automation tools available in
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1667
  Isabelle/HOL.  There is however a tantalising connection between our
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1668
  work and very recent work by Jensen et al \cite{Jensen13} on
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1669
  verifying X86 assembly code that might change that. They observed a
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1670
  similar phenomenon with assembly programs where Hoare-style
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1671
  reasoning is sometimes possible, but sometimes it is not. In order
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1672
  to ease their reasoning, they introduced a more primitive
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1673
  specification logic, on which Hoare-rules can be provided for
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1674
  special cases.  It remains to be seen whether their specification
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1675
  logic for assembly code can make it easier to reason about our
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1676
  Turing programs. That would be an attractive result, because Turing
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1677
  machine programs are very much like assembly programs and it would
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1678
  connect some very classic work on Turing machines to very
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1679
  cutting-edge work on machine code verification. In order to try out
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1680
  such ideas, our formalisation provides the ``playground''. The code
06a6db387cd2 updated and small modification
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 236
diff changeset
  1681
  of our formalisation is available from the
145
38d8e0e37b7d updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 144
diff changeset
  1682
  Mercurial repository at \url{http://www.dcs.kcl.ac.uk/staff/urbanc/cgi-bin/repos.cgi/tm/}.
233
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
  1683
  \medskip
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
  1684
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
  1685
  \noindent
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
  1686
  {\bf Acknowledgements:} We are very grateful for the extremely helpful
e0a7ee9842d6 changed the introduction adn cited Zammit
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 232
diff changeset
  1687
  comments by the anonymous reviewers.
114
120091653998 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 113
diff changeset
  1688
*}
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1689
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1690
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1691
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1692
(*<*)
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1693
end
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1694
end
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1695
(*>*)