utm/turing_basic.thy
author zhang
Sat, 29 Sep 2012 12:38:12 +0000
changeset 370 1ce04eb1c8ad
child 371 48b231495281
permissions -rw-r--r--
Initial upload of the formal construction of Universal Turing Machine.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
370
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     1
theory turing_basic
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     2
imports Main
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     3
begin
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     4
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     5
section {* Basic definitions of Turing machine *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     6
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     7
(* Title: Turing machine's definition and its charater
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     8
   Author: Xu Jian <xujian817@hotmail.com>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
     9
   Maintainer: Xu Jian
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    10
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    11
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    12
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    13
\label{description of turing machine}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    14
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    15
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    16
section {* Basic definitions of Turing machine *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    17
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    18
(* Title: Turing machine's definition and its charater
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    19
   Author: Xu Jian <xujian817@hotmail.com>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    20
   Maintainer: Xu Jian
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    21
*)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    22
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    23
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    24
  Actions of Turing machine (Abbreviated TM in the following* ).
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    25
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    26
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    27
datatype taction = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    28
  -- {* Write zero *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    29
  W0 | 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    30
  -- {* Write one *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    31
  W1 | 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    32
  -- {* Move left *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    33
  L | 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    34
  -- {* Move right *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    35
  R | 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    36
  -- {* Do nothing *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    37
  Nop
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    38
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    39
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    40
  Tape contents in every block.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    41
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    42
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    43
datatype block = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    44
  -- {* Blank *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    45
  Bk | 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    46
  -- {* Occupied *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    47
  Oc
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    48
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    49
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    50
  Tape is represented as a pair of lists $(L_{left}, L_{right})$,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    51
  where $L_left$, named {\em left list}, is used to represent
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    52
  the tape to the left of RW-head and
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    53
  $L_{right}$, named {\em right list}, is used to represent the tape
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    54
  under and to the right of RW-head.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    55
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    56
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    57
type_synonym tape = "block list \<times> block list"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    58
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    59
text {* The state of turing machine.*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    60
type_synonym tstate = nat
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    61
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    62
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    63
  Turing machine instruction is represented as a 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    64
  pair @{text "(action, next_state)"},
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    65
  where @{text "action"} is the action to take at the current state 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    66
  and @{text "next_state"} is the next state the machine is getting into
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    67
  after the action.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    68
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    69
type_synonym tinst = "taction \<times> tstate"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    70
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    71
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    72
  Program of Turing machine is represented as a list of Turing instructions
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    73
  and the execution of the program starts from the head of the list.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    74
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    75
type_synonym tprog = "tinst list"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    76
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    77
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    78
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    79
  Turing machine configuration, which consists of the current state 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    80
  and the tape.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    81
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    82
type_synonym t_conf = "tstate \<times> tape"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    83
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    84
fun nth_of ::  "'a list \<Rightarrow> nat \<Rightarrow> 'a option"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    85
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    86
  "nth_of xs n = (if n < length xs then Some (xs!n)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    87
                  else None)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    88
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    89
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    90
  The function used to fetech instruction out of Turing program.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    91
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    92
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    93
fun fetch :: "tprog \<Rightarrow> tstate \<Rightarrow> block \<Rightarrow> tinst"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    94
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    95
  "fetch p s b = (if s = 0 then (Nop, 0) else
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    96
                  case b of 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    97
                     Bk \<Rightarrow> case nth_of p (2 * (s - 1)) of
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    98
                          Some i \<Rightarrow> i
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
    99
                        | None \<Rightarrow> (Nop, 0) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   100
                   | Oc \<Rightarrow> case nth_of p (2 * (s - 1) +1) of
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   101
                          Some i \<Rightarrow> i
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   102
                        | None \<Rightarrow> (Nop, 0))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   103
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   104
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   105
fun new_tape :: "taction \<Rightarrow> tape \<Rightarrow> tape"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   106
where 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   107
   "new_tape action (leftn, rightn) = (case action of
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   108
                                         W0 \<Rightarrow> (leftn, Bk#(tl rightn)) |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   109
                                         W1 \<Rightarrow> (leftn, Oc#(tl rightn)) |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   110
                                         L  \<Rightarrow>  (if leftn = [] then (tl leftn, Bk#rightn)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   111
                                               else (tl leftn, (hd leftn) # rightn)) |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   112
                                         R  \<Rightarrow> if rightn = [] then (Bk#leftn,tl rightn) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   113
                                               else ((hd rightn)#leftn, tl rightn) |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   114
                                         Nop \<Rightarrow> (leftn, rightn)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   115
                                       )"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   116
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   117
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   118
  The one step function used to transfer Turing machine configuration.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   119
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   120
fun tstep :: "t_conf \<Rightarrow> tprog \<Rightarrow> t_conf"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   121
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   122
  "tstep c p = (let (s, l, r) = c in 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   123
                     let (ac, ns) = (fetch p s (case r of [] \<Rightarrow> Bk |     
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   124
                                                               x # xs \<Rightarrow> x)) in
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   125
                       (ns, new_tape ac (l, r)))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   126
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   127
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   128
  The many-step function.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   129
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   130
fun steps :: "t_conf \<Rightarrow> tprog \<Rightarrow> nat \<Rightarrow> t_conf"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   131
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   132
  "steps c p 0 = c" |
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   133
  "steps c p (Suc n) = steps (tstep c p) p n"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   134
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   135
lemma tstep_red: "steps c p (Suc n) = tstep (steps c p n) p"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   136
proof(induct n arbitrary: c)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   137
  fix c
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   138
  show "steps c p (Suc 0) = tstep (steps c p 0) p" by(simp add: steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   139
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   140
  fix n c
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   141
  assume ind: "\<And> c. steps c p (Suc n) = tstep (steps c p n) p"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   142
  have "steps (tstep c p) p (Suc n) = tstep (steps (tstep c p) p n) p"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   143
    by(rule ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   144
  thus "steps c p (Suc (Suc n)) = tstep (steps c p (Suc n)) p" by(simp add: steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   145
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   146
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   147
declare Let_def[simp] option.split[split]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   148
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   149
definition 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   150
  "iseven n \<equiv> \<exists> x. n = 2 * x"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   151
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   152
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   153
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   154
  The following @{text "t_correct"} function is used to specify the wellformedness of Turing
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   155
  machine.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   156
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   157
fun t_correct :: "tprog \<Rightarrow> bool"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   158
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   159
  "t_correct p = (length p \<ge> 2 \<and> iseven (length p) \<and> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   160
                   list_all (\<lambda> (acn, s). s \<le> length p div 2) p)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   161
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   162
declare t_correct.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   163
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   164
lemma allimp: "\<lbrakk>\<forall>x. P x \<longrightarrow> Q x; \<forall>x. P x\<rbrakk> \<Longrightarrow> \<forall>x. Q x"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   165
by(auto elim: allE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   166
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   167
lemma halt_lemma: "\<lbrakk>wf LE; \<forall> n. (\<not> P (f n) \<longrightarrow> (f (Suc n), (f n)) \<in> LE)\<rbrakk> \<Longrightarrow> \<exists> n. P (f n)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   168
apply(rule exCI, drule allimp, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   169
apply(drule_tac f = f  in wf_inv_image, simp add: inv_image_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   170
apply(erule wf_induct, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   171
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   172
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   173
lemma steps_add: "steps c t (x + y) = steps (steps c t x) t y"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   174
by(induct x arbitrary: c, auto simp: steps.simps tstep_red)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   175
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   176
lemma listall_set: "list_all p t \<Longrightarrow> \<forall> a \<in> set t. p a"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   177
by(induct t, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   178
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   179
lemma fetch_ex: "\<exists>b a. fetch T aa ab = (b, a)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   180
by(simp add: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   181
definition exponent :: "'a \<Rightarrow> nat \<Rightarrow> 'a list" ("_\<^bsup>_\<^esup>" [0, 0]100)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   182
  where "exponent x n = replicate n x"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   183
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   184
text {* 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   185
  @{text "tinres l1 l2"} means left list @{text "l1"} is congruent with
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   186
  @{text "l2"} with respect to the execution of Turing machine. 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   187
  Appending Blank to the right of eigther one does not affect the 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   188
  outcome of excution. 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   189
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   190
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   191
definition tinres :: "block list \<Rightarrow> block list \<Rightarrow> bool"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   192
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   193
  "tinres bx by = (\<exists> n. bx = by@Bk\<^bsup>n\<^esup> \<or> by = bx @ Bk\<^bsup>n\<^esup>)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   194
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   195
lemma exp_zero: "a\<^bsup>0\<^esup> = []"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   196
by(simp add: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   197
lemma exp_ind_def: "a\<^bsup>Suc x \<^esup> = a # a\<^bsup>x\<^esup>"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   198
by(simp add: exponent_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   199
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   200
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   201
  The following lemma shows the meaning of @{text "tinres"} with respect to 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   202
  one step execution.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   203
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   204
lemma tinres_step: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   205
  "\<lbrakk>tinres l l'; tstep (ss, l, r) t = (sa, la, ra); tstep (ss, l', r) t = (sb, lb, rb)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   206
    \<Longrightarrow> tinres la lb \<and> ra = rb \<and> sa = sb"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   207
apply(auto simp: tstep.simps fetch.simps new_tape.simps 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   208
        split: if_splits taction.splits list.splits
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   209
                 block.splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   210
apply(case_tac [!] "t ! (2 * (ss - Suc 0))", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   211
     auto simp: exponent_def tinres_def split: if_splits taction.splits list.splits
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   212
                 block.splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   213
apply(case_tac [!] "t ! (2 * (ss - Suc 0) + Suc 0)", 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   214
     auto simp: exponent_def tinres_def split: if_splits taction.splits list.splits
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   215
                 block.splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   216
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   217
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   218
declare tstep.simps[simp del] steps.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   219
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   220
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   221
  The following lemma shows the meaning of @{text "tinres"} with respect to 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   222
  many step execution.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   223
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   224
lemma tinres_steps: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   225
  "\<lbrakk>tinres l l'; steps (ss, l, r) t stp = (sa, la, ra); steps (ss, l', r) t stp = (sb, lb, rb)\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   226
    \<Longrightarrow> tinres la lb \<and> ra = rb \<and> sa = sb"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   227
apply(induct stp arbitrary: sa la ra sb lb rb, simp add: steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   228
apply(simp add: tstep_red)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   229
apply(case_tac "(steps (ss, l, r) t stp)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   230
apply(case_tac "(steps (ss, l', r) t stp)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   231
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   232
  fix stp sa la ra sb lb rb a b c aa ba ca
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   233
  assume ind: "\<And>sa la ra sb lb rb. \<lbrakk>steps (ss, l, r) t stp = (sa, la, ra); 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   234
          steps (ss, l', r) t stp = (sb, lb, rb)\<rbrakk> \<Longrightarrow> tinres la lb \<and> ra = rb \<and> sa = sb"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   235
  and h: " tinres l l'" "tstep (steps (ss, l, r) t stp) t = (sa, la, ra)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   236
         "tstep (steps (ss, l', r) t stp) t = (sb, lb, rb)" "steps (ss, l, r) t stp = (a, b, c)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   237
         "steps (ss, l', r) t stp = (aa, ba, ca)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   238
  have "tinres b ba \<and> c = ca \<and> a = aa"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   239
    apply(rule_tac ind, simp_all add: h)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   240
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   241
  thus "tinres la lb \<and> ra = rb \<and> sa = sb"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   242
    apply(rule_tac l = b and l' = ba and r = c  and ss = a   
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   243
            and t = t in tinres_step)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   244
    using h
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   245
    apply(simp, simp, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   246
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   247
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   248
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   249
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   250
  The following function @{text "tshift tp n"} is used to shift Turing programs 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   251
  @{text "tp"} by @{text "n"} when it is going to be combined with others.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   252
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   253
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   254
fun tshift :: "tprog \<Rightarrow> nat \<Rightarrow> tprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   255
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   256
  "tshift tp off = (map (\<lambda> (action, state). (action, (if state = 0 then 0
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   257
                                                      else state + off))) tp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   258
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   259
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   260
  When two Turing programs are combined, the end state (state @{text "0"}) of the one 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   261
  at the prefix position needs to be connected to the start state 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   262
  of the one at postfix position. If @{text "tp"} is the Turing program
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   263
  to be at the prefix, @{text "change_termi_state tp"} is the transformed Turing program.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   264
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   265
fun change_termi_state :: "tprog \<Rightarrow> tprog"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   266
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   267
  "change_termi_state t = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   268
       (map (\<lambda> (acn, ns). if ns = 0 then (acn, Suc ((length t) div 2)) else (acn, ns)) t)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   269
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   270
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   271
  @{text "t_add tp1 tp2"} is the combined Truing program.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   272
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   273
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   274
fun t_add :: "tprog \<Rightarrow> tprog \<Rightarrow> tprog" ("_ |+| _" [0, 0] 100)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   275
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   276
  "t_add t1 t2 = ((change_termi_state t1) @ (tshift t2 ((length t1) div 2)))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   277
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   278
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   279
  Tests whether the current configuration is at state @{text "0"}.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   280
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   281
definition isS0 :: "t_conf \<Rightarrow> bool"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   282
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   283
  "isS0 c = (let (s, l, r) = c in s = 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   284
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   285
declare tstep.simps[simp del] steps.simps[simp del] 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   286
        t_add.simps[simp del] fetch.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   287
        new_tape.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   288
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   289
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   290
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   291
  Single step execution starting from state @{text "0"} will not make any progress.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   292
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   293
lemma tstep_0: "tstep (0, tp) p = (0, tp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   294
apply(simp add: tstep.simps fetch.simps new_tape.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   295
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   296
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   297
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   298
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   299
  Many step executions starting from state @{text "0"} will not make any progress.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   300
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   301
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   302
lemma steps_0: "steps (0, tp) p stp = (0, tp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   303
apply(induct stp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   304
apply(simp add: steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   305
apply(simp add: tstep_red tstep_0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   306
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   307
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   308
lemma s_keep_step: "\<lbrakk>a \<le> length A div 2; tstep (a, b, c) A = (s, l, r); t_correct A\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   309
  \<Longrightarrow> s \<le> length A div 2"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   310
apply(simp add: tstep.simps fetch.simps t_correct.simps iseven_def 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   311
  split: if_splits block.splits list.splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   312
apply(case_tac [!] a, auto simp: list_all_length)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   313
apply(erule_tac x = "2 * nat" in allE, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   314
apply(erule_tac x = "2 * nat" in allE, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   315
apply(erule_tac x = "Suc (2 * nat)" in allE, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   316
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   317
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   318
lemma s_keep: "\<lbrakk>steps (Suc 0, tp) A stp = (s, l, r);  t_correct A\<rbrakk> \<Longrightarrow> s \<le> length A div 2"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   319
proof(induct stp arbitrary: s l r)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   320
  case 0 thus "?case" by(auto simp: t_correct.simps steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   321
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   322
  fix stp s l r
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   323
  assume ind: "\<And>s l r. \<lbrakk>steps (Suc 0, tp) A stp = (s, l, r); t_correct A\<rbrakk> \<Longrightarrow> s \<le> length A div 2"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   324
  and h1: "steps (Suc 0, tp) A (Suc stp) = (s, l, r)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   325
  and h2: "t_correct A"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   326
  from h1 h2 show "s \<le> length A div 2"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   327
  proof(simp add: tstep_red, cases "(steps (Suc 0, tp) A stp)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   328
    fix a b c 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   329
    assume h3: "tstep (a, b, c) A = (s, l, r)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   330
    and h4: "steps (Suc 0, tp) A stp = (a, b, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   331
    have "a \<le> length A div 2"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   332
      using h2 h4
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   333
      by(rule_tac l = b and r = c in ind, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   334
    thus "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   335
      using h3 h2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   336
      by(simp add: s_keep_step)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   337
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   338
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   339
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   340
lemma t_merge_fetch_pre:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   341
  "\<lbrakk>fetch A s b = (ac, ns); s \<le> length A div 2; t_correct A; s \<noteq> 0\<rbrakk> \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   342
  fetch (A |+| B) s b = (ac, if ns = 0 then Suc (length A div 2) else ns)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   343
apply(subgoal_tac "2 * (s - Suc 0) < length A \<and> Suc (2 * (s - Suc 0)) < length A")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   344
apply(auto simp: fetch.simps t_add.simps split: if_splits block.splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   345
apply(simp_all add: nth_append change_termi_state.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   346
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   347
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   348
lemma [simp]:  "\<lbrakk>\<not> a \<le> length A div 2; t_correct A\<rbrakk> \<Longrightarrow> fetch A a b = (Nop, 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   349
apply(auto simp: fetch.simps del: nth_of.simps split: block.splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   350
apply(case_tac [!] a, auto simp: t_correct.simps iseven_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   351
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   352
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   353
lemma  [elim]: "\<lbrakk>t_correct A; \<not> isS0 (tstep (a, b, c) A)\<rbrakk> \<Longrightarrow> a \<le> length A div 2"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   354
apply(rule_tac classical, auto simp: tstep.simps new_tape.simps isS0_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   355
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   356
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   357
lemma [elim]: "\<lbrakk>t_correct A; \<not> isS0 (tstep (a, b, c) A)\<rbrakk> \<Longrightarrow> 0 < a"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   358
apply(rule_tac classical, simp add: tstep_0 isS0_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   359
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   360
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   361
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   362
lemma t_merge_pre_eq_step: "\<lbrakk>tstep (a, b, c) A = cf; t_correct A; \<not> isS0 cf\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   363
        \<Longrightarrow> tstep (a, b, c) (A |+| B) = cf"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   364
apply(subgoal_tac "a \<le> length A div 2 \<and> a \<noteq> 0")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   365
apply(simp add: tstep.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   366
apply(case_tac "fetch A a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   367
apply(drule_tac B = B in t_merge_fetch_pre, simp, simp, simp, simp add: isS0_def, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   368
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   369
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   370
lemma t_merge_pre_eq:  "\<lbrakk>steps (Suc 0, tp) A stp = cf; \<not> isS0 cf; t_correct A\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   371
    \<Longrightarrow> steps (Suc 0, tp) (A |+| B) stp = cf"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   372
proof(induct stp arbitrary: cf)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   373
  case 0 thus "?case" by(simp add: steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   374
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   375
  fix stp cf
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   376
  assume ind: "\<And>cf. \<lbrakk>steps (Suc 0, tp) A stp = cf; \<not> isS0 cf; t_correct A\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   377
                 \<Longrightarrow> steps (Suc 0, tp) (A |+| B) stp = cf"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   378
  and h1: "steps (Suc 0, tp) A (Suc stp) = cf"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   379
  and h2: "\<not> isS0 cf"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   380
  and h3: "t_correct A"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   381
  from h1 h2 h3 show "steps (Suc 0, tp) (A |+| B) (Suc stp) = cf"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   382
  proof(simp add: tstep_red, cases "steps (Suc 0, tp) (A) stp", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   383
    fix a b c
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   384
    assume h4: "tstep (a, b, c) A = cf"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   385
    and h5: "steps (Suc 0, tp) A stp = (a, b, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   386
    have "steps (Suc 0, tp) (A |+| B) stp = (a, b, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   387
    proof(cases a)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   388
      case 0 thus "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   389
        using h4 h2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   390
        apply(simp add: tstep_0, cases cf, simp add: isS0_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   391
        done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   392
    next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   393
      case (Suc n) thus "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   394
        using h5 h3
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   395
        apply(rule_tac ind, auto simp: isS0_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   396
        done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   397
    qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   398
    thus "tstep (steps (Suc 0, tp) (A |+| B) stp) (A |+| B) = cf"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   399
      using h4 h5 h3 h2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   400
      apply(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   401
      apply(rule t_merge_pre_eq_step, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   402
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   403
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   404
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   405
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   406
declare nth.simps[simp del] tshift.simps[simp del] change_termi_state.simps[simp del]
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   407
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   408
lemma [simp]: "length (change_termi_state A) = length A"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   409
by(simp add: change_termi_state.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   410
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   411
lemma first_halt_point: "steps (Suc 0, tp) A stp = (0, tp')
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   412
 \<Longrightarrow> \<exists>stp. \<not> isS0 (steps (Suc 0, tp) A stp) \<and> steps (Suc 0, tp) A (Suc stp) = (0, tp')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   413
proof(induct stp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   414
  case 0  thus "?case" by(simp add: steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   415
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   416
  case (Suc n) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   417
  fix stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   418
  assume ind: "steps (Suc 0, tp) A stp = (0, tp') \<Longrightarrow> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   419
       \<exists>stp. \<not> isS0 (steps (Suc 0, tp) A stp) \<and> steps (Suc 0, tp) A (Suc stp) = (0, tp')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   420
    and h: "steps (Suc 0, tp) A (Suc stp) = (0, tp')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   421
  from h show "\<exists>stp. \<not> isS0 (steps (Suc 0, tp) A stp) \<and> steps (Suc 0, tp) A (Suc stp) = (0, tp')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   422
  proof(simp add: tstep_red, cases "steps (Suc 0, tp) A stp", simp, case_tac a)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   423
    fix a b c
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   424
    assume g1: "a = (0::nat)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   425
    and g2: "tstep (a, b, c) A = (0, tp')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   426
    and g3: "steps (Suc 0, tp) A stp = (a, b, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   427
    have "steps (Suc 0, tp) A stp = (0, tp')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   428
      using g2 g1 g3
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   429
      by(simp add: tstep_0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   430
    hence "\<exists> stp. \<not> isS0 (steps (Suc 0, tp) A stp) \<and> steps (Suc 0, tp) A (Suc stp) = (0, tp')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   431
      by(rule ind)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   432
    thus "\<exists>stp. \<not> isS0 (steps (Suc 0, tp) A stp) \<and> tstep (steps (Suc 0, tp) A stp) A = (0, tp')" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   433
      apply(simp add: tstep_red)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   434
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   435
  next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   436
    fix a b c nat
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   437
    assume g1: "steps (Suc 0, tp) A stp = (a, b, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   438
    and g2: "steps (Suc 0, tp) A (Suc stp) = (0, tp')" "a= Suc nat"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   439
    thus "\<exists>stp. \<not> isS0 (steps (Suc 0, tp) A stp) \<and> tstep (steps (Suc 0, tp) A stp) A = (0, tp')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   440
      apply(rule_tac x = stp in exI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   441
      apply(simp add: isS0_def tstep_red)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   442
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   443
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   444
qed 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   445
   
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   446
lemma t_merge_pre_halt_same': 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   447
  "\<lbrakk>\<not> isS0 (steps (Suc 0, tp) A stp) ; steps (Suc 0, tp) A (Suc stp) = (0, tp'); t_correct A\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   448
  \<Longrightarrow> steps (Suc 0, tp) (A |+| B) (Suc stp) = (Suc (length A div 2), tp')"    
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   449
proof(simp add: tstep_red, cases "steps (Suc 0, tp) A stp", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   450
  fix a b c 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   451
  assume h1: "\<not> isS0 (a, b, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   452
  and h2: "tstep (a, b, c) A = (0, tp')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   453
  and h3: "t_correct A"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   454
  and h4: "steps (Suc 0, tp) A stp = (a, b, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   455
  have "steps (Suc 0, tp) (A |+| B) stp = (a, b, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   456
    using h1 h4 h3
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   457
    apply(rule_tac  t_merge_pre_eq, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   458
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   459
  moreover have "tstep (a, b, c) (A |+| B) = (Suc (length A div 2), tp')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   460
    using h2 h3 h1 h4 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   461
    apply(simp add: tstep.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   462
    apply(case_tac " fetch A a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   463
    apply(drule_tac B = B in t_merge_fetch_pre, auto simp: isS0_def intro: s_keep)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   464
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   465
  ultimately show "tstep (steps (Suc 0, tp) (A |+| B) stp) (A |+| B) = (Suc (length A div 2), tp')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   466
    by(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   467
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   468
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   469
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   470
  When Turing machine @{text "A"} and @{text "B"} are combined and the execution
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   471
  of @{text "A"} can termination within @{text "stp"} steps, 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   472
  the combined machine @{text "A |+| B"} will eventually get into the starting 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   473
  state of machine @{text "B"}.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   474
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   475
lemma t_merge_pre_halt_same: "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   476
  \<lbrakk>steps (Suc 0, tp) A stp = (0, tp'); t_correct A; t_correct B\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   477
     \<Longrightarrow> \<exists> stp. steps (Suc 0, tp) (A |+| B) stp = (Suc (length A div 2), tp')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   478
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   479
  assume a_wf: "t_correct A"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   480
  and b_wf: "t_correct B"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   481
  and a_ht: "steps (Suc 0, tp) A stp = (0, tp')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   482
  have halt_point: "\<exists> stp. \<not> isS0 (steps (Suc 0, tp) A stp) \<and> steps (Suc 0, tp) A (Suc stp) = (0, tp')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   483
    using a_ht
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   484
    by(erule_tac first_halt_point)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   485
  then obtain stp' where "\<not> isS0 (steps (Suc 0, tp) A stp') \<and> steps (Suc 0, tp) A (Suc stp') = (0, tp')"..
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   486
  hence "steps (Suc 0, tp) (A |+| B) (Suc stp') = (Suc (length A div 2), tp')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   487
    using a_wf
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   488
    apply(rule_tac t_merge_pre_halt_same', auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   489
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   490
  thus "?thesis" ..
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   491
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   492
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   493
lemma fetch_0: "fetch p 0 b = (Nop, 0)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   494
by(simp add: fetch.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   495
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   496
lemma [simp]: "length (tshift B x) = length B"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   497
by(simp add: tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   498
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   499
lemma [simp]: "t_correct A \<Longrightarrow> 2 * (length A div 2) = length A"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   500
apply(simp add: t_correct.simps iseven_def, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   501
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   502
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   503
lemma t_merge_fetch_snd: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   504
  "\<lbrakk>fetch B a b = (ac, ns); t_correct A; t_correct B; a > 0 \<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   505
  \<Longrightarrow> fetch (A |+| B) (a + length A div 2) b
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   506
  = (ac, if ns = 0 then 0 else ns + length A div 2)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   507
apply(auto simp: fetch.simps t_add.simps split: if_splits block.splits)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   508
apply(case_tac [!] a, simp_all)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   509
apply(simp_all add: nth_append change_termi_state.simps tshift.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   510
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   511
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   512
lemma t_merge_snd_eq_step: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   513
  "\<lbrakk>tstep (s, l, r) B = (s', l', r'); t_correct A; t_correct B; s > 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   514
    \<Longrightarrow> tstep (s + length A div 2, l, r) (A |+| B) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   515
       (if s' = 0 then 0 else s' + length A div 2, l' ,r') "
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   516
apply(simp add: tstep.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   517
apply(cases "fetch B s (case r of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   518
apply(auto simp: t_merge_fetch_snd)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   519
apply(frule_tac [!] t_merge_fetch_snd, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   520
done 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   521
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   522
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   523
  Relates the executions of TM @{text "B"}, one is when @{text "B"} is executed alone,
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   524
  the other is the execution when @{text "B"} is in the combined TM.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   525
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   526
lemma t_merge_snd_eq_steps: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   527
  "\<lbrakk>t_correct A; t_correct B; steps (s, l, r) B stp = (s', l', r'); s > 0\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   528
  \<Longrightarrow> steps (s + length A div 2, l, r) (A |+| B) stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   529
      (if s' = 0 then 0 else s' + length A div 2, l', r')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   530
proof(induct stp arbitrary: s' l' r')
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   531
  case 0 thus "?case" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   532
    by(simp add: steps.simps)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   533
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   534
  fix stp s' l' r'
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   535
  assume ind: "\<And>s' l' r'. \<lbrakk>t_correct A; t_correct B; steps (s, l, r) B stp = (s', l', r'); 0 < s\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   536
                   \<Longrightarrow> steps (s + length A div 2, l, r) (A |+| B) stp = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   537
                          (if s' = 0 then 0 else s' + length A div 2, l', r')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   538
  and h1: "steps (s, l, r) B (Suc stp) = (s', l', r')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   539
  and h2: "t_correct A"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   540
  and h3: "t_correct B"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   541
  and h4: "0 < s"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   542
  from h1 show "steps (s + length A div 2, l, r) (A |+| B) (Suc stp) 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   543
            = (if s' = 0 then 0 else s' + length A div 2, l', r')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   544
  proof(simp only: tstep_red, cases "steps (s, l, r) B stp")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   545
    fix a b c 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   546
    assume h5: "steps (s, l, r) B stp = (a, b, c)" "tstep (steps (s, l, r) B stp) B = (s', l', r')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   547
    hence h6: "(steps (s + length A div 2, l, r) (A |+| B) stp) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   548
                ((if a = 0 then 0 else a + length A div 2, b, c))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   549
      using h2 h3 h4
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   550
      by(rule_tac ind, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   551
    thus "tstep (steps (s + length A div 2, l, r) (A |+| B) stp) (A |+| B) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   552
       (if s' = 0 then 0 else s'+ length A div 2, l', r')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   553
      using h5
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   554
    proof(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   555
      assume "tstep (0, b, c) B = (0, l', r')" thus "tstep (0, b, c) (A |+| B) = (0, l', r')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   556
        by(simp add: tstep_0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   557
    next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   558
      assume "tstep (0, b, c) B = (s', l', r')" "0 < s'"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   559
      thus "tstep (0, b, c) (A |+| B) = (s' + length A div 2, l', r')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   560
        by(simp add: tstep_0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   561
    next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   562
      assume "tstep (a, b, c) B = (0, l', r')" "0 < a"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   563
      thus "tstep (a + length A div 2, b, c) (A |+| B) = (0, l', r')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   564
        using h2 h3
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   565
        by(drule_tac t_merge_snd_eq_step, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   566
    next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   567
      assume "tstep (a, b, c) B = (s', l', r')" "0 < a" "0 < s'"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   568
      thus "tstep (a + length A div 2, b, c) (A |+| B) = (s' + length A div 2, l', r')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   569
        using h2 h3
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   570
        by(drule_tac t_merge_snd_eq_step, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   571
    qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   572
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   573
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   574
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   575
lemma t_merge_snd_halt_eq: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   576
  "\<lbrakk>steps (Suc 0, tp) B stp = (0, tp'); t_correct A; t_correct B\<rbrakk>
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   577
  \<Longrightarrow> \<exists>stp. steps (Suc (length A div 2), tp) (A |+| B) stp = (0, tp')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   578
apply(case_tac tp, cases tp', simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   579
apply(drule_tac  s = "Suc 0" in t_merge_snd_eq_steps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   580
done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   581
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   582
lemma t_inj: "\<lbrakk>steps (Suc 0, tp) A stpa = (0, tp1); steps (Suc 0, tp) A stpb = (0, tp2)\<rbrakk> 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   583
      \<Longrightarrow> tp1 = tp2"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   584
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   585
  assume h1: "steps (Suc 0, tp) A stpa = (0, tp1)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   586
  and h2: "steps (Suc 0, tp) A stpb = (0, tp2)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   587
  thus "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   588
  proof(cases "stpa < stpb")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   589
    case True thus "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   590
      using h1 h2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   591
      apply(drule_tac less_imp_Suc_add, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   592
      apply(simp del: add_Suc_right add_Suc add: add_Suc_right[THEN sym] steps_add steps_0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   593
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   594
  next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   595
    case False thus "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   596
      using h1 h2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   597
      apply(drule_tac leI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   598
      apply(case_tac "stpb = stpa", auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   599
      apply(subgoal_tac "stpb < stpa")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   600
      apply(drule_tac less_imp_Suc_add, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   601
      apply(simp del: add_Suc_right add_Suc add: add_Suc_right[THEN sym] steps_add steps_0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   602
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   603
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   604
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   605
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   606
type_synonym t_assert = "tape \<Rightarrow> bool"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   607
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   608
definition t_imply :: "t_assert \<Rightarrow> t_assert \<Rightarrow> bool" ("_ \<turnstile>-> _" [0, 0] 100)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   609
  where
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   610
  "t_imply a1 a2 = (\<forall> tp. a1 tp \<longrightarrow> a2 tp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   611
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   612
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   613
locale turing_merge =
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   614
  fixes A :: "tprog" and B :: "tprog" and P1 :: "t_assert"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   615
  and P2 :: "t_assert"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   616
  and P3 :: "t_assert"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   617
  and P4 :: "t_assert"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   618
  and Q1:: "t_assert"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   619
  and Q2 :: "t_assert"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   620
  assumes 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   621
  A_wf : "t_correct A"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   622
  and B_wf : "t_correct B"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   623
  and A_halt : "P1 tp \<Longrightarrow> \<exists> stp. let (s, tp') = steps (Suc 0, tp) A stp in s = 0 \<and> Q1 tp'"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   624
  and B_halt : "P2 tp \<Longrightarrow> \<exists> stp. let (s, tp') = steps (Suc 0, tp) B stp in s = 0 \<and> Q2 tp'"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   625
  and A_uhalt : "P3 tp \<Longrightarrow> \<not> (\<exists> stp. isS0 (steps (Suc 0, tp) A stp))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   626
  and B_uhalt: "P4 tp \<Longrightarrow> \<not> (\<exists> stp. isS0 (steps (Suc 0, tp) B stp))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   627
begin
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   628
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   629
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   630
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   631
  The following lemma tries to derive the Hoare logic rule for sequentially combined TMs.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   632
  It deals with the situtation when both @{text "A"} and @{text "B"} are terminated.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   633
*}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   634
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   635
lemma t_merge_halt: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   636
  assumes aimpb: "Q1 \<turnstile>-> P2"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   637
  shows "P1 \<turnstile>->  \<lambda> tp. (\<exists> stp tp'. steps (Suc 0, tp) (A |+| B)  stp = (0, tp') \<and> Q2 tp')"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   638
proof(simp add: t_imply_def, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   639
  fix a b
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   640
  assume h: "P1 (a, b)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   641
  hence "\<exists> stp. let (s, tp') = steps (Suc 0, a, b) A stp in s = 0 \<and> Q1 tp'"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   642
    using A_halt by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   643
  from this obtain stp1 where "let (s, tp') = steps (Suc 0, a, b) A stp1 in s = 0 \<and> Q1 tp'" ..
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   644
  thus "\<exists>stp aa ba. steps (Suc 0, a, b) (A |+| B) stp = (0, aa, ba) \<and> Q2 (aa, ba)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   645
  proof(case_tac "steps (Suc 0, a, b) A stp1", simp, erule_tac conjE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   646
    fix aa ba c
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   647
    assume g1: "Q1 (ba, c)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   648
      and g2: "steps (Suc 0, a, b) A stp1 = (0, ba, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   649
    hence "P2 (ba, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   650
      using aimpb apply(simp add: t_imply_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   651
      done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   652
    hence "\<exists> stp. let (s, tp') = steps (Suc 0, ba, c) B stp in s = 0 \<and> Q2 tp'"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   653
      using B_halt by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   654
    from this obtain stp2 where "let (s, tp') = steps (Suc 0, ba, c) B stp2 in s = 0 \<and> Q2 tp'" ..
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   655
    thus "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   656
    proof(case_tac "steps (Suc 0, ba, c) B stp2", simp, erule_tac conjE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   657
      fix aa bb ca
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   658
      assume g3: " Q2 (bb, ca)" "steps (Suc 0, ba, c) B stp2 = (0, bb, ca)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   659
      have "\<exists> stp. steps (Suc 0, a, b) (A |+| B) stp = (Suc (length A div 2), ba , c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   660
        using g2 A_wf B_wf
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   661
        by(rule_tac t_merge_pre_halt_same, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   662
      moreover have "\<exists> stp. steps (Suc (length A div 2), ba, c) (A |+| B) stp = (0, bb, ca)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   663
        using g3 A_wf B_wf
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   664
        apply(rule_tac t_merge_snd_halt_eq, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   665
        done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   666
      ultimately show "\<exists>stp aa ba. steps (Suc 0, a, b) (A |+| B) stp = (0, aa, ba) \<and> Q2 (aa, ba)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   667
        apply(erule_tac exE, erule_tac exE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   668
        apply(rule_tac x = "stp + stpa" in exI, simp add: steps_add)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   669
        using g3 by simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   670
    qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   671
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   672
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   673
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   674
lemma  t_merge_uhalt_tmp:
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   675
  assumes B_uh: "\<forall>stp. \<not> isS0 (steps (Suc 0, b, c) B stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   676
  and merge_ah: "steps (Suc 0, tp) (A |+| B) stpa = (Suc (length A div 2), b, c)" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   677
  shows "\<forall> stp. \<not> isS0 (steps (Suc 0, tp) (A |+| B) stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   678
  using B_uh merge_ah
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   679
apply(rule_tac allI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   680
apply(case_tac "stp > stpa")
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   681
apply(erule_tac x = "stp - stpa" in allE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   682
apply(case_tac "(steps (Suc 0, b, c) B (stp - stpa))", simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   683
proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   684
  fix stp a ba ca 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   685
  assume h1: "\<not> isS0 (a, ba, ca)" "stpa < stp"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   686
  and h2: "steps (Suc 0, b, c) B (stp - stpa) = (a, ba, ca)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   687
  have "steps (Suc 0 + length A div 2, b, c) (A |+| B) (stp - stpa) = 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   688
      (if a = 0 then 0 else a + length A div 2, ba, ca)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   689
    using A_wf B_wf h2
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   690
    by(rule_tac t_merge_snd_eq_steps, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   691
  moreover have "a > 0" using h1 by(simp add: isS0_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   692
  moreover have "\<exists> stpb. stp = stpa + stpb" 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   693
    using h1 by(rule_tac x = "stp - stpa" in exI, simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   694
  ultimately show "\<not> isS0 (steps (Suc 0, tp) (A |+| B) stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   695
    using merge_ah
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   696
    by(auto simp: steps_add isS0_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   697
next
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   698
  fix stp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   699
  assume h: "steps (Suc 0, tp) (A |+| B) stpa = (Suc (length A div 2), b, c)" "\<not> stpa < stp"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   700
  hence "\<exists> stpb. stpa = stp + stpb" apply(rule_tac x = "stpa - stp" in exI, auto) done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   701
  thus "\<not> isS0 (steps (Suc 0, tp) (A |+| B) stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   702
    using h
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   703
    apply(auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   704
    apply(cases "steps (Suc 0, tp) (A |+| B) stp", simp add: steps_add isS0_def steps_0)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   705
    done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   706
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   707
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   708
text {*
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   709
  The following lemma deals with the situation when TM @{text "B"} can not terminate.
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   710
  *}
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   711
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   712
lemma t_merge_uhalt: 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   713
  assumes aimpb: "Q1 \<turnstile>-> P4"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   714
  shows "P1 \<turnstile>-> \<lambda> tp. \<not> (\<exists> stp. isS0 (steps (Suc 0, tp) (A |+| B) stp))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   715
proof(simp only: t_imply_def, rule_tac allI, rule_tac impI)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   716
  fix tp 
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   717
  assume init_asst: "P1 tp"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   718
  show "\<not> (\<exists>stp. isS0 (steps (Suc 0, tp) (A |+| B) stp))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   719
  proof -
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   720
    have "\<exists> stp. let (s, tp') = steps (Suc 0, tp) A stp in s = 0 \<and> Q1 tp'"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   721
      using A_halt[of tp] init_asst
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   722
      by(simp)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   723
    from this obtain stpx where "let (s, tp') = steps (Suc 0, tp) A stpx in s = 0 \<and> Q1 tp'" ..
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   724
    thus "?thesis"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   725
    proof(cases "steps (Suc 0, tp) A stpx", simp, erule_tac conjE)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   726
      fix a b c
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   727
      assume "Q1 (b, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   728
        and h3: "steps (Suc 0, tp) A stpx = (0, b, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   729
      hence h2: "P4 (b, c)"  using aimpb
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   730
        by(simp add: t_imply_def)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   731
      have "\<exists> stp. steps (Suc 0, tp) (A |+| B) stp = (Suc (length A div 2), b, c)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   732
        using h3 A_wf B_wf
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   733
        apply(rule_tac stp = stpx in t_merge_pre_halt_same, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   734
        done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   735
      from this obtain stpa where h4:"steps (Suc 0, tp) (A |+| B) stpa = (Suc (length A div 2), b, c)" ..
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   736
      have " \<not> (\<exists> stp. isS0 (steps (Suc 0, b, c) B stp))"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   737
        using B_uhalt [of "(b, c)"] h2 apply simp
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   738
        done
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   739
      from this and h4 show "\<forall>stp. \<not> isS0 (steps (Suc 0, tp) (A |+| B) stp)"
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   740
        by(rule_tac t_merge_uhalt_tmp, auto)
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   741
    qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   742
  qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   743
qed
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   744
end
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   745
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   746
end
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
   747