Paper/Paper.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Wed, 06 Feb 2013 14:14:35 +0000
changeset 140 7f5243700f25
parent 138 7fa1b8e88d76
child 141 4d7a568bd911
permissions -rw-r--r--
updated
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
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
     3
imports "../thys/recursive"
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
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
     6
25
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
     7
hide_const (open) s 
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
     8
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
     9
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    10
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
    11
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
    12
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
    13
abbreviation
48
559e5c6e5113 updated to ITP and updated directories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
    14
  "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
    15
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    16
consts DUMMY::'a
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    17
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    18
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    19
(* THEOREMS *)
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    20
notation (Rule output)
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    21
  "==>"  ("\<^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
    22
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    23
syntax (Rule output)
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    24
  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    25
  ("\<^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
    26
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    27
  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" 
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    28
  ("\<^raw:\mbox{>_\<^raw:}\\>/ _")
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    29
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    30
  "_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
    31
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    32
notation (Axiom output)
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    33
  "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
    34
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    35
notation (IfThen output)
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    36
  "==>"  ("\<^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
    37
syntax (IfThen output)
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    38
  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    39
  ("\<^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
    40
  "_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
    41
  "_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
    42
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    43
notation (IfThenNoBox output)
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    44
  "==>"  ("\<^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
    45
syntax (IfThenNoBox output)
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    46
  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    47
  ("\<^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
    48
  "_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
    49
  "_asm" :: "prop \<Rightarrow> asms" ("_")
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    50
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
context uncomputable
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    53
begin
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
    54
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    55
notation (latex output)
100
dfe852aacae6 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 99
diff changeset
    56
  Cons ("_::_" [48,47] 48) and
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
    57
  set ("") and
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    58
  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
    59
  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
    60
  update2 ("update") and
63
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
    61
  tm_wf0 ("wf") and
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
    62
  (*is_even ("iseven") 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
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
    70
  exponent ("_\<^bsup>_\<^esup>") and
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
    71
  haltP ("halts") 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 
48
559e5c6e5113 updated to ITP and updated directories
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 38
diff changeset
    74
  tm_comp ("_ \<oplus> _") 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
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
    76
  inv_begin0 ("I\<^isub>0") and
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
    77
  inv_begin1 ("I\<^isub>1") and
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
    78
  inv_begin2 ("I\<^isub>2") and
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
    79
  inv_begin3 ("I\<^isub>3") and
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
    80
  inv_begin4 ("I\<^isub>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
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
    82
  inv_loop1 ("J\<^isub>1") and
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
    83
  inv_loop0 ("J\<^isub>0") and
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
    84
  inv_end1 ("K\<^isub>1") and
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
    85
  inv_end0 ("K\<^isub>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
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
    89
  recf.id ("id\<^raw:\makebox[0mm]{\,\,\,\,>\<^isup>_\<^raw:}>\<^isub>_") and
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
    90
  Pr ("Pr\<^isup>_") and
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
    91
  Cn ("Cn\<^isup>_") and
130
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
    92
  Mn ("Mn\<^isup>_") and
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
    93
  rec_calc_rel ("eval") and 
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
    94
  tm_of_rec ("translate")
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
    95
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
lemma inv_begin_print:
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
    98
  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
    99
        "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
   100
        "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
   101
        "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
   102
        "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
   103
        "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
   104
apply(case_tac [!] tp)
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   105
by (auto)
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   106
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   107
lemma inv1: 
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   108
  shows "0 < n \<Longrightarrow> inv_begin0 n \<mapsto> inv_loop1 n"
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   109
unfolding assert_imp_def
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   110
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
   111
apply(auto)
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   112
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
   113
apply(auto simp add: replicate.simps)
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   114
done
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   115
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   116
lemma inv2: 
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   117
  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
   118
apply(rule ext)
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   119
apply(case_tac x)
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   120
apply(simp add: inv_end1.simps)
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   121
done
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   122
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   123
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
   124
lemma measure_begin_print:
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
   125
  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
   126
        "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
   127
        "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
   128
        "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
   129
by (simp_all)
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
   130
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   131
declare [[show_question_marks = false]]
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   132
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   133
lemma nats2tape:
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   134
  shows "<([]::nat list)> = []"
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   135
  and "<[n]> = <n>"
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   136
  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
   137
  and "<(n, m)> = <n> @ [Bk] @ <m>"
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   138
  and "<[n, m]> = <(n, m)>"
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   139
  and "<n> = Oc \<up> (n + 1)" 
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   140
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
   141
apply(case_tac ns)
97
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_nat_abv)
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   143
done 
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   144
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   145
lemmas HR1 = 
99
fe7a257bdff4 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 98
diff changeset
   146
  Hoare_plus_halt[where ?S.0="R\<iota>" and ?A="p\<^isub>1" and B="p\<^isub>2"]
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   147
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   148
lemmas HR2 =
99
fe7a257bdff4 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 98
diff changeset
   149
  Hoare_plus_unhalt[where ?A="p\<^isub>1" and B="p\<^isub>2"]
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   150
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   151
lemma inv_begin01:
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   152
  assumes "n > 1"
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   153
  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
   154
using assms by auto                          
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   155
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   156
lemma inv_begin02:
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   157
  assumes "n = 1"
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   158
  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
   159
using assms by auto
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   160
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
   161
115
653426ed4b38 started with abacus section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 114
diff changeset
   162
lemma layout:
653426ed4b38 started with abacus section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 114
diff changeset
   163
  shows "layout_of [] = []"
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   164
  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
   165
  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
   166
  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
   167
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
   168
653426ed4b38 started with abacus section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 114
diff changeset
   169
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   170
(*>*)
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   171
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   172
section {* Introduction *}
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   173
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   174
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   175
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   176
text {*
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   177
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
   178
%\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
   179
%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
   180
%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
   181
%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
   182
%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
   183
%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
   184
%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
   185
%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
   186
%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
   187
%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
   188
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
   189
8
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   190
\noindent
79
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   191
Suppose you want to mechanise a proof for whether a predicate @{term
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   192
P}, say, is decidable or not. Decidability of @{text P} usually
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   193
amounts to showing whether \mbox{@{term "P \<or> \<not>P"}} holds. But this
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   194
does \emph{not} work in Isabelle/HOL and other HOL theorem provers,
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   195
since they are based on classical logic where the law of excluded
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   196
middle ensures that \mbox{@{term "P \<or> \<not>P"}} is always provable no
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   197
matter whether @{text P} is constructed by computable means. We hit on
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   198
this limitation previously when we mechanised the correctness proofs
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   199
of two algorithms \cite{UrbanCheneyBerghofer11,WuZhangUrban12}, but
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
   200
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
   201
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
   202
%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
   203
%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
   204
%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
   205
%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
   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
The only satisfying way out of this problem in a theorem prover based
b388dceee892 shortening a bit the paper and updating various things
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 48
diff changeset
   208
on classical logic is to formalise a theory of computability. Norrish
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   209
provided such a formalisation for HOL4. He choose the
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   210
$\lambda$-calculus as the starting point for his formalisation because
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   211
of its ``simplicity'' \cite[Page 297]{Norrish11}.  Part of his
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   212
formalisation is a clever infrastructure for reducing
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   213
$\lambda$-terms. He also established the computational equivalence
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   214
between the $\lambda$-calculus and recursive functions.  Nevertheless
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   215
he concluded that it would be appealing to have formalisations for
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   216
more operational models of computations, such as Turing machines or
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   217
register machines.  One reason is that many proofs in the literature
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   218
use them.  He noted however that \cite[Page 310]{Norrish11}:
8
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   219
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   220
\begin{quote}
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   221
\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
   222
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
   223
daunting prospect.''
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   224
\end{quote}
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   225
c216ae455c90 more on the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   226
\noindent
33
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   227
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
   228
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
   229
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
   230
involved with this work, one has to understand that Turing machine
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   231
programs can be completely \emph{unstructured}, behaving similar to
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   232
Basic programs containing the infamous gotos \cite{Dijkstra68}. This
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   233
precludes in the general case a compositional Hoare-style reasoning
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   234
about Turing programs.  We provide such Hoare-rules for when it
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   235
\emph{is} possible to reason in a compositional manner (which is
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   236
fortunately quite often), but also tackle the more complicated case
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   237
when we translate abacus programs into Turing programs.  This
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   238
reasoning about concrete Turing machine programs is usually
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   239
left out in the informal literature, e.g.~\cite{Boolos87}.
12
dd400b5797e1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 10
diff changeset
   240
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
   241
%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
   242
%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
   243
%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
   244
%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
   245
%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
   246
%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
   247
%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
   248
%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
   249
%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
   250
%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
   251
%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
   252
%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
   253
%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
   254
%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
   255
%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
   256
%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
   257
%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
   258
%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
   259
%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
   260
%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
   261
%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
   262
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   263
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
   264
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
   265
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
   266
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
   267
\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
   268
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
   269
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
   270
``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
   271
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
   272
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
   273
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
   274
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
   275
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
   276
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
   277
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
   278
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
   279
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
   280
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
   281
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
   282
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
   283
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
   284
%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
   285
%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
   286
%\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
   287
%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
   288
%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
   289
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   290
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
   291
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
   292
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
   293
23]{AspertiRicciotti12}:
10
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 9
diff changeset
   294
15
90bc8cccc218 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 13
diff changeset
   295
\begin{quote}\it
13
a7ec585d7f20 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   296
``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
   297
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
   298
annoying.'' 
a7ec585d7f20 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   299
\end{quote}
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   300
15
90bc8cccc218 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 13
diff changeset
   301
\noindent
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   302
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
   303
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
   304
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
   305
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
   306
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
   307
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
   308
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
   309
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
   310
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
   311
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
   312
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
   313
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
   314
\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
   315
\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
   316
turn recursive functions to abacus machines. The universal Turing
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
   317
machine can then be constructed as a recursive function.
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   318
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   319
\smallskip
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   320
\noindent
38
8f8db701f69f updated contribution section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   321
{\bf Contributions:} We formalised in Isabelle/HOL Turing machines following the
8f8db701f69f updated contribution section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   322
description of Boolos et al \cite{Boolos87} where tapes only have blank or
8f8db701f69f updated contribution section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   323
occupied cells. We mechanise the undecidability of the halting problem and
8f8db701f69f updated contribution section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   324
prove the correctness of concrete Turing machines that are needed
8f8db701f69f updated contribution section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   325
in this proof; such correctness proofs are left out in the informal literature.  
72
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   326
For reasoning about Turing machine programs we derive Hoare-rules.
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   327
We also construct the universal Turing machine from \cite{Boolos87} by
79
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   328
translating recursive functions to abacus machines and abacus machines to
38
8f8db701f69f updated contribution section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 37
diff changeset
   329
Turing machines. Since we have set up in Isabelle/HOL a very general computability
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   330
model and undecidability result, we are able to formalise other
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   331
results: we describe a proof of the computational equivalence
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   332
of single-sided Turing machines, which is not given in \cite{Boolos87},
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   333
but needed for example for formalising the undecidability proof of 
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   334
Wang's tiling problem \cite{Robinson71}.
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   335
%We are not aware of any other
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   336
%formalisation of a substantial undecidability problem.
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   337
*}
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   338
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   339
section {* Turing Machines *}
9
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
   340
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
   341
text {* \noindent
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   342
  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
   343
  ``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
   344
  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
   345
  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
   346
  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
   347
  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
   348
  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
   349
  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
   350
  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
   351
  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
   352
  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
   353
  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
   354
  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
   355
  %
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   356
  \begin{center}
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   357
  \begin{tikzpicture}[scale=0.9]
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   358
  \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
   359
  \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
   360
  \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
   361
  \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
   362
  \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
   363
  \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
   364
  \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
   365
  \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
   366
  \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
   367
  \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
   368
  \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
   369
  \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
   370
  \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
   371
  \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
   372
  \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
   373
  \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
   374
  \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
   375
  \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
   376
  \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
   377
  \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
   378
  \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
   379
  \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
   380
  \end{tikzpicture}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   381
  \end{center}
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   382
  
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   383
  \noindent
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   384
  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
   385
  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
   386
  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
   387
  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
   388
  the Turing machine can perform:
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   389
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   390
  \begin{center}
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   391
  \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
   392
  @{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
   393
  & $\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
   394
  \end{tabular}
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   395
  \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
   396
  & $\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
   397
  & $\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
   398
  \end{tabular}
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   399
  \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
   400
  & $\mid$ & @{term Nop} & (do-nothing operation)\\
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   401
  \end{tabular}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   402
  \end{center}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   403
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   404
  \noindent
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 19
diff changeset
   405
  We slightly deviate
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   406
  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
   407
  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
   408
  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
   409
  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
   410
  following tape updating function:
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   411
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   412
  \begin{center}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   413
  \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
   414
  @{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
   415
  @{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
   416
  @{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
   417
  @{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
   418
  @{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
   419
  \end{tabular}
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   420
  \end{center}
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
   421
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   422
  \noindent
89
c67e8ed4c865 updated uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 88
diff changeset
   423
  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
   424
  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
   425
  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
   426
  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
   427
  such that @{term "tl [] == []"} holds. The third clause 
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   428
  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
   429
  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
   430
  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
   431
  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
   432
  in the fourth clause for a right move action. The @{term Nop} operation
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   433
  leaves the the tape unchanged.
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   434
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
   435
  %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
   436
  %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
   437
  %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
   438
  %\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
   439
  %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
   440
  %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
   441
  %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
   442
  %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
   443
  %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
   444
  %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
   445
  %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
   446
  %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
   447
  %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
   448
  %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
   449
  %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
   450
  %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
   451
  %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
   452
50
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   453
  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
   454
  %Given
816e84ca16d6 updated turing_basic by Jian
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 49
diff changeset
   455
  %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
   456
  %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
   457
  %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
   458
  %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
   459
  %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
   460
  %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
   461
  %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
   462
  %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
   463
  %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
   464
  %quite cumbersome to reason about. 
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   465
  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
   466
  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
   467
  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
   468
  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
   469
  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
   470
  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
   471
63
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
   472
  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
   473
  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
   474
  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
   475
  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
   476
  %
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   477
  \begin{equation}
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   478
  \begin{tikzpicture}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   479
  \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
   480
  \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
   481
  {$\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
   482
  = starting state\end{tabular}}}$};
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   483
  
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   484
  \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
   485
  \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
   486
  \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
   487
  \end{tikzpicture}
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   488
  \label{dither}
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   489
  \end{equation}
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   490
  %
29
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 28
diff changeset
   491
  \noindent
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   492
  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
   493
  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
   494
  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
   495
  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
   496
  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
   497
  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
   498
  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
   499
  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
   500
  ``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
   501
  0}-state, but it will always perform a @{term Nop}-operation and
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   502
  remain in the @{text 0}-state.  Unlike Asperti and Riccioti
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   503
  \cite{AspertiRicciotti12}, we have chosen a very concrete
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   504
  representation for programs, because when constructing a universal
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   505
  Turing machine, we need to define a coding function for programs.
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   506
  This can be directly done for our programs-as-lists, but is
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   507
  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
   508
29
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 28
diff changeset
   509
  Given a program @{term p}, a state
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   510
  and the cell being read by the head, we need to fetch
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   511
  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
   512
  the function @{term fetch}
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   513
 
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   514
  \begin{equation}\label{fetch}
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   515
  \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
   516
  \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
   517
  @{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
   518
  \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
   519
  @{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
   520
  \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
   521
  \end{tabular}}
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   522
  \end{equation}
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   523
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   524
  \noindent
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   525
  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
   526
  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
   527
  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
   528
  (@{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
   529
  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
   530
  satisfies the following three properties:
33
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   531
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   532
  \begin{center}
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
   533
  @{thm tm_wf.simps[where p="p" and off="0::nat", simplified, THEN eq_reflection]}
33
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   534
  \end{center}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   535
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   536
  \noindent
75
97eaf7514988 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   537
  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
   538
  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
   539
  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
   540
  the program or being the @{text 0}-state.
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   541
72
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   542
  We need to be able to sequentially compose Turing machine programs. Given our
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   543
  concrete representation, this is relatively straightforward, if
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   544
  slightly fiddly. We use the following two auxiliary functions:
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   545
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   546
  \begin{center}
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   547
  \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}}
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   548
  @{thm (lhs) shift.simps} @{text "\<equiv>"} @{thm (rhs) shift.simps}\\
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   549
  @{thm (lhs) adjust.simps} @{text "\<equiv>"} @{thm (rhs) adjust.simps}\\
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   550
  \end{tabular}
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   551
  \end{center}
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   552
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   553
  \noindent
130
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   554
  The first adds @{text n} to all states, except the @{text 0}-state,
72
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   555
  thus moving all ``regular'' states to the segment starting at @{text
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   556
  n}; the second adds @{term "Suc(length p div 2)"} to the @{text
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   557
  0}-state, thus redirecting all references to the ``halting state''
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   558
  to the first state after the program @{text p}.  With these two
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   559
  functions in place, we can define the \emph{sequential composition}
79
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   560
  of two Turing machine programs @{text "p\<^isub>1"} and @{text "p\<^isub>2"} as
72
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   561
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   562
  \begin{center}
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   563
  @{thm tm_comp.simps[where ?p1.0="p\<^isub>1" and ?p2.0="p\<^isub>2", THEN eq_reflection]}
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   564
  \end{center}
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   565
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   566
  \noindent
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   567
  %This means @{text "p\<^isub>1"} is executed first. Whenever it originally
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   568
  %transitioned to the @{text 0}-state, it will in the composed program transition to the starting
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   569
  %state of @{text "p\<^isub>2"} instead. All the states of @{text "p\<^isub>2"}
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   570
  %have been shifted in order to make sure that the states of the composed 
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   571
  %program @{text "p\<^isub>1 \<oplus> p\<^isub>2"} still only ``occupy'' 
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   572
  %an initial segment of the natural numbers.
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   573
49f8e5cf82c5 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 71
diff changeset
   574
  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
   575
  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
   576
  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
   577
  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
   578
  "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
   579
  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
   580
  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
   581
  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
   582
  the function @{term step}
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   583
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   584
  \begin{center}
63
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
   585
  \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
   586
  @{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
   587
  & & @{text "in (s', update (l, r) a)"}
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   588
  \end{tabular}
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   589
  \end{center}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   590
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   591
  \noindent
104
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   592
  where @{term "read r"} returns the head of the list @{text r}, or if
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   593
  @{text r} is empty it returns @{term Bk}.  It is impossible in
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   594
  Isabelle/HOL to lift the @{term step}-function in order to realise a
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   595
  general evaluation function for Turing machines programs. The reason
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   596
  is that functions in HOL-based provers need to be terminating, and
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   597
  clearly there are programs that are not. We can however define a
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   598
  recursive evaluation function that performs exactly @{text n} steps:
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   599
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
   600
  \begin{center}
30
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   601
  \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 29
diff changeset
   602
  @{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
   603
  @{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
   604
  \end{tabular}
22
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 21
diff changeset
   605
  \end{center}
18
a961c2e4dcea updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   606
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   607
  \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
   608
  \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
   609
  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
   610
  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
   611
  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
   612
  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
   613
  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
   614
  @{text n} is reached. 
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   615
  
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   616
  \begin{figure}[t]
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   617
  \begin{center}
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   618
  \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
   619
  \begin{tabular}[t]{@ {}l@ {}}
92
b9d0dd18c81e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 91
diff changeset
   620
  @{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
   621
  \hspace{2mm}@{text "["}@{text "(W\<^bsub>Bk\<^esub>, 0), (R, 2), (R, 3),"}\\
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   622
  \hspace{2mm}\phantom{@{text "["}}@{text "(R, 2), (W1, 3), (L, 4),"}\\
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   623
  \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
   624
  \end{tabular}
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   625
  &
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   626
  \begin{tabular}[t]{@ {}l@ {}}
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   627
  @{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
   628
  \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
   629
  \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
   630
  \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
   631
  \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
   632
  \end{tabular}
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   633
  &
87
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   634
  \begin{tabular}[t]{@ {}l@ {}}
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   635
  @{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
   636
  \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
   637
  \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
   638
  \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
   639
  \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
   640
  \end{tabular}
87
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   641
  \end{tabular}\\[2mm]
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   642
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   643
  \begin{tikzpicture}[scale=0.7]
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   644
  \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
   645
  \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
   646
  \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
   647
  \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
   648
  \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
   649
  \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
   650
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   651
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   652
  \begin{scope}[shift={(0.5,0)}]
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   653
  \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
   654
  \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
   655
  \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
   656
  \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
   657
  \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
   658
  \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
   659
  \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
   660
  \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
   661
  \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
   662
  \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
   663
  \end{scope}
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   664
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   665
  \begin{scope}[shift={(2.9,0)}]
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   666
  \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
   667
  \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
   668
  \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
   669
  \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
   670
  \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
   671
  \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
   672
  \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
   673
  \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
   674
  \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
   675
  \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
   676
  \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
   677
  \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
   678
  \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
   679
  \end{scope}
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   680
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   681
  \begin{scope}[shift={(6.8,0)}]
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   682
  \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
   683
  \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
   684
  \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
   685
  \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
   686
  \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
   687
  \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
   688
  \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
   689
  \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
   690
  \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
   691
  \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
   692
  \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
   693
  \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
   694
  \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
   695
  \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
   696
  \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
   697
  \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
   698
  \end{scope}
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   699
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   700
  \begin{scope}[shift={(11.7,0)}]
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   701
  \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
   702
  \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
   703
  \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
   704
  \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
   705
  \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
   706
  \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
   707
  \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
   708
  \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
   709
  \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
   710
  \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
   711
  \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
   712
  \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
   713
  \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
   714
  \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
   715
  \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
   716
  \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
   717
  \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
   718
  \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
   719
  \end{scope}
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   720
  \end{tikzpicture}\\[-8mm]\mbox{} 
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   721
  \end{center}
94
aeaf1374dc67 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 93
diff changeset
   722
  \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
   723
  (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
   724
  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
   725
  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
   726
  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
   727
  \label{copy}
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   728
  \end{figure}
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   729
87
cf6e89b5f702 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 86
diff changeset
   730
104
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   731
  We often need to restrict tapes to be in standard form, which means 
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   732
  the left list of the tape is either empty or only contains @{text "Bk"}s, and 
130
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   733
  the right list contains some ``clusters'' of @{text "Oc"}s separated by single 
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   734
  blanks. To make this formal we define the following overloaded function
104
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   735
  encoding natural numbers into lists of @{term "Oc"}s and @{term Bk}s.
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   736
  % 
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   737
  \begin{equation}
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   738
  \mbox{\begin{tabular}[t]{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}}
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   739
  @{thm (lhs) nats2tape(6)} & @{text "\<equiv>"} & @{thm (rhs) nats2tape(6)}\\
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   740
  @{thm (lhs) nats2tape(4)} & @{text "\<equiv>"} & @{thm (rhs) nats2tape(4)}\\
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   741
  \end{tabular}\hspace{6mm}
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   742
  \begin{tabular}[t]{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}}
85
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   743
  @{thm (lhs) nats2tape(1)} & @{text "\<equiv>"} & @{thm (rhs) nats2tape(1)}\\
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   744
  @{thm (lhs) nats2tape(2)} & @{text "\<equiv>"} & @{thm (rhs) nats2tape(2)}\\
e6f395eccfe7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 84
diff changeset
   745
  @{thm (lhs) nats2tape(3)} & @{text "\<equiv>"} & @{thm (rhs) nats2tape(3)}
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   746
  \end{tabular}}\label{standard}
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   747
  \end{equation}
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   748
  %
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   749
  \noindent
104
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   750
  A \emph{standard tape} is then of the form @{text "(Bk\<^isup>l,\<langle>[n\<^isub>1,...,n\<^isub>m]\<rangle>)"} for some @{text l} 
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   751
  and @{text "n\<^bsub>1...m\<^esub>"}. Note that the head in a standard tape ``points'' to the 
94
aeaf1374dc67 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 93
diff changeset
   752
  leftmost @{term "Oc"} on the tape. Note also that the natural number @{text 0} 
aeaf1374dc67 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 93
diff changeset
   753
  is represented by a single filled cell on a standard tape, @{text 1} by two filled cells and so on.
79
bc54c5e648d7 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 78
diff changeset
   754
104
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   755
  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
   756
  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
   757
  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
   758
  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
   759
  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
   760
  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
   761
  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
   762
  \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
   763
91
2068654bdf54 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 90
diff changeset
   764
  \begin{equation}
2068654bdf54 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 90
diff changeset
   765
  \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
   766
  @{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
   767
  \end{tabular}}\label{tcopy}
2068654bdf54 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 90
diff changeset
   768
  \end{equation}
73
a673539f2f75 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 72
diff changeset
   769
32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 31
diff changeset
   770
  \noindent
102
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   771
  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
   772
  correctness proofs, we introduce the notion of total correctness
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   773
  defined in terms of \emph{Hoare-triples}, written @{term "{P} p
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   774
  {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
   775
  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
   776
  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
   777
  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
   778
  similar to the notion of \emph{realisability} in \cite{AspertiRicciotti12}. We
102
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   779
  also have \emph{Hoare-pairs} of the form @{term "{P} p \<up>"}
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   780
  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
   781
  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
   782
  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
   783
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   784
  \begin{center}
76
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   785
  \begin{tabular}{@ {}c@ {\hspace{4mm}}c@ {}}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   786
  \begin{tabular}[t]{@ {}l@ {}}
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   787
  \colorbox{mygrey}{@{thm (lhs) Hoare_halt_def}} @{text "\<equiv>"}\\[1mm]
76
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   788
  \hspace{5mm}@{text "\<forall>"} @{term "(l, r)"}.\\
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   789
  \hspace{7mm}if @{term "P (l, r)"} holds then\\
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   790
  \hspace{7mm}@{text "\<exists>"} @{term n}. such that\\
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   791
  \hspace{7mm}@{text "is_final (steps (1, (l, r)) p n)"} \hspace{1mm}@{text "\<and>"}\\
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   792
  \hspace{7mm}@{text "Q holds_for (steps (1, (l, r)) p n)"}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   793
  \end{tabular} &
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   794
  \begin{tabular}[t]{@ {}l@ {}}
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   795
  \colorbox{mygrey}{@{thm (lhs) Hoare_unhalt_def}} @{text "\<equiv>"}\\[1mm]
76
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   796
  \hspace{5mm}@{text "\<forall>"} @{term "(l, r)"}.\\ 
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   797
  \hspace{7mm}if @{term "P (l, r)"} holds then\\
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   798
  \hspace{7mm}@{text "\<forall>"} @{term n}. @{text "\<not> is_final (steps (1, (l, r)) p n)"}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   799
  \end{tabular}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   800
  \end{tabular}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   801
  \end{center}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   802
  
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   803
  \noindent
102
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   804
  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
   805
fe7a257bdff4 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 98
diff changeset
   806
  \begin{equation}
fe7a257bdff4 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 98
diff changeset
   807
  @{thm[mode=Rule] Hoare_consequence}
fe7a257bdff4 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 98
diff changeset
   808
  \end{equation}
fe7a257bdff4 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 98
diff changeset
   809
104
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   810
  \noindent
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   811
  where
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   812
  @{term "P' \<mapsto> 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
   813
  @{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
   814
90
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   815
  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
   816
  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
   817
  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
   818
  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
   819
  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
   820
  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
   821
  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
   822
  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
   823
  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
   824
  "tcopy_end"}.
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   825
130
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   826
  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
   827
  @{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
   828
  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
   829
  @{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
   830
  below.
90
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   831
76
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   832
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   833
  \begin{center}
81
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   834
  \begin{tabular}{l@ {\hspace{3mm}}lcl}
2e9881578cb2 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 80
diff changeset
   835
  & \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
   836
  \raisebox{2mm}{halting case:} &
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   837
  \begin{tikzpicture}[scale=0.8]
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   838
  \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
   839
  \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
   840
  \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
   841
  \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
   842
  \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
   843
  \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
   844
  \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
   845
  \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
   846
  \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
   847
  \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
   848
  \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
   849
  \end{tikzpicture} 
90
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   850
  & \raisebox{2mm}{$\;\;\large\Rightarrow\;\;$} &
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   851
  \begin{tikzpicture}[scale=0.8]
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   852
  \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
   853
  \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
   854
  \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
   855
  \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
   856
  \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
   857
  \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
   858
  \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
   859
  \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
   860
  \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
   861
  \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
   862
  \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
   863
  \end{tikzpicture}\\
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   864
90
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   865
  \raisebox{2mm}{non-halting case:} &
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   866
  \begin{tikzpicture}[scale=0.8]
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   867
  \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
   868
  \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
   869
  \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
   870
  \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
   871
  \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
   872
  \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
   873
  \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
   874
  \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
   875
  \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
   876
  \end{tikzpicture} 
90
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   877
  & \raisebox{2mm}{$\;\;\large\Rightarrow\;\;$} &
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   878
  \raisebox{2mm}{loops}
80
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   879
  \end{tabular}
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   880
  \end{center}
eb589fa73fc1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 79
diff changeset
   881
90
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   882
  \noindent
91
2068654bdf54 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 90
diff changeset
   883
  We can prove the following Hoare-statements:
90
d2f4b775cd15 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 89
diff changeset
   884
 
76
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   885
  \begin{center}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   886
  \begin{tabular}{l}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   887
  @{thm dither_halts}\\
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   888
  @{thm dither_loops}
04399b471108 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 75
diff changeset
   889
  \end{tabular}
34
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   890
  \end{center}
22e5804b135c updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   891
77
04e5850818c8 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 76
diff changeset
   892
  \noindent
102
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   893
  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
   894
  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
   895
96
bd320b5365e2 updated uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 94
diff changeset
   896
  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
   897
  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
   898
  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
   899
  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
   900
  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
   901
  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
   902
  composed programs.
75
97eaf7514988 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 73
diff changeset
   903
91
2068654bdf54 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 90
diff changeset
   904
  \begin{center}
94
aeaf1374dc67 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 93
diff changeset
   905
  \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
   906
  $\inferrule*[Right=@{thm (prem 3) HR1}]
aeaf1374dc67 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 93
diff changeset
   907
  {@{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
   908
  {@{thm (concl) HR1}}
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   909
  $ &
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   910
  $
94
aeaf1374dc67 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 93
diff changeset
   911
  \inferrule*[Right=@{thm (prem 3) HR2}]
aeaf1374dc67 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 93
diff changeset
   912
  {@{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
   913
  {@{thm (concl) HR2}}
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   914
  $
91
2068654bdf54 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 90
diff changeset
   915
  \end{tabular}
2068654bdf54 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 90
diff changeset
   916
  \end{center}
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   917
91
2068654bdf54 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 90
diff changeset
   918
  \noindent
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   919
  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
   920
  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
   921
  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
   922
  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
   923
  needed in order to ensure that the redirection of the halting and
107
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
   924
  initial state in @{term "p\<^isub>1"} and @{term "p\<^isub>2"}, respectively, match
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
   925
  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
   926
  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
   927
  @{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
   928
  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
   929
  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
   930
  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
   931
  "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
   932
  our formalisation.
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
   933
107
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
   934
  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
   935
  @{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
   936
  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
   937
  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
   938
104
01f688735b9b updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 103
diff changeset
   939
  \begin{figure}[t]
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   940
  \begin{center}
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
   941
  \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
   942
  \hline
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   943
  @{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
   944
  @{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
   945
  @{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
   946
  @{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
   947
  @{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
   948
                                &             & @{thm (rhs) inv_begin02}\smallskip \\
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
   949
   \hline
100
dfe852aacae6 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 99
diff changeset
   950
  @{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
   951
                               &             & @{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
   952
  @{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
   953
   \hline
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
   954
  @{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
   955
  @{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
   956
  \hline
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   957
  \end{tabular}
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   958
  \end{center}
106
c3155e9e4f63 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 105
diff changeset
   959
  \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
   960
  @{term tcopy_begin}. Below, the invariants only for the starting and halting states of
c3155e9e4f63 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 105
diff changeset
   961
  @{term tcopy_loop} and @{term tcopy_end} are shown. In each invariant the parameter 
c3155e9e4f63 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 105
diff changeset
   962
  @{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
   963
  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
   964
  \end{figure}
96
bd320b5365e2 updated uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 94
diff changeset
   965
bd320b5365e2 updated uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 94
diff changeset
   966
  \begin{center}
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   967
  \begin{tabular}{rcl}
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   968
  @{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
   969
  @{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
   970
  & & @{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
   971
 & & @{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
   972
  & & @{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
   973
   & & @{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
   974
  & & @{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
   975
  \end{tabular}
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   976
  \end{center}
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   977
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   978
  \noindent
102
cdef5b1072fe updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 100
diff changeset
   979
  This invariant depends on @{term n} representing the number of
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   980
  @{term Oc}s@{text "+1"} (or encoded number) 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
   981
  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
   982
  (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
   983
  @{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
   984
  for @{term "tcopy_begin"}. 
97
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   985
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
   986
  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
   987
  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
   988
  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
   989
  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
   990
  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
   991
  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
   992
  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
   993
d6f04e3e9894 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 96
diff changeset
   994
  \begin{center}
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
   995
  \begin{tabular}{rcl}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
   996
  @{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
   997
  @{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
   998
  & & @{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
   999
  @{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
  1000
 & & @{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
  1001
  & & @{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
  1002
  \end{tabular}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1003
  \end{center}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1004
  
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1005
  \noindent
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1006
  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
  1007
  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
  1008
  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
  1009
  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
  1010
  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
  1011
   
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1012
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1013
  \begin{center}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1014
  @{thm (concl) begin_correct}\hspace{6mm}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1015
  @{thm (concl) loop_correct}\hspace{6mm}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1016
  @{thm (concl) end_correct}
96
bd320b5365e2 updated uncomputable
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 94
diff changeset
  1017
  \end{center}
36
4b35e0e0784b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 35
diff changeset
  1018
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1019
  \noindent 
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1020
  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
  1021
  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
  1022
  the halting state of @{term tcopy_begin} implies the invariant of
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1023
  the starting state of @{term tcopy_loop}, that is @{term "inv_begin0
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1024
  n \<mapsto> 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
  1025
  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
  1026
  of @{term tcopy}:
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
  1027
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1028
  \begin{center}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1029
  @{thm tcopy_correct}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1030
  \end{center}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1031
107
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1032
  \noindent
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1033
  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
  1034
  @{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
  1035
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1036
  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
  1037
  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
  1038
  @{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
  1039
  number is defined as
24
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
  1040
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1041
  \begin{center}
99
fe7a257bdff4 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 98
diff changeset
  1042
  @{thm haltP_def}
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1043
  \end{center}
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1044
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1045
  \noindent
107
c2a7a99bf554 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 106
diff changeset
  1046
  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
  1047
  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
  1048
  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
  1049
  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
  1050
  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
  1051
  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
  1052
  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
  1053
  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
  1054
  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
  1055
  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
  1056
  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
  1057
  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
  1058
  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
  1059
  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
  1060
  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
  1061
  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
  1062
  natural numbers @{term ns}:
106
c3155e9e4f63 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 105
diff changeset
  1063
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1064
  \begin{center}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1065
  \begin{tabular}{r}
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1066
  @{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
  1067
  @{thm (concl) uncomputable.h_case}\\
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1068
  
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1069
  @{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
  1070
  @{thm (concl) uncomputable.nh_case}
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1071
  \end{tabular}
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1072
  \end{center}
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1073
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1074
  \noindent
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1075
  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
  1076
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1077
  \begin{center}
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1078
  @{thm tcontra_def}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1079
  \end{center}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1080
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1081
  \noindent
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1082
  Suppose @{thm (prem 1) "tcontra_halt"} holds. Given the invariants @{text "P\<^isub>1"}\ldots@{text "P\<^isub>3"} 
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1083
  shown on the 
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1084
  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
  1085
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1086
  \begin{center}\small
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1087
  \begin{tabular}{@ {}c@ {\hspace{-10mm}}c@ {}}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1088
  \begin{tabular}[t]{@ {}l@ {}}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1089
  @{term "P\<^isub>1 \<equiv> \<lambda>tp. tp = ([]::cell list, <code_tcontra>)"}\\
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1090
  @{term "P\<^isub>2 \<equiv> \<lambda>tp. tp = ([Bk], <(code_tcontra, code_tcontra)>)"}\\
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1091
  @{term "P\<^isub>3 \<equiv> \<lambda>tp. \<exists>k. tp = (Bk \<up> k, <0::nat>)"}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1092
  \end{tabular}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1093
  &
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1094
  \begin{tabular}[b]{@ {}l@ {}}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1095
  \raisebox{-20mm}{$\inferrule*{
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1096
    \inferrule*{@{term "{P\<^isub>1} tcopy {P\<^isub>2}"} \\ @{term "{P\<^isub>2} H {P\<^isub>3}"}}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1097
    {@{term "{P\<^isub>1} (tcopy |+| H) {P\<^isub>3}"}}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1098
    \\ @{term "{P\<^isub>3} dither \<up>"}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1099
   }
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1100
   {@{term "{P\<^isub>1} tcontra \<up>"}}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1101
  $}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1102
  \end{tabular}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1103
  \end{tabular}
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1104
  \end{center}
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1105
105
2cae8a39803e updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 104
diff changeset
  1106
  \noindent
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1107
  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
  1108
  with @{term "<(code tcontra)>"} halts.
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1109
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1110
  Suppose @{thm (prem 1) "tcontra_unhalt"} holds. Again, given the invariants 
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1111
  @{text "Q\<^isub>1"}\ldots@{text "Q\<^isub>3"}
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1112
  shown on the 
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1113
  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
  1114
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1115
  \begin{center}\small
116
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
  1116
  \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
  1117
  \begin{tabular}[t]{@ {}l@ {}}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1118
  @{term "Q\<^isub>1 \<equiv> \<lambda>tp. tp = ([]::cell list, <code_tcontra>)"}\\
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1119
  @{term "Q\<^isub>2 \<equiv> \<lambda>tp. tp = ([Bk], <(code_tcontra, code_tcontra)>)"}\\
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1120
  @{term "Q\<^isub>3 \<equiv> \<lambda>tp. \<exists>k. tp = (Bk \<up> k, <1::nat>)"}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1121
  \end{tabular}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1122
  &
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1123
  \begin{tabular}[t]{@ {}l@ {}}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1124
  \raisebox{-20mm}{$\inferrule*{
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1125
    \inferrule*{@{term "{Q\<^isub>1} tcopy {Q\<^isub>2}"} \\ @{term "{Q\<^isub>2} H {Q\<^isub>3}"}}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1126
    {@{term "{Q\<^isub>1} (tcopy |+| H) {Q\<^isub>3}"}}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1127
    \\ @{term "{Q\<^isub>3} dither {Q\<^isub>3}"}
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1128
   }
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1129
   {@{term "{Q\<^isub>1} tcontra {Q\<^isub>3}"}}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1130
  $}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1131
  \end{tabular}
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1132
  \end{tabular}
93
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1133
  \end{center}
f2bda6ba4952 updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 92
diff changeset
  1134
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1135
  \noindent
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1136
  This time the Hoare-triple states that @{term tcontra} terminates 
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1137
  with the ``output'' @{term "<(1::nat)>"}. In both case we come
130
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1138
  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
  1139
  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
  1140
  whether Turing machines terminate.
9
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
  1141
*}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
  1142
63
35fe8fe12e65 small updates
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 61
diff changeset
  1143
17
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
  1144
section {* Abacus Machines *}
66cebc19ef18 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
  1145
25
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
  1146
text {*
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
  1147
  \noindent
112
fea23f9a9d85 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 111
diff changeset
  1148
  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
  1149
  for making it less laborious to write Turing machine
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1150
  programs. Abacus machines operate over a set of registers @{text "R\<^isub>0"},
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1151
  @{text "R\<^isub>1"}, \ldots{}, @{text "R\<^isub>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
  1152
  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
  1153
  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
  1154
  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
  1155
  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
  1156
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
  1157
  \begin{center}
111
dfc629cd11de made uncomputable compatible with abacus
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 109
diff changeset
  1158
  \begin{tabular}{rcl@ {\hspace{10mm}}l}
117
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1159
  @{text "i"} & $::=$  & @{term "Inc R\<iota>"} & increment register $R$ by one\\
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1160
  & $\mid$ & @{term "Dec R\<iota> l"} & if content of $R$ is non-zero, then decrement it by one\\
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1161
  & & & otherwise jump to instruction $l$\\
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1162
  & $\mid$ & @{term "Goto l"} & jump to instruction $l$
25
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
  1163
  \end{tabular}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
  1164
  \end{center}
27
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
  1165
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
  1166
  \noindent
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1167
  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
  1168
  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
  1169
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
  1170
  \begin{center}
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1171
  @{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
  1172
  \end{center}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
  1173
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
  1174
  \noindent
113
8ef94047e6e2 abacus section updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 112
diff changeset
  1175
  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
  1176
  then execute one instructions after the other, unless there is a jump.  For
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1177
  example the second instruction @{term "Goto 0"} means
122
db518aba152a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 121
diff changeset
  1178
  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
  1179
  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
  1180
  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
  1181
  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
  1182
  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
  1183
  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
  1184
  all registers.
844e149696d4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 119
diff changeset
  1185
  By convention
119
892a3d7493aa updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 118
diff changeset
  1186
  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
  1187
  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
  1188
  
115
653426ed4b38 started with abacus section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 114
diff changeset
  1189
  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
  1190
  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
  1191
  standard tapes (see definition shown in \eqref{standard}). Because of the jumps in abacus programs, it
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1192
  is impossible to build a Turing machine programs out of components 
123
d8f04ed7489e updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 122
diff changeset
  1193
  using our @{text "\<oplus>"}-operation shown in the previous section.
124
bda714532263 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
  1194
  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
  1195
  abacus program as follows
117
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1196
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1197
  \begin{center}
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1198
  \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
  1199
  @{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
  1200
  @{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
  1201
  @{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
  1202
  @{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
  1203
  \end{tabular}
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1204
  \end{center}
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1205
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1206
  \noindent
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1207
  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
  1208
  are needed to translate each abacus instruction. This information
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1209
  is needed in order to calculate the state where the Turing program
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1210
  code of one abacus instruction ends.
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1211
  The @{text Goto}
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1212
  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
  1213
  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
  1214
653426ed4b38 started with abacus section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 114
diff changeset
  1215
  \begin{center}
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1216
  @{text "tm_of_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
  1217
  \end{center}
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1218
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1219
  \noindent
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1220
  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
  1221
  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
  1222
  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
  1223
  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
  1224
  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
  1225
  @{text "R"}. This needs a machine
117
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1226
  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
  1227
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1228
  \begin{center}
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1229
  \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
  1230
  @{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
  1231
  @{thm (lhs) findnth.simps(2)} & @{text "\<equiv>"}\\
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1232
  \multicolumn{3}{@ {}l@ {}}{\hspace{8mm}@{thm (rhs) findnth.simps(2)}}\\
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1233
  \end{tabular}
115
653426ed4b38 started with abacus section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 114
diff changeset
  1234
  \end{center}
653426ed4b38 started with abacus section
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 114
diff changeset
  1235
117
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1236
  \noindent
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1237
  Then we need to increase the ``number'' on the tape by one,
124
bda714532263 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
  1238
  and adjust the following ``registers''. By adjusting we only need to 
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1239
  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
  1240
  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
  1241
  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
  1242
  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
  1243
  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
  1244
  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
  1245
  corresponding register is @{text 0}. For this we have a Turing machine program
124
bda714532263 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
  1246
  with @{text 16} states (again details are omitted). 
117
b27f4bd98078 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 116
diff changeset
  1247
121
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
  1248
  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
  1249
  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
  1250
  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
  1251
  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
  1252
  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
  1253
  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
  1254
  ``mops up'' the tape (cleaning all @{text Oc}s) except for the
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1255
  @{term Oc}s of the last number represented on the tape. This needs
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1256
  a Turing machine program with @{text "2 * R + 12"} states, assuming @{text R}
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1257
  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
  1258
124
bda714532263 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
  1259
  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
  1260
  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
  1261
  @{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
  1262
  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
  1263
  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
  1264
  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
  1265
  prove as invariant that it performs
bda714532263 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 123
diff changeset
  1266
  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
  1267
  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
  1268
  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
  1269
  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
  1270
  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
  1271
  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
  1272
  
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
  1273
  %We were able to simplify the proof somewhat
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
  1274
*}
29
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 28
diff changeset
  1275
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 28
diff changeset
  1276
121
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
  1277
section {* Recursive Functions and a Universal Turing Machine *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
  1278
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
  1279
text {*
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1280
  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
  1281
  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
  1282
  function. This is different from Norrish \cite{Norrish11} who gives a universal 
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1283
  function for Church numbers, and also from Asperti and Ricciotti 
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1284
  \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
  1285
  directly, but for simulating Turing machines with a more restricted alphabet.
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1286
  \emph{Recursive functions} @{term r} are defined as the datatype
121
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
  1287
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1288
  \begin{center}
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1289
  \begin{tabular}{c@ {\hspace{4mm}}c}
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1290
  \begin{tabular}{rcl@ {\hspace{4mm}}l}
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1291
  @{term r} & @{text "::="} & @{term z} & (zero-functions)\\
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1292
            & @{text "|"}   & @{term s} & (successor-function)\\
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1293
            & @{text "|"}   & @{term "id n m"} & (projection)\\
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1294
  \end{tabular} &
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1295
  \begin{tabular}{cl@ {\hspace{4mm}}l}
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1296
  @{text "|"} & @{term "Cn n r rs"} & (composition)\\
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1297
  @{text "|"} & @{term "Pr n r\<^isub>1 r\<^isub>2"} & (primitive recursion)\\
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1298
  @{text "|"} & @{term "Mn n r"} & (minimisation)\\
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1299
  \end{tabular}
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1300
  \end{tabular}
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1301
  \end{center}
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1302
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1303
  \noindent 
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1304
  where @{text n} indicates the function expects @{term n} arguments
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1305
  (@{text z} and @{term s} expect one argument), and @{text rs} stands
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1306
  for a list of recursive functions. Since we know in each case 
130
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1307
  the arity, say @{term l}, we can define an inductive evaluation relation that  
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1308
  relates a recursive function @{text r} and a list @{term ns} of natural numbers of length @{text l},
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1309
  to what the result of the recursive function is, say @{text n}---we omit the straightforward
133
ca7fb6848715 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 132
diff changeset
  1310
  definition of @{term "rec_calc_rel r ns n"}. Because of space reasons, we also omit the 
130
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1311
  definition of translating
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1312
  recursive functions into abacus programs. We can prove the following
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1313
  theorem about the translation: Assuming 
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1314
  @{thm (prem 1) recursive_compile_to_tm_correct3[where recf="r" and args="ns" and r="n"]}
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1315
  then the following Hoare-triple holds
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1316
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1317
  \begin{center}
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1318
  @{thm (concl) recursive_compile_to_tm_correct3[where recf="r" and args="ns" and r="n"]}
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1319
  \end{center}
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1320
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1321
  \noindent
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1322
  which means that if the recursive function @{text r} with arguments @{text ns} evaluates
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1323
  to @{text n}, then the corresponding Turing machine @{term "tm_of_rec r"} if started
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1324
  with the standard tape @{term "([Bk, Bk], <ns::nat list>)"} will terminate
132
264ff7014657 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
  1325
  with the standard tape @{term "(Bk \<up> k, <n::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
  1326
1e89c65f844b added UTM
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
  1327
  and the also the definition of the
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1328
  universal function (we refer the reader to our formalisation).
134
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1329
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1330
  \cite[Page 32]{Boolos87}
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1331
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1332
  \begin{quote}\it
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1333
  ``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
  1334
  are represented initially on the tape, then the machine either will never halt, 
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1335
  or will halt in some nonstandard configuration\ldots''
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1336
  \end{quote}
126
0b302c0b449a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
  1337
  
134
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1338
  This means that if you encode the plus function but only give one argument,
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1339
  then the TM will either loop {\bf or} stop with a non-standard tape
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1340
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1341
  But in the definition of the universal function the TMs will never stop
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1342
  with non-standard tapes. 
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1343
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1344
  SO the following TM calculates something according to def from chap 8,
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1345
  but not with chap 3. For example:
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1346
  
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1347
  \begin{center}
135
ba63ba7d282b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 134
diff changeset
  1348
  @{term "[(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
  1349
  \end{center}
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1350
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1351
  If started with @{term "([], [Oc])"} it halts with the
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1352
  non-standard tape @{term "([Oc], [])"} according to the definition in Chap 3;
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1353
  but with @{term "([], [Oc])"} according to def Chap 8
140
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1354
*}
134
f47f1ef313d1 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 133
diff changeset
  1355
140
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1356
section {* XYZ *}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1357
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1358
text {*
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1359
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
  1360
\begin{enumerate}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1361
    \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
  1362
        $
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1363
        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
  1364
        $.
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1365
    \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
  1366
\end{enumerate}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1367
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1368
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
  1369
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1370
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1371
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
  1372
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1373
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
  1374
\begin{quote}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1375
(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
  1376
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
  1377
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
  1378
\end{quote}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1379
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
  1380
	\begin{equation}\label{stdh_def}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1381
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
  1382
\end{equation}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1383
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
  1384
    \begin{equation}\label{contrived_tm}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1385
        [(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
  1386
    \end{equation}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1387
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
  1388
\[
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1389
(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
  1390
\]
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1391
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
  1392
\begin{equation}\label{fetch-def}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1393
\begin{aligned}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1394
	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
  1395
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
  1396
\end{aligned}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1397
\end{equation}
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1398
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
  1399
\[
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1400
(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
  1401
\]
7f5243700f25 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 138
diff changeset
  1402
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
  1403
*}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
  1404
121
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
  1405
(*
13
a7ec585d7f20 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  1406
section {* Wang Tiles\label{Wang} *}
7
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1407
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1408
text {*
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1409
  Used in texture mapings - graphics
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1410
*}
121
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
  1411
*)
7
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1412
121
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 120
diff changeset
  1413
section {* Conclusion *}
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1414
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1415
text {*
136
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1416
  In previous works we were unable to formalise results about
138
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1417
  computability because in Isabelle/HOL we cannot represent the
136
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1418
  decidability of a predicate @{text P}, say, as the formula @{term "P
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1419
  \<or> \<not>P"}. For reasoning about computability we need to formalise a
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1420
  concrete model of computations. We could have followed Norrish
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1421
  \cite{Norrish11} using the $\lambda$-calculus as the starting point,
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1422
  but then we would have to reimplement his infrastructure for
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1423
  reducing $\lambda$-terms on the ML-level. We would still need to
138
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1424
  connect his work to Turing machines for proofs that make essential use of them
136
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1425
  (for example the proof of Wang's tiling problem \cite{Robinson71}).
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1426
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1427
  We therefore have formalised Turing machines and the main
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1428
  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
  1429
  et al \cite{Boolos87}.  For this we did not need to implement
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1430
  anything on the ML-level of Isabelle/HOL. While formalising
138
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1431
  \cite{Boolos87} we have found an inconsistency in Bolos et al's 
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1432
  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
  1433
  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
  1434
  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
  1435
  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
  1436
  Turing machines will \emph{not} halt unless the tape is in standard
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1437
  form. Following in the footsteps of another paper \cite{Nipkow98}
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1438
  formalising the results from a semantics textbook, we could
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1439
  therefore have titled our paper ``Boolos et al are (almost)
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1440
  Right''. We have not attempted to formalise everything precisely as
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1441
  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
  1442
  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
  1443
  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
  1444
  non-standard, but very much suited to a formalisation in a theorem
8fa9e018abe4 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 135
diff changeset
  1445
  prover where the @{term steps}-function need to be total.
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
  1446
138
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1447
  The most closely related work is by Norrish \cite{Norrish11}, and by
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1448
  Asperti and Ricciotti \cite{AspertiRicciotti12}. Norrish mentions
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1449
  that formalising Turing machines would be a ``daunting prospect''
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1450
  \cite[Page 310]{Norrish11}. While $\lambda$-terms indeed lead to
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1451
  some slick mechanised proofs, our experience is that Turing machines
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1452
  are not too daunting if one is only concerned with formalising the
125
1ce74a77fa2a conclusion
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 124
diff changeset
  1453
  undecidability of the halting problem for Turing machines.  This
1ce74a77fa2a conclusion
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 124
diff changeset
  1454
  took us around 1500 loc of Isar-proofs, which is just one-and-a-half
138
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1455
  times of a mechanised proof pearl about the Myhill-Nerode
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1456
  theorem. So our conclusion is it not as daunting as we estimated
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1457
  when reading the paper by Norrish \cite{Norrish11}. The work
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1458
  involved with constructing a universal Turing machine via recursive
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1459
  functions and abacus machines, on the other hand, is not a project
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1460
  one wants to undertake too many times (our formalisation of abacus
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1461
  machines and their correct translation is approximately 4300 loc;
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1462
  recursive functions XX loc and the universal Turing machine XX loc).
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
  1463
  
125
1ce74a77fa2a conclusion
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 124
diff changeset
  1464
  Our work was also very much inspired by the formalisation of Turing
1ce74a77fa2a conclusion
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 124
diff changeset
  1465
  machines by Asperti and Ricciotti \cite{AspertiRicciotti12} in the
1ce74a77fa2a conclusion
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 124
diff changeset
  1466
  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
  1467
  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
  1468
  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
  1469
  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
  1470
  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
  1471
  (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
  1472
  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
  1473
  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
  1474
  for them.
125
1ce74a77fa2a conclusion
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 124
diff changeset
  1475
  
138
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1476
  For us the most interesting aspect of our work are the correctness
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1477
  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
  1478
  theory often leave the constructions of particular Turing machines
138
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1479
  as exercise to the reader, as for example \cite{Boolos87}, deeming
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1480
  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
  1481
  presentations leave out any arguments why these Turing machines
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1482
  should be correct.  This means the reader or formalsiser is left
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1483
  with the task of finding appropriate invariants and measures for
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1484
  showing 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
  1485
  Whenever we can use Hoare-style reasoning, the invariants are
138
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1486
  relatively straightforward and much smaller than for example the
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1487
  invariants used by Myreen in a correctness proof of a garbage collector
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1488
  written in machine code \cite[Page 76]{Myreen09}. Howvere, the invariant 
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1489
  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
  1490
  similar in size as the one by Myreen and finding a sufficiently
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1491
  strong one took us, like Myreen for his, something on the magnitude of
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1492
  weeks.
71
8c7f10b3da7b updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 63
diff changeset
  1493
138
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1494
  Our reasoning about the invariants is not much supported by the
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1495
  automation in Isabelle. There is however a tantalising connection
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1496
  between our work and very recent work \cite{Jensen13} on verifying
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1497
  X86 assembly code. They observed a similar phenomenon with assembly
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1498
  programs that Hoare-style reasoning is sometimes possible, but
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1499
  sometimes it is not. In order to ease their reasoning they
125
1ce74a77fa2a conclusion
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 124
diff changeset
  1500
  introduced a more primitive specification logic, on which
138
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1501
  Hoare-rules can be provided for special cases.  It remains to be
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1502
  seen whether their specification logic for assembly code can make it
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1503
  easier to reason about our Turing programs. That would be an
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1504
  attractive result, because Turing machine programs are very much
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1505
  like assmbly programs and it would connect some very classic work on
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1506
  Turing machines with very cutting-edge work on machine code
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1507
  verification. In order to try out such ideas, our formalisation provides the
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1508
  ``playground''. The code of our formalisation is available from the
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1509
  Mercurial repository at\\
7fa1b8e88d76 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
  1510
  \url{http://www.dcs.kcl.ac.uk/staff/urbanc/cgi-bin/repos.cgi/tm/}.
114
120091653998 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 113
diff changeset
  1511
*}
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1512
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1513
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1514
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1515
(*<*)
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1516
end
109
4635641e77cb updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 107
diff changeset
  1517
end
6
50880fcda34d added paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1518
(*>*)